aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc
diff options
context:
space:
mode:
Diffstat (limited to 'powerpc')
-rw-r--r--powerpc/Archi.v2
-rw-r--r--powerpc/AsmToJSON.ml1
-rw-r--r--powerpc/Asmgen.v14
-rw-r--r--powerpc/Asmgenproof.v8
-rw-r--r--powerpc/Asmgenproof1.v5
l---------powerpc/BTL_SEsimplify.v1
-rw-r--r--powerpc/CSE2deps.v35
-rw-r--r--powerpc/CSE2depsproof.v214
-rw-r--r--powerpc/DuplicateOpcodeHeuristic.ml41
l---------powerpc/ExpansionOracle.ml1
-rw-r--r--powerpc/Machregsaux.ml5
-rw-r--r--powerpc/Machregsaux.mli2
-rw-r--r--powerpc/Op.v109
l---------powerpc/PrepassSchedulingOracle.ml1
-rw-r--r--powerpc/SelectLong.vp2
-rw-r--r--powerpc/SelectLongproof.v1
-rw-r--r--powerpc/SelectOp.vp11
-rw-r--r--powerpc/SelectOpproof.v29
-rw-r--r--powerpc/TargetPrinter.ml8
19 files changed, 470 insertions, 20 deletions
diff --git a/powerpc/Archi.v b/powerpc/Archi.v
index 28859051..5b0af3b6 100644
--- a/powerpc/Archi.v
+++ b/powerpc/Archi.v
@@ -71,3 +71,5 @@ Global Opaque ptr64 big_endian splitlong
default_nan_32 choose_nan_32
fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
+
+Definition has_notrap_loads := false.
diff --git a/powerpc/AsmToJSON.ml b/powerpc/AsmToJSON.ml
index 09cfc28d..1f32dd62 100644
--- a/powerpc/AsmToJSON.ml
+++ b/powerpc/AsmToJSON.ml
@@ -362,6 +362,7 @@ let pp_instructions pp ic =
| EF_annot_val _
| EF_builtin _
| EF_debug _
+ | EF_profiling _
| EF_external _
| EF_free
| EF_malloc
diff --git a/powerpc/Asmgen.v b/powerpc/Asmgen.v
index 7b6ac9af..924773e7 100644
--- a/powerpc/Asmgen.v
+++ b/powerpc/Asmgen.v
@@ -831,8 +831,13 @@ Definition transl_memory_access
Error(msg "Asmgen.transl_memory_access")
end.
-Definition transl_load (chunk: memory_chunk) (addr: addressing)
- (args: list mreg) (dst: mreg) (k: code) :=
+Definition transl_load
+ (trap : trapping_mode)
+ (chunk: memory_chunk) (addr: addressing)
+ (args: list mreg) (dst: mreg) (k: code) :=
+ match trap with
+ | NOTRAP => Error (msg "Asmgen.transl_load non-trapping loads unsupported on PPC")
+ | TRAP =>
match chunk with
| Mint8signed =>
do r <- ireg_of dst;
@@ -860,6 +865,7 @@ Definition transl_load (chunk: memory_chunk) (addr: addressing)
transl_memory_access (Plfd r) (Plfdx r) true addr args GPR12 k
| _ =>
Error (msg "Asmgen.transl_load")
+ end
end.
Definition transl_store (chunk: memory_chunk) (addr: addressing)
@@ -917,8 +923,8 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction)
loadind GPR1 f.(fn_link_ofs) Tint R11 k1)
| Mop op args res =>
transl_op op args res k
- | Mload chunk addr args dst =>
- transl_load chunk addr args dst k
+ | Mload trap chunk addr args dst =>
+ transl_load trap chunk addr args dst k
| Mstore chunk addr args src =>
transl_store chunk addr args src k
| Mcall sig (inl r) =>
diff --git a/powerpc/Asmgenproof.v b/powerpc/Asmgenproof.v
index 85541118..e25ae583 100644
--- a/powerpc/Asmgenproof.v
+++ b/powerpc/Asmgenproof.v
@@ -339,6 +339,7 @@ Proof.
eapply loadind_label; eauto.
eapply tail_nolabel_trans; eapply loadind_label; eauto.
eapply transl_op_label; eauto.
+ destruct t; try discriminate.
destruct m; monadInv H; (eapply tail_nolabel_trans; [eapply transl_memory_access_label; TailNoLabel|TailNoLabel]).
destruct m; monadInv H; eapply transl_memory_access_label; TailNoLabel.
destruct s0; monadInv H; TailNoLabel.
@@ -668,6 +669,13 @@ Opaque loadind.
split. simpl; congruence.
apply R; auto with asmgen.
+
+- (* Mload notrap *) (* isn't there a nicer way? *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
+- (* Mload notrap *)
+ inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate.
+
- (* Mstore *)
assert (eval_addressing tge sp addr rs##args = Some a).
rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
diff --git a/powerpc/Asmgenproof1.v b/powerpc/Asmgenproof1.v
index 6ae520ef..cc9fc1a7 100644
--- a/powerpc/Asmgenproof1.v
+++ b/powerpc/Asmgenproof1.v
@@ -1738,8 +1738,8 @@ Qed.
(** Translation of loads *)
Lemma transl_load_correct:
- forall chunk addr args dst k c (rs: regset) m a v,
- transl_load chunk addr args dst k = OK c ->
+ forall trap chunk addr args dst k c (rs: regset) m a v,
+ transl_load trap chunk addr args dst k = OK c ->
eval_addressing ge (rs#GPR1) addr (map rs (map preg_of args)) = Some a ->
Mem.loadv chunk m a = Some v ->
exists rs',
@@ -1748,6 +1748,7 @@ Lemma transl_load_correct:
/\ forall r, r <> PC -> r <> GPR12 -> r <> GPR0 -> r <> preg_of dst -> rs' r = rs r.
Proof.
intros.
+ destruct trap; try discriminate.
assert (LD: forall v, Val.lessdef a v -> v = a).
{ intros. inv H2; auto. discriminate H1. }
assert (BASE: forall mk1 mk2 unaligned k' chunk' v',
diff --git a/powerpc/BTL_SEsimplify.v b/powerpc/BTL_SEsimplify.v
new file mode 120000
index 00000000..f190e6d5
--- /dev/null
+++ b/powerpc/BTL_SEsimplify.v
@@ -0,0 +1 @@
+../aarch64/BTL_SEsimplify.v \ No newline at end of file
diff --git a/powerpc/CSE2deps.v b/powerpc/CSE2deps.v
new file mode 100644
index 00000000..4592f408
--- /dev/null
+++ b/powerpc/CSE2deps.v
@@ -0,0 +1,35 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* David Monniaux CNRS, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+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
+ | (Ainstack ofs), (Ainstack ofs'), _, _ =>
+ negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk)
+ | _, _, _, _ => true
+ end.
diff --git a/powerpc/CSE2depsproof.v b/powerpc/CSE2depsproof.v
new file mode 100644
index 00000000..ede09dd6
--- /dev/null
+++ b/powerpc/CSE2depsproof.v
@@ -0,0 +1,214 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* David Monniaux CNRS, VERIMAG *)
+(* *)
+(* Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+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%nat 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.
+
+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.
+ 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 : 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 = 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 (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: 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 (Int.unsigned ofsr) chunkr (Int.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.
+
+Section STACK_WRITE.
+ Variable m m2 : mem.
+ Variable chunkw chunkr : memory_chunk.
+
+ 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
+ (Ainstack ofsw) nil = Some addrw.
+ Hypothesis ADDRR : eval_addressing genv sp
+ (Ainstack ofsr) nil = Some addrr.
+
+ Lemma stack_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 *.
+
+ inv ADDRR.
+ inv ADDRW.
+
+ destruct sp; 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 stack_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 stack_load_store_away1.
+ all: tauto.
+ Qed.
+ End INDEXED_AWAY.
+End STACK_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 (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.
+- (* Ainstack / Ainstack *)
+ destruct args. 2: discriminate.
+ destruct args'. 2: discriminate.
+ cbn in OVERLAP.
+ destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP.
+ 2: discriminate.
+ cbn in *.
+ eapply stack_load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption.
+Qed.
+
+End SOUNDNESS.
diff --git a/powerpc/DuplicateOpcodeHeuristic.ml b/powerpc/DuplicateOpcodeHeuristic.ml
new file mode 100644
index 00000000..c48fdfba
--- /dev/null
+++ b/powerpc/DuplicateOpcodeHeuristic.ml
@@ -0,0 +1,41 @@
+(* *************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Sylvain Boulmé Grenoble-INP, VERIMAG *)
+(* David Monniaux CNRS, VERIMAG *)
+(* Cyril Six Kalray *)
+(* *)
+(* Copyright Kalray. Copyright VERIMAG. All rights reserved. *)
+(* This file is distributed under the terms of the INRIA *)
+(* Non-Commercial License Agreement. *)
+(* *)
+(* *************************************************************)
+
+(* 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
diff --git a/powerpc/ExpansionOracle.ml b/powerpc/ExpansionOracle.ml
new file mode 120000
index 00000000..ee2674bf
--- /dev/null
+++ b/powerpc/ExpansionOracle.ml
@@ -0,0 +1 @@
+../aarch64/ExpansionOracle.ml \ No newline at end of file
diff --git a/powerpc/Machregsaux.ml b/powerpc/Machregsaux.ml
index 9d3a2243..d17382ad 100644
--- a/powerpc/Machregsaux.ml
+++ b/powerpc/Machregsaux.ml
@@ -13,3 +13,8 @@
(** Auxiliary functions on machine registers *)
let is_scratch_register s = s = "R0" || s = "r0"
+
+let class_of_type = function
+ | AST.Tint | AST.Tlong -> 0
+ | AST.Tfloat | AST.Tsingle -> 1
+ | AST.Tany32 | AST.Tany64 -> assert false
diff --git a/powerpc/Machregsaux.mli b/powerpc/Machregsaux.mli
index f3d52849..01b0f9fd 100644
--- a/powerpc/Machregsaux.mli
+++ b/powerpc/Machregsaux.mli
@@ -13,3 +13,5 @@
(** Auxiliary functions on machine registers *)
val is_scratch_register: string -> bool
+
+val class_of_type: AST.typ -> int
diff --git a/powerpc/Op.v b/powerpc/Op.v
index a96a0439..505b7545 100644
--- a/powerpc/Op.v
+++ b/powerpc/Op.v
@@ -569,6 +569,35 @@ Proof with (try exact I; try reflexivity).
unfold Val.select. destruct (eval_condition c vl m). apply Val.normalize_type. exact I.
Qed.
+Definition is_trapping_op (op : operation) :=
+ match op with
+ | Odiv | Odivl | Odivu | Odivlu
+ | Oshrximm _ | Oshrxlimm _
+ | Ointoffloat
+ | Olongoffloat
+ | Ofloatoflong => true
+ | _ => false
+ end.
+
+Definition args_of_operation op :=
+ if eq_operation op Omove
+ then 1%nat
+ else List.length (fst (type_of_operation op)).
+
+
+Lemma is_trapping_op_sound:
+ forall op vl sp m,
+ is_trapping_op op = false ->
+ (List.length vl) = args_of_operation op ->
+ eval_operation genv sp op vl m <> None.
+Proof.
+ unfold args_of_operation.
+ destruct op; destruct eq_operation; intros; simpl in *; try congruence.
+ all: try (destruct vl as [ | vh1 vl1]; try discriminate).
+ all: try (destruct vl1 as [ | vh2 vl2]; try discriminate).
+ all: try (destruct vl2 as [ | vh3 vl3]; try discriminate).
+ all: try (destruct vl3 as [ | vh4 vl4]; try discriminate).
+Qed.
End SOUNDNESS.
(** * Manipulating and transforming operations *)
@@ -719,7 +748,7 @@ Definition is_trivial_op (op: operation) : bool :=
(** Operations that depend on the memory state. *)
-Definition condition_depends_on_memory (c: condition) : bool :=
+Definition cond_depends_on_memory (c: condition) : bool :=
match c with
| Ccompu _ => true
| Ccompuimm _ _ => true
@@ -730,14 +759,14 @@ Definition condition_depends_on_memory (c: condition) : bool :=
Definition op_depends_on_memory (op: operation) : bool :=
match op with
- | Ocmp c => condition_depends_on_memory c
- | Osel c ty => condition_depends_on_memory c
+ | Ocmp c => cond_depends_on_memory c
+ | Osel c ty => cond_depends_on_memory c
| _ => false
end.
-Lemma condition_depends_on_memory_correct:
+Lemma cond_depends_on_memory_correct:
forall c args m1 m2,
- condition_depends_on_memory c = false ->
+ cond_depends_on_memory c = false ->
eval_condition c args m1 = eval_condition c args m2.
Proof.
intros. destruct c; simpl; auto; discriminate.
@@ -749,12 +778,36 @@ Lemma op_depends_on_memory_correct:
eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
Proof.
intros until m2. destruct op; simpl; try congruence; intros C.
-- f_equal; f_equal; apply condition_depends_on_memory_correct; auto.
+- f_equal; f_equal; apply cond_depends_on_memory_correct; auto.
- destruct args; auto. destruct args; auto.
- rewrite (condition_depends_on_memory_correct c args m1 m2 C).
+ rewrite (cond_depends_on_memory_correct c args m1 m2 C).
auto.
Qed.
+Lemma cond_valid_pointer_eq:
+ forall cond args m1 m2,
+ (forall b z, Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) ->
+ eval_condition cond args m1 = eval_condition cond args m2.
+Proof.
+ intros until m2. intro MEM. destruct cond eqn:COND; simpl; try congruence.
+ all: repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
+Lemma op_valid_pointer_eq:
+ forall (F V: Type) (ge: Genv.t F V) sp op args m1 m2,
+ (forall b z, Mem.valid_pointer m1 b z = Mem.valid_pointer m2 b z) ->
+ eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
+Proof.
+ intros until m2. destruct op eqn:OP; simpl; try congruence.
+ - intros MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+ - intro MEM; destruct c; simpl; try congruence;
+ repeat (destruct args; simpl; try congruence);
+ erewrite cmpu_bool_valid_pointer_eq || erewrite cmplu_bool_valid_pointer_eq; eauto.
+Qed.
+
(** Global variables mentioned in an operation or addressing mode *)
Definition globals_operation (op: operation) : list ident :=
@@ -1016,6 +1069,21 @@ Proof.
apply Val.add_inject; auto. apply H; simpl; auto.
Qed.
+
+Lemma eval_addressing_inj_none:
+ forall addr sp1 vl1 sp2 vl2,
+ (forall id ofs,
+ In id (globals_addressing addr) ->
+ Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
+ Val.inject f sp1 sp2 ->
+ Val.inject_list f vl1 vl2 ->
+ eval_addressing ge1 sp1 addr vl1 = None ->
+ eval_addressing ge2 sp2 addr vl2 = None.
+Proof.
+ intros until vl2. intros Hglobal Hinjsp Hinjvl.
+ destruct addr; simpl in *;
+ inv Hinjvl; trivial; try discriminate; inv H0; trivial; try discriminate; inv H2; trivial; try discriminate.
+Qed.
End EVAL_COMPAT.
(** Compatibility of the evaluation functions with the ``is less defined'' relation over values. *)
@@ -1082,6 +1150,20 @@ Proof.
rewrite <- val_inject_list_lessdef. eauto. auto.
Qed.
+
+Lemma eval_addressing_lessdef_none:
+ forall sp addr vl1 vl2,
+ Val.lessdef_list vl1 vl2 ->
+ eval_addressing genv sp addr vl1 = None ->
+ eval_addressing genv sp addr vl2 = None.
+Proof.
+ intros until vl2. intros Hlessdef Heval1.
+ destruct addr; simpl in *;
+ inv Hlessdef; trivial; try discriminate;
+ inv H0; trivial; try discriminate;
+ inv H2; trivial; try discriminate.
+Qed.
+
Lemma eval_operation_lessdef:
forall sp op vl1 vl2 v1 m1 m2,
Val.lessdef_list vl1 vl2 ->
@@ -1173,6 +1255,19 @@ Proof.
econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
Qed.
+Lemma eval_addressing_inject_none:
+ forall addr vl1 vl2,
+ Val.inject_list f vl1 vl2 ->
+ eval_addressing genv (Vptr sp1 Ptrofs.zero) addr vl1 = None ->
+ eval_addressing genv (Vptr sp2 Ptrofs.zero) (shift_stack_addressing delta addr) vl2 = None.
+Proof.
+ intros.
+ rewrite eval_shift_stack_addressing.
+ eapply eval_addressing_inj_none with (sp1 := Vptr sp1 Ptrofs.zero); eauto.
+ intros. apply symbol_address_inject.
+ econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
+Qed.
+
Lemma eval_operation_inject:
forall op vl1 vl2 v1 m1 m2,
Val.inject_list f vl1 vl2 ->
diff --git a/powerpc/PrepassSchedulingOracle.ml b/powerpc/PrepassSchedulingOracle.ml
new file mode 120000
index 00000000..9885fd52
--- /dev/null
+++ b/powerpc/PrepassSchedulingOracle.ml
@@ -0,0 +1 @@
+../x86/PrepassSchedulingOracle.ml \ No newline at end of file
diff --git a/powerpc/SelectLong.vp b/powerpc/SelectLong.vp
index 5f13774b..e4274ba5 100644
--- a/powerpc/SelectLong.vp
+++ b/powerpc/SelectLong.vp
@@ -16,7 +16,7 @@ Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats.
Require Import Op CminorSel.
-Require Import SelectOp SplitLong.
+Require Import OpHelpers SelectOp SplitLong.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
diff --git a/powerpc/SelectLongproof.v b/powerpc/SelectLongproof.v
index ea14668f..2264451d 100644
--- a/powerpc/SelectLongproof.v
+++ b/powerpc/SelectLongproof.v
@@ -16,6 +16,7 @@ Require Import String Coqlib Maps Zbits Integers Floats Errors.
Require Archi.
Require Import AST Values Memory Globalenvs Events.
Require Import Cminor Op CminorSel.
+Require Import OpHelpers OpHelpersproof.
Require Import SelectOp SelectOpproof SplitLong SplitLongproof.
Require Import SelectLong.
diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp
index cd9a65df..fe8b5453 100644
--- a/powerpc/SelectOp.vp
+++ b/powerpc/SelectOp.vp
@@ -39,7 +39,7 @@
Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats Builtins.
-Require Import Op CminorSel.
+Require Import Op OpHelpers CminorSel.
Require Archi.
Local Open Scope cminorsel_scope.
@@ -472,7 +472,7 @@ Definition intuoffloat (e: expr) :=
else
Elet e
(Elet (Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil)
- (Econdition (CEcond (Ccompf Clt) (Eletvar 1 ::: Eletvar 0 ::: Enil))
+ (Econdition (CEcond (Ccompf Clt) None (Eletvar 1 ::: Eletvar 0 ::: Enil))
(intoffloat (Eletvar 1))
(addimm Float.ox8000_0000 (intoffloat (subf (Eletvar 1) (Eletvar 0))))))%nat.
@@ -566,6 +566,13 @@ Nondetfunction builtin_arg (e: expr) :=
| _ => BA e
end.
+(* floats *)
+Definition divf_base (e1: expr) (e2: expr) :=
+ Eop Odivf (e1 ::: e2 ::: Enil).
+
+Definition divfs_base (e1: expr) (e2: expr) :=
+ Eop Odivfs (e1 ::: e2 ::: Enil).
+
(** Platform-specific known builtins *)
Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v
index adac6c34..edc935d4 100644
--- a/powerpc/SelectOpproof.v
+++ b/powerpc/SelectOpproof.v
@@ -18,6 +18,8 @@ Require Import Values Memory Builtins Globalenvs.
Require Import Cminor Op CminorSel.
Require Import Compopts.
Require Import SelectOp.
+Require Import OpHelpers.
+Require Import OpHelpersproof.
Local Open Scope cminorsel_scope.
Local Transparent Archi.ptr64.
@@ -70,8 +72,10 @@ Ltac TrivialExists :=
(** * Correctness of the smart constructors *)
Section CMCONSTR.
-
-Variable ge: genv.
+Variable prog: program.
+Variable hf: helper_functions.
+Hypothesis HELPERS: helper_functions_declared prog hf.
+Let ge := Genv.globalenv prog.
Variable sp: val.
Variable e: env.
Variable m: mem.
@@ -1064,6 +1068,27 @@ Proof.
- constructor; auto.
Qed.
+(* floating-point division without HELPERS *)
+Theorem eval_divf_base:
+ forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v.
+Proof.
+ intros; unfold divf_base.
+ TrivialExists.
+Qed.
+
+Theorem eval_divfs_base:
+ forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v.
+Proof.
+ intros; unfold divfs_base.
+ TrivialExists.
+Qed.
+
(** Platform-specific known builtins *)
Theorem eval_platform_builtin:
diff --git a/powerpc/TargetPrinter.ml b/powerpc/TargetPrinter.ml
index 52d30e33..a82fa5d9 100644
--- a/powerpc/TargetPrinter.ml
+++ b/powerpc/TargetPrinter.ml
@@ -117,7 +117,9 @@ module Linux_System : SYSTEM =
let name_of_section = function
| Section_text -> ".text"
- | Section_data i ->
+ | Section_data(i, true) ->
+ failwith "_Thread_local unsupported on this platform"
+ | Section_data(i, false) ->
variable_section ~sec:".data" ~bss:".section .bss" i
| Section_small_data i ->
variable_section
@@ -212,7 +214,9 @@ module Diab_System : SYSTEM =
let name_of_section = function
| Section_text -> ".text"
- | Section_data i ->
+ | Section_data(i, true) ->
+ failwith "_Thread_local unsupported on this platform"
+ | Section_data (i, false) ->
variable_section ~sec:".data" ~bss:".bss" i
| Section_small_data i ->
variable_section ~sec:".sdata" ~bss:".sbss" ~common:false i