diff --git a/Makefile b/Makefile index 6465825538..e87f315c10 100644 --- a/Makefile +++ b/Makefile @@ -86,7 +86,9 @@ BACKEND=\ CSEdomain.v CombineOp.v CSE.v CombineOpproof.v CSEproof.v \ NeedDomain.v NeedOp.v Deadcode.v Deadcodeproof.v \ Unusedglob.v Unusedglobproof.v \ - Machregs.v Locations.v Conventions1.v Conventions.v LTL.v \ + Machregs.v FragmentBlock.v Registerfile.v Locations.v \ + Conventions1.v Conventions.v \ + LTL.v LTLtyping.v \ Allocation.v Allocproof.v \ Tunneling.v Tunnelingproof.v \ Linear.v Lineartyping.v \ diff --git a/arm/Asm.v b/arm/Asm.v index 07dea75608..194ff840e8 100644 --- a/arm/Asm.v +++ b/arm/Asm.v @@ -59,6 +59,22 @@ Proof. decide equality. Defined. Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}. Proof. decide equality. Defined. +Definition ireg_index (i: ireg): Z := + 1 + match i with + | IR0 => 0 | IR1 => 1 | IR2 => 2 | IR3 => 3 + | IR4 => 4 | IR5 => 5 | IR6 => 6 | IR7 => 7 + | IR8 => 8 | IR9 => 9 | IR10 => 10 | IR11 => 11 + | IR12 => 12 | IR13 => 13 | IR14 => 14 + end. + +Definition freg_index (f: freg): Z := + 16 + match f with + | FR0 => 0 | FR1 => 2 | FR2 => 4 | FR3 => 6 + | FR4 => 8 | FR5 => 10 | FR6 => 12 | FR7 => 14 + | FR8 => 16 | FR9 => 18 | FR10 => 20 | FR11 => 22 + | FR12 => 24 | FR13 => 26 | FR14 => 28 | FR15 => 30 + end. + (** Bits in the condition register. *) Inductive crbit: Type := @@ -85,12 +101,123 @@ Coercion CR: crbit >-> preg. Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}. Proof. decide equality. apply ireg_eq. apply freg_eq. apply crbit_eq. Defined. -Module PregEq. - Definition t := preg. - Definition eq := preg_eq. -End PregEq. +Definition crbit_index (c: crbit): Z := + 16 + 32 + match c with + | CN => 0 + | CZ => 1 + | CC => 2 + | CV => 3 + end. + +Definition preg_index (p: preg): Z := + match p with + | IR i => ireg_index i + | FR f => freg_index f + | CR c => crbit_index c + | PC => 16 + 32 + 4 + 1 + end. + +Module Preg <: REGISTER_MODEL. + + Definition reg := preg. + Definition eq_dec := preg_eq. + + Definition type pr := + match pr with + | IR _ | CR _ | PC => Tany32 + | FR _ => Tany64 + end. + + Definition quantity_of pr := + match pr with + | IR _ | CR _ | PC => Q32 + | FR _ => Q64 + end. + + Definition chunk_of pr := + match pr with + | IR _ | CR _ | PC => Many32 + | FR _ => Many64 + end. + + Lemma type_cases: forall r, type r = Tany32 \/ type r = Tany64. + Proof. + destruct r; auto. + Qed. + + Definition ofs (r: preg): Z := + preg_index r. + + (* A register's address: The index of its first byte. *) + Definition addr (r: preg): Z := + preg_index r * 4. + + (* The address one byte past the end of register [r]. The next nonoverlapping + register may start here. *) + Definition next_addr (r: preg): Z := addr r + AST.typesize (type r). + + Remark addr_pos: forall p, addr p > 0. + Proof. + intros. unfold addr. destruct p as [x|x|x|]; try destruct x; simpl; omega. + Qed. + + Lemma addr_compat: forall p, FragBlock.addr (ofs p) = addr p. + Proof. + reflexivity. + Qed. + + Lemma size_compat: + forall p, + AST.typesize (type p) = FragBlock.quantity_size (quantity_of p). + Proof. + intros. unfold quantity_of. destruct (type p) eqn:TYP, p; simpl; auto; inv TYP. + Qed. + + Lemma next_addr_compat: forall p, FragBlock.next_addr (ofs p) (quantity_of p) = next_addr p. + Proof. + unfold next_addr, addr, ofs, FragBlock.next_addr, FragBlock.addr; intros. + rewrite size_compat. auto. + Qed. + + Lemma quantity_of_compat: + forall p, + quantity_of p = quantity_of_typ (type p). + Proof. + destruct p; reflexivity. + Qed. + + Lemma chunk_of_reg_compat: + forall p, + chunk_of p = chunk_of_quantity (quantity_of p). + Proof. + destruct p; simpl; auto. + Qed. + + Lemma chunk_of_reg_type: + forall p, + chunk_of p = chunk_of_type (type p). + Proof. + destruct p; reflexivity. + Qed. + + Lemma diff_outside_interval: + forall r s, r <> s -> next_addr r <= addr s \/ next_addr s <= addr r. + Proof. + intros. + destruct s as [x|x|x|] eqn:S; try destruct x; simpl in H; + destruct r as [y|y|y|] eqn:R; try destruct y; simpl in H; compute; + try (left; congruence); try (right; congruence); + compute in H; tauto. + Qed. + +End Preg. + +Lemma pc_type: subtype Tptr (Preg.type PC) = true. +Proof. + simpl; auto. +Qed. -Module Pregmap := EMap(PregEq). +Module Pregmap := Regfile(Preg). (** Conventional names for stack pointer ([SP]) and return address ([RA]) *) @@ -210,9 +337,11 @@ Inductive instruction : Type := | Pfldd: freg -> ireg -> int -> instruction (**r float64 load *) | Pfldd_a: freg -> ireg -> int -> instruction (**r any64 load to FP reg *) | Pflds: freg -> ireg -> int -> instruction (**r float32 load *) + | Pflds_a: freg -> ireg -> int -> instruction (**r any32 load to FP reg *) | Pfstd: freg -> ireg -> int -> instruction (**r float64 store *) | Pfstd_a: freg -> ireg -> int -> instruction (**r any64 store from FP reg *) | Pfsts: freg -> ireg -> int -> instruction (**r float32 store *) + | Pfsts_a: freg -> ireg -> int -> instruction (**r any32 store from FP reg *) (* Pseudo-instructions *) | Pallocframe: Z -> ptrofs -> instruction (**r allocate new stack frame *) @@ -323,12 +452,14 @@ Definition program := AST.program fundef unit. type [Tint], float registers to values of type [Tfloat], and condition bits to either [Vzero] or [Vone]. *) -Definition regset := Pregmap.t val. +Definition regset := Pregmap.t. Definition genv := Genv.t fundef unit. -Notation "a # b" := (a b) (at level 1, only parsing) : asm. +Notation "a # b" := (Pregmap.get b a) (at level 1, only parsing) : asm. Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm. +Definition pregmap_read rs := fun r => Pregmap.get r rs. + Open Scope asm. (** Undefining some registers *) @@ -336,13 +467,33 @@ Open Scope asm. Fixpoint undef_regs (l: list preg) (rs: regset) : regset := match l with | nil => rs - | r :: l' => undef_regs l' (rs#r <- Vundef) + | r :: l' => (undef_regs l' rs)#r <- Vundef end. +Lemma undef_regs_other: + forall r rl rs, ~In r rl -> (undef_regs rl rs) # r = rs # r. +Proof. + induction rl; simpl; intros. auto. rewrite Pregmap.gso. apply IHrl. intuition. intuition. +Qed. + +Lemma undef_regs_same: + forall r rl rs, In r rl -> (undef_regs rl rs) # r = Vundef. +Proof. + induction rl; simpl; intros. tauto. + destruct H. + - subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + - destruct (Preg.eq_dec r a). + + subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + + rewrite Pregmap.gso; auto. +Qed. + (** Undefining the condition codes *) Definition undef_flags (rs: regset) : regset := - fun r => match r with CR _ => Vundef | _ => rs r end. + Pregmap.set (CR CN) Vundef + (Pregmap.set (CR CZ) Vundef + (Pregmap.set (CR CC) Vundef + (Pregmap.set (CR CV) Vundef rs))). (** Assigning a register pair *) @@ -354,11 +505,11 @@ Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := (** Assigning the result of a builtin *) -Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := +Definition set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := match res with | BR r => rs#r <- v | BR_none => rs - | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) + | BR_splitlong hi lo => rs #hi <- (Val.hiword v) #lo <- (Val.loword v) end. Section RELSEM. @@ -430,11 +581,11 @@ Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) := Definition eval_shift_op (so: shift_op) (rs: regset) := match so with | SOimm n => Vint n - | SOreg r => rs r - | SOlsl r n => Val.shl (rs r) (Vint n) - | SOlsr r n => Val.shru (rs r) (Vint n) - | SOasr r n => Val.shr (rs r) (Vint n) - | SOror r n => Val.ror (rs r) (Vint n) + | SOreg r => rs # r + | SOlsl r n => Val.shl (rs # r) (Vint n) + | SOlsr r n => Val.shru (rs # r) (Vint n) + | SOasr r n => Val.shr (rs # r) (Vint n) + | SOror r n => Val.ror (rs # r) (Vint n) end. (** Auxiliaries for memory accesses *) @@ -448,7 +599,7 @@ Definition exec_load (chunk: memory_chunk) (addr: val) (r: preg) Definition exec_store (chunk: memory_chunk) (addr: val) (r: preg) (rs: regset) (m: mem) := - match Mem.storev chunk m addr (rs r) with + match Mem.storev chunk m addr (rs # r) with | None => Stuck | Some m' => Next (nextinstr rs) m' end. @@ -503,62 +654,62 @@ Definition compare_float32 (rs: regset) (v1 v2: val) := Definition eval_testcond (c: testcond) (rs: regset) : option bool := match c with | TCeq => (**r equal *) - match rs CZ with + match rs # CZ with | Vint n => Some (Int.eq n Int.one) | _ => None end | TCne => (**r not equal *) - match rs CZ with + match rs # CZ with | Vint n => Some (Int.eq n Int.zero) | _ => None end | TClo => (**r unsigned less than *) - match rs CC with + match rs # CC with | Vint n => Some (Int.eq n Int.zero) | _ => None end | TCls => (**r unsigned less or equal *) - match rs CC, rs CZ with + match rs # CC, rs # CZ with | Vint c, Vint z => Some (Int.eq c Int.zero || Int.eq z Int.one) | _, _ => None end | TChs => (**r unsigned greater or equal *) - match rs CC with + match rs # CC with | Vint n => Some (Int.eq n Int.one) | _ => None end | TChi => (**r unsigned greater *) - match rs CC, rs CZ with + match rs # CC, rs # CZ with | Vint c, Vint z => Some (Int.eq c Int.one && Int.eq z Int.zero) | _, _ => None end | TClt => (**r signed less than *) - match rs CV, rs CN with + match rs # CV, rs # CN with | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.one) | _, _ => None end | TCle => (**r signed less or equal *) - match rs CV, rs CN, rs CZ with + match rs # CV, rs # CN, rs # CZ with | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.one || Int.eq z Int.one) | _, _, _ => None end | TCge => (**r signed greater or equal *) - match rs CV, rs CN with + match rs # CV, rs # CN with | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.zero) | _, _ => None end | TCgt => (**r signed greater *) - match rs CV, rs CN, rs CZ with + match rs # CV, rs # CN, rs # CZ with | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.zero && Int.eq z Int.zero) | _, _, _ => None end | TCpl => (**r positive *) - match rs CN with + match rs # CN with | Vint n => Some (Int.eq n Int.zero) | _ => None end | TCmi => (**r negative *) - match rs CN with + match rs # CN with | Vint n => Some (Int.eq n Int.one) | _ => None end @@ -743,22 +894,26 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out exec_load Many64 (Val.add rs#r2 (Vint n)) r1 rs m | Pflds r1 r2 n => exec_load Mfloat32 (Val.add rs#r2 (Vint n)) r1 rs m + | Pflds_a r1 r2 n => + exec_load Many32 (Val.add rs#r2 (Vint n)) r1 rs m | Pfstd r1 r2 n => exec_store Mfloat64 (Val.add rs#r2 (Vint n)) r1 rs m | Pfstd_a r1 r2 n => exec_store Many64 (Val.add rs#r2 (Vint n)) r1 rs m | Pfsts r1 r2 n => exec_store Mfloat32 (Val.add rs#r2 (Vint n)) r1 rs m + | Pfsts_a r1 r2 n => + exec_store Many32 (Val.add rs#r2 (Vint n)) r1 rs m (* Pseudo-instructions *) | Pallocframe sz pos => let (m1, stk) := Mem.alloc m 0 sz in let sp := (Vptr stk Ptrofs.zero) in - match Mem.storev Mint32 m1 (Val.offset_ptr sp pos) rs#IR13 with + match Mem.storev Many32 m1 (Val.offset_ptr sp pos) rs#IR13 with | None => Stuck | Some m2 => Next (nextinstr (rs #IR12 <- (rs#IR13) #IR13 <- sp)) m2 end | Pfreeframe sz pos => - match Mem.loadv Mint32 m (Val.offset_ptr rs#IR13 pos) with + match Mem.loadv Many32 m (Val.offset_ptr rs#IR13 pos) with | None => Stuck | Some v => match rs#IR13 with @@ -854,25 +1009,54 @@ Definition preg_of (r: mreg) : preg := (** Undefine all registers except SP and callee-save registers *) -Definition undef_caller_save_regs (rs: regset) : regset := +Definition undef_caller_save_regs_spec (rs: preg -> val) : preg -> val := fun r => if preg_eq r SP || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) then rs r else Vundef. +Definition pregs_destroyed_at_call := + PC :: CR CN :: CR CZ :: CR CC :: CR CV :: IR RA + :: (map preg_of (filter (fun m => negb (is_callee_save m)) all_mregs)). + +Lemma pregs_destroyed_at_call_correct: + forall r, + preg_eq r SP + || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) + = + negb (In_dec preg_eq r pregs_destroyed_at_call). +Proof. + intros. + destruct r as [x|x|x|]; try destruct x; auto. +Qed. + +Definition undef_caller_save_regs (rs: regset) : regset := + undef_regs pregs_destroyed_at_call rs. + +Lemma undef_caller_save_regs_correct: + forall rs r, + (undef_caller_save_regs rs) # r = undef_caller_save_regs_spec (fun r' => rs # r) r. +Proof. + intros. unfold undef_caller_save_regs, undef_caller_save_regs_spec. + rewrite pregs_destroyed_at_call_correct. + destruct (In_dec preg_eq r pregs_destroyed_at_call) as [IN | NOTIN]. + - rewrite undef_regs_same; auto. + - rewrite undef_regs_other; auto. +Qed. + (** Extract the values of the arguments of an external call. We exploit the calling conventions from module [Conventions], except that we use ARM registers instead of locations. *) Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop := | extcall_arg_reg: forall r, - extcall_arg rs m (R r) (rs (preg_of r)) - | extcall_arg_stack: forall ofs ty bofs v, + extcall_arg rs m (R r) (rs # (preg_of r)) + | extcall_arg_stack: forall ofs q bofs v, bofs = Stacklayout.fe_ofs_arg + 4 * ofs -> - Mem.loadv (chunk_of_type ty) m - (Val.offset_ptr (rs (IR IR13)) (Ptrofs.repr bofs)) = Some v -> - extcall_arg rs m (S Outgoing ofs ty) v. + Mem.loadv (chunk_of_quantity q) m + (Val.offset_ptr (rs # (IR IR13)) (Ptrofs.repr bofs)) = Some v -> + extcall_arg rs m (S Outgoing ofs q) v. Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop := | extcall_arg_one: forall l v, @@ -898,17 +1082,17 @@ Inductive state: Type := Inductive step: state -> trace -> state -> Prop := | exec_step_internal: forall b ofs f i rs m rs' m', - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) (fn_code f) = Some i -> exec_instr f i rs m = Next rs' m' -> step (State rs m) E0 (State rs' m') | exec_step_builtin: forall b ofs f ef args res rs m vargs t vres rs' m', - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) -> - eval_builtin_args ge rs (rs SP) m args vargs -> + eval_builtin_args ge (pregmap_read rs) (rs # SP) m args vargs -> external_call ef ge vargs m t vres m' -> rs' = nextinstr (set_res res vres @@ -916,11 +1100,11 @@ Inductive step: state -> trace -> state -> Prop := step (State rs m) t (State rs' m') | exec_step_external: forall b ef args res rs m t rs' m', - rs PC = Vptr b Ptrofs.zero -> + rs # PC = Vptr b Ptrofs.zero -> Genv.find_funct_ptr ge b = Some (External ef) -> external_call ef ge args m t res m' -> extcall_arguments rs m (ef_sig ef) args -> - rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs IR14) -> + rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs#IR14) -> step (State rs m) t (State rs' m'). End RELSEM. @@ -931,7 +1115,7 @@ Inductive initial_state (p: program): state -> Prop := | initial_state_intro: forall m0, let ge := Genv.globalenv p in let rs0 := - (Pregmap.init Vundef) + Pregmap.init # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero) # IR14 <- Vzero # IR13 <- Vzero in diff --git a/arm/AsmToJSON.ml b/arm/AsmToJSON.ml index 276ceecc29..3ee0b9f816 100644 --- a/arm/AsmToJSON.ml +++ b/arm/AsmToJSON.ml @@ -225,7 +225,7 @@ let pp_instructions pp ic = | Pfdivd(r1, r2, r3) -> instruction pp "Pfdivd" [DFreg r1; DFreg r2; DFreg r3] | Pfdivs(r1, r2, r3) -> instruction pp "Pfdivs" [SFreg r1; SFreg r2; SFreg r3] | Pfldd(r1, r2, n) | Pfldd_a(r1, r2, n) -> instruction pp "Pfldd" [DFreg r1; Ireg r2; Long n] - | Pflds(r1, r2, n) -> instruction pp "Pflds" [SFreg r1; Ireg r2; Long n] + | Pflds(r1, r2, n) | Pflds_a(r1, r2, n) -> instruction pp "Pflds" [SFreg r1; Ireg r2; Long n] | Pflid(r1, f) -> instruction pp "Pflid" [DFreg r1; Float64 f] | Pfmuld(r1, r2, r3) -> instruction pp "Pfmuld" [DFreg r1; DFreg r2; DFreg r3] | Pfmuls(r1, r2, r3) -> instruction pp "Pfmuls" [SFreg r1; SFreg r2; SFreg r3] @@ -235,7 +235,7 @@ let pp_instructions pp ic = | Pfsitos(r1, r2) -> instruction pp "Pfsitos" [SFreg r1; Ireg r2] | Pfsqrt(r1, r2) -> instruction pp "Pfsqrt" [DFreg r1; DFreg r2] | Pfstd(r1, r2, n) | Pfstd_a(r1, r2, n) -> instruction pp "Pfstd" [DFreg r1; Ireg r2; Long n] - | Pfsts(r1, r2, n) -> instruction pp "Pfsts" [SFreg r1; Ireg r2; Long n] + | Pfsts(r1, r2, n) | Pfsts_a(r1, r2, n) -> instruction pp "Pfsts" [SFreg r1; Ireg r2; Long n] | Pfsubd(r1, r2, r3) -> instruction pp "Pfsubd" [DFreg r1; DFreg r2; DFreg r3] | Pfsubs(r1, r2, r3) -> instruction pp "Pfsubs" [SFreg r1; SFreg r2; SFreg r3] | Pftosizd(r1, r2) -> instruction pp "Pftosizd" [Ireg r1; DFreg r2] diff --git a/arm/Asmexpand.ml b/arm/Asmexpand.ml index 7c18be6ba6..23e7c9cc77 100644 --- a/arm/Asmexpand.ml +++ b/arm/Asmexpand.ml @@ -187,7 +187,7 @@ let expand_builtin_vload_common chunk base ofs res = emit (Pldrsh (res, base, SOimm ofs)) | Mint32, BR(IR res) -> emit (Pldr (res, base, SOimm ofs)) - | Mint64, BR_splitlong(BR(IR res1), BR(IR res2)) -> + | Mint64, BR_splitlong(IR res1, IR res2) -> let ofs_hi = if Archi.big_endian then ofs else Int.add ofs _4 in let ofs_lo = if Archi.big_endian then Int.add ofs _4 else ofs in if base <> res2 then begin @@ -350,7 +350,7 @@ let expand_builtin_inline name args res = emit (Pfsqrt (res,a1)) (* 64-bit integer arithmetic *) | "__builtin_negl", [BA_splitlong(BA(IR ah), BA(IR al))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah ) rl (fun rl -> emit (Prsbs (rl,al,SOimm _0)); (* No "rsc" instruction in Thumb2. Emulate based on @@ -364,20 +364,20 @@ let expand_builtin_inline name args res = end) | "__builtin_addl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah || rl = bh) rl (fun rl -> emit (Padds (rl,al,SOreg bl)); emit (Padc (rh,ah,SOreg bh))) | "__builtin_subl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah || rl = bh) rl (fun rl -> emit (Psubs (rl,al,SOreg bl)); emit (Psbc (rh,ah,SOreg bh))) | "__builtin_mull", [BA(IR a); BA(IR b)], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> emit (Pumull (rl,rh,a,b)) (* Memory accesses *) | "__builtin_read16_reversed", [BA(IR a1)], BR(IR res) -> diff --git a/arm/Asmgen.v b/arm/Asmgen.v index 1d2f360fdb..d7b9c15758 100644 --- a/arm/Asmgen.v +++ b/arm/Asmgen.v @@ -570,37 +570,29 @@ Definition indexed_memory_access else addimm IR14 base (Int.sub n n1) (mk_instr IR14 n1 :: k). Definition loadind_int (base: ireg) (ofs: ptrofs) (dst: ireg) (k: code) := - indexed_memory_access (fun base n => Pldr dst base (SOimm n)) mk_immed_mem_word base (Ptrofs.to_int ofs) k. + indexed_memory_access (fun base n => Pldr_a dst base (SOimm n)) mk_immed_mem_word base (Ptrofs.to_int ofs) k. -Definition loadind (base: ireg) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) := +Definition loadind (base: ireg) (ofs: ptrofs) (q: quantity) (dst: mreg) (k: code) := let ofs := Ptrofs.to_int ofs in - match ty, preg_of dst with - | Tint, IR r => - OK (indexed_memory_access (fun base n => Pldr r base (SOimm n)) mk_immed_mem_word base ofs k) - | Tany32, IR r => + match q, preg_of dst with + | Q32, IR r => OK (indexed_memory_access (fun base n => Pldr_a r base (SOimm n)) mk_immed_mem_word base ofs k) - | Tsingle, FR r => - OK (indexed_memory_access (Pflds r) mk_immed_mem_float base ofs k) - | Tfloat, FR r => - OK (indexed_memory_access (Pfldd r) mk_immed_mem_float base ofs k) - | Tany64, FR r => + | Q32, FR r => + OK (indexed_memory_access (Pflds_a r) mk_immed_mem_float base ofs k) + | Q64, FR r => OK (indexed_memory_access (Pfldd_a r) mk_immed_mem_float base ofs k) | _, _ => Error (msg "Asmgen.loadind") end. -Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) := +Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (q: quantity) (k: code) := let ofs := Ptrofs.to_int ofs in - match ty, preg_of src with - | Tint, IR r => - OK (indexed_memory_access (fun base n => Pstr r base (SOimm n)) mk_immed_mem_word base ofs k) - | Tany32, IR r => + match q, preg_of src with + | Q32, IR r => OK (indexed_memory_access (fun base n => Pstr_a r base (SOimm n)) mk_immed_mem_word base ofs k) - | Tsingle, FR r => - OK (indexed_memory_access (Pfsts r) mk_immed_mem_float base ofs k) - | Tfloat, FR r => - OK (indexed_memory_access (Pfstd r) mk_immed_mem_float base ofs k) - | Tany64, FR r => + | Q32, FR r => + OK (indexed_memory_access (Pfsts_a r) mk_immed_mem_float base ofs k) + | Q64, FR r => OK (indexed_memory_access (Pfstd_a r) mk_immed_mem_float base ofs k) | _, _ => Error (msg "Asmgen.storeind") @@ -614,8 +606,8 @@ Definition save_lr (ofs: ptrofs) (k: code) := let n := Ptrofs.to_int ofs in let n1 := mk_immed_mem_word n in if Int.eq n n1 - then Pstr IR14 IR13 (SOimm n) :: k - else addimm IR12 IR13 (Int.sub n n1) (Pstr IR14 IR12 (SOimm n1) :: k). + then Pstr_a IR14 IR13 (SOimm n) :: k + else addimm IR12 IR13 (Int.sub n n1) (Pstr_a IR14 IR12 (SOimm n1) :: k). Definition save_lr_preserves_R12 (ofs: ptrofs) : bool := let n := Ptrofs.to_int ofs in @@ -722,12 +714,12 @@ Definition transl_store (chunk: memory_chunk) (addr: addressing) Definition transl_instr (f: Mach.function) (i: Mach.instruction) (r12_is_parent: bool) (k: code) := match i with - | Mgetstack ofs ty dst => - loadind IR13 ofs ty dst k - | Msetstack src ofs ty => - storeind src IR13 ofs ty k - | Mgetparam ofs ty dst => - do c <- loadind IR12 ofs ty dst k; + | Mgetstack ofs q dst => + loadind IR13 ofs q dst k + | Msetstack src ofs q => + storeind src IR13 ofs q k + | Mgetparam ofs q dst => + do c <- loadind IR12 ofs q dst k; OK (if r12_is_parent then c else loadind_int IR13 f.(fn_link_ofs) IR12 c) @@ -769,8 +761,8 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction) Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool := match i with - | Msetstack src ofs ty => before - | Mgetparam ofs ty dst => negb (mreg_eq dst R12) + | Msetstack src ofs q => before + | Mgetparam ofs q dst => negb (mreg_eq dst R12) | Mop Omove args res => before && negb (mreg_eq res R12) | _ => false end. diff --git a/arm/Asmgenproof.v b/arm/Asmgenproof.v index 2c001f45e2..babcf67747 100644 --- a/arm/Asmgenproof.v +++ b/arm/Asmgenproof.v @@ -73,7 +73,7 @@ Qed. Lemma exec_straight_exec: forall fb f c ep tf tc c' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs # PC) fb f c ep tf tc -> exec_straight tge tf tc rs m c' rs' m' -> plus step tge (State rs m) E0 (State rs' m'). Proof. @@ -85,10 +85,10 @@ Qed. Lemma exec_straight_at: forall fb f c ep tf tc c' ep' tc' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs # PC) fb f c ep tf tc -> transl_code f c' ep' = OK tc' -> exec_straight tge tf tc rs m tc' rs' m' -> - transl_code_at_pc ge (rs' PC) fb f c' ep' tf tc'. + transl_code_at_pc ge (rs' # PC) fb f c' ep' tf tc'. Proof. intros. inv H. exploit exec_straight_steps_2; eauto. @@ -362,11 +362,11 @@ Lemma find_label_goto_label: forall f tf lbl rs m c' b ofs, Genv.find_funct_ptr ge b = Some (Internal f) -> transf_function f = OK tf -> - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> exists tc', exists rs', goto_label tf lbl rs m = Next rs' m - /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc' + /\ transl_code_at_pc ge (rs' # PC) b f c' false tf tc' /\ forall r, r <> PC -> rs'#r = rs#r. Proof. intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2. @@ -379,6 +379,7 @@ Proof. rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q. auto. omega. generalize (transf_function_no_overflow _ _ H0). omega. + simpl; auto. intros. apply Pregmap.gso; auto. Qed. @@ -422,7 +423,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (FIND: Genv.find_funct_ptr ge fb = Some (Internal f)) (MEXT: Mem.extends m m') - (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc) + (AT: transl_code_at_pc ge (rs # PC) fb f c ep tf tc) (AG: agree ms sp rs) (DXP: ep = true -> rs#IR12 = parent_sp s), match_states (Mach.State s fb sp c ms m) @@ -432,8 +433,8 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = Vptr fb Ptrofs.zero) - (ATLR: rs RA = parent_ra s), + (ATPC: rs # PC = Vptr fb Ptrofs.zero) + (ATLR: rs # RA = parent_ra s), match_states (Mach.Callstate s fb ms m) (Asm.State rs m') | match_states_return: @@ -441,7 +442,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = parent_ra s), + (ATPC: rs # PC = parent_ra s), match_states (Mach.Returnstate s ms m) (Asm.State rs m'). @@ -450,7 +451,7 @@ Lemma exec_straight_steps: match_stack ge s -> Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1 # PC) fb f (i :: c) ep tf tc -> (forall k c (TR: transl_instr f i ep k = OK c), exists rs2, exec_straight tge tf c rs1 m1' k rs2 m2' @@ -473,7 +474,7 @@ Lemma exec_straight_steps_goto: Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1 # PC) fb f (i :: c) ep tf tc -> it1_is_parent ep i = false -> (forall k c (TR: transl_instr f i ep k = OK c), exists jmp, exists k', exists rs2, @@ -534,11 +535,13 @@ Theorem step_simulation: (exists S2', plus step tge S1' t S2' /\ match_states S2 S2') \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat. Proof. - induction 1; intros; inv MS. + intros. set (IND := Mach.step return_address_offset ge S1 t S2). + induction H; intros; inv MS. - (* Mlabel *) left; eapply exec_straight_steps; eauto; intros. - monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto. + monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. + auto using nextinstr_pc. split. apply agree_nextinstr; auto. simpl; congruence. - (* Mgetstack *) @@ -553,13 +556,15 @@ Proof. - (* Msetstack *) unfold store_stack in H. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (regmap_get src rs) (rs0 # (preg_of src))). eapply preg_val; eauto. 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]]. exists rs'; split. eauto. - split. eapply agree_undef_regs; eauto with asmgen. +Local Transparent destroyed_by_setstack. + split. + simpl; eapply agree_exten; eauto with asmgen. simpl; intros. rewrite Q; auto with asmgen. - (* Mgetparam *) @@ -588,7 +593,13 @@ Opaque loadind. split. eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto. instantiate (1 := rs1#IR12 <- (rs2#IR12)). intros. rewrite Pregmap.gso; auto with asmgen. - congruence. intros. unfold Pregmap.set. destruct (PregEq.eq r' IR12). congruence. auto with asmgen. + congruence. intros. + destruct (Preg.eq_dec r' IR12). + subst; rewrite Pregmap.gss. + change (Preg.chunk_of IR12) with (chunk_of_type Tany32). + rewrite Val.load_result_same; auto. + apply Pregmap.get_has_type. + rewrite Pregmap.gso; auto with asmgen. simpl; intros. rewrite U; auto with asmgen. apply preg_of_not_R12; auto. @@ -599,7 +610,7 @@ Opaque loadind. 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]]]. - assert (S: Val.lessdef v (rs2 (preg_of res))) by (eapply Val.lessdef_trans; eauto). + assert (S: Val.lessdef v (rs2 # (preg_of res))) by (eapply Val.lessdef_trans; eauto). exists rs2; split. eauto. split. eapply agree_set_undef_mreg; eauto with asmgen. simpl. destruct op; try congruence. destruct ep; simpl; try congruence. intros. @@ -622,7 +633,7 @@ Opaque loadind. rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1. intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (Regmap.get src rs) (rs0 # (preg_of src))). eapply preg_val; eauto. exploit Mem.storev_extends; eauto. intros [m2' [C D]]. left; eapply exec_straight_steps; eauto. intros. simpl in TR. @@ -638,10 +649,10 @@ Opaque loadind. eapply transf_function_no_overflow; eauto. destruct ros as [rf|fid]; simpl in H; monadInv H5. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (Pregmap.get x0 rs0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto. generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x). @@ -653,8 +664,10 @@ Opaque loadind. simpl. eauto. econstructor; eauto. econstructor; eauto. - eapply agree_sp_def; eauto. + eapply agree_sp_def; eauto. eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simpl. + rewrite Pregmap.gss. change (Preg.chunk_of PC) with (chunk_of_type Tany32). + rewrite Val.load_result_same; auto. apply Pregmap.get_has_type. Simpl. rewrite <- H2. auto. + (* Direct call *) generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. @@ -667,9 +680,9 @@ Opaque loadind. simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. eauto. econstructor; eauto. econstructor; eauto. - eapply agree_sp_def; eauto. + eapply agree_sp_def; eauto. eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simpl. - Simpl. rewrite <- H2. auto. + Simpl. rewrite <- H2. Simpl. - (* Mtailcall *) assert (f0 = f) by congruence. subst f0. @@ -697,15 +710,18 @@ Opaque loadind. econstructor; split. eapply exec_straight_trans. eexact P. apply exec_straight_one. simpl. rewrite R; auto with asmgen. unfold chunk_of_type in A; simpl in A. rewrite A. - rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. auto. - split. Simpl. split. Simpl. intros. Simpl. + rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. auto using if_preg_nextinstr_PC. + split. Simpl. + change (Preg.chunk_of IR13) with (chunk_of_type Tany32). + apply Val.load_result_same. simpl in A. apply Mem.loadv_type in A. exact A. + split. Simpl. intros. Simpl. } destruct ros as [rf|fid]; simpl in H; monadInv H7. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0 # x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H7; intros LD; inv LD; auto. destruct (X (Pbreg x0 sig :: x)) as [rs2 [P [Q [R S]]]]. exploit exec_straight_steps_2. eexact P. eauto. eauto. eapply functions_transl; eauto. eauto. @@ -717,11 +733,15 @@ Opaque loadind. simpl. reflexivity. traceEq. econstructor; eauto. - split. Simpl. eapply parent_sp_def; eauto. - intros. Simpl. rewrite S; auto with asmgen. eapply preg_val; eauto. - Simpl. rewrite S; auto with asmgen. + split. Simpl. eapply parent_sp_def; eauto. eapply parent_sp_type; eauto. + intros. rewrite Pregmap.gso; eauto with asmgen. + rewrite S; auto with asmgen. eapply preg_val; eauto. + Simpl. change (Preg.chunk_of PC) with (chunk_of_type Tany32). + rewrite S, H9. auto using Val.load_result_same. + rewrite <- (ireg_of_eq _ _ EQ1); auto with asmgen. rewrite <- (ireg_of_eq _ _ EQ1); auto with asmgen. rewrite <- (ireg_of_eq _ _ EQ1); auto with asmgen. + Simpl. + (* Direct call *) destruct (X (Pbsymb fid sig :: x)) as [rs2 [P [Q [R S]]]]. exploit exec_straight_steps_2. eexact P. eauto. eauto. eapply functions_transl; eauto. eauto. @@ -733,13 +753,15 @@ Opaque loadind. simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. reflexivity. traceEq. econstructor; eauto. - split. Simpl. eapply parent_sp_def; eauto. + split. Simpl. eapply parent_sp_def; eauto. eapply parent_sp_type; eauto. intros. Simpl. rewrite S; auto with asmgen. eapply preg_val; eauto. + Simpl. + Simpl. - (* Mbuiltin *) - inv AT. monadInv H4. + inv AT. monadInv H5. exploit functions_transl; eauto. intro FN. - generalize (transf_function_no_overflow _ _ H3); intro NOOV. + generalize (transf_function_no_overflow _ _ H4); intro NOOV. exploit builtin_args_match; eauto. intros [vargs' [P Q]]. exploit external_call_mem_extends; eauto. intros [vres' [m2' [A [B [C D]]]]]. @@ -754,12 +776,19 @@ Opaque loadind. instantiate (2 := tf); instantiate (1 := x). unfold nextinstr. rewrite Pregmap.gss. rewrite set_res_other. rewrite undef_regs_other_2. - rewrite <- H1. simpl. econstructor; eauto. + rewrite <- H2. simpl. econstructor; eauto. eapply code_tail_next_int; eauto. rewrite preg_notin_charact. intros. auto with asmgen. auto with asmgen. apply agree_nextinstr. eapply agree_set_res; auto. eapply agree_undef_regs; eauto. intros; apply undef_regs_other_2; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + destruct res; simpl in H1; InvBooleans; simpl; auto. + eapply Val.has_subtype; eauto. + split. + destruct vres'; simpl; auto; destruct (mreg_type lo); auto; congruence. + destruct vres'; simpl; auto; destruct (mreg_type hi); auto; congruence. congruence. - (* Mgoto *) @@ -781,6 +810,7 @@ Opaque loadind. left; eapply exec_straight_steps_goto; eauto. intros. simpl in TR. destruct (transl_cond_correct tge tf cond args _ rs0 m' _ TR) as [rs' [A [B C]]]. + unfold pregmap_read in EC. rewrite EC in B. destruct B as [Bpos Bneg]. econstructor; econstructor; econstructor; split. eexact A. split. eapply agree_undef_regs; eauto with asmgen. @@ -790,10 +820,11 @@ Opaque loadind. exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC. left; eapply exec_straight_steps; eauto. intros. simpl in TR. destruct (transl_cond_correct tge tf cond args _ rs0 m' _ TR) as [rs' [A [B C]]]. + unfold pregmap_read in EC. rewrite EC in B. destruct B as [Bpos Bneg]. econstructor; split. eapply exec_straight_trans. eexact A. - apply exec_straight_one. simpl. rewrite Bpos. reflexivity. auto. + apply exec_straight_one. simpl. rewrite Bpos. reflexivity. apply nextinstr_pc. split. eapply agree_undef_regs; eauto with asmgen. intros; Simpl. simpl. congruence. @@ -843,8 +874,10 @@ Opaque loadind. econstructor; split. eapply exec_straight_trans. eexact P. apply exec_straight_one. simpl. rewrite R; auto with asmgen. rewrite A. - rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. auto. + rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. apply if_preg_nextinstr_PC; auto. split. Simpl. + assert (SP_TYP: Val.has_type (parent_sp s) Tany32) by (apply Mem.loadv_type in A; auto). + change (Preg.chunk_of IR13) with (chunk_of_type Tany32). apply Val.load_result_same; auto. split. Simpl. intros. Simpl. } @@ -858,8 +891,11 @@ Opaque loadind. simpl. reflexivity. traceEq. econstructor; eauto. - split. Simpl. eapply parent_sp_def; eauto. + split. Simpl. eapply parent_sp_def; eauto. eapply parent_sp_type; eauto. intros. Simpl. rewrite S; auto with asmgen. eapply preg_val; eauto. + Simpl. + change (Preg.chunk_of PC) with (chunk_of_type Tany32). rewrite Val.load_result_same; auto. + apply Pregmap.get_has_type. - (* internal function *) exploit functions_translated; eauto. intros [tf [A B]]. monadInv B. @@ -881,7 +917,10 @@ Opaque loadind. (* Execution of function prologue *) set (rs2 := nextinstr (rs0#IR12 <- (parent_sp s) #IR13 <- (Vptr stk Ptrofs.zero))). edestruct (save_lr_correct tge tf ra_ofs (Pcfi_rel_offset ra_ofs' :: x0) rs2) as (rs3 & X & Y & Z). - change (rs2 IR13) with sp. change (rs2 IR14) with (rs0 IR14). rewrite ATLR. eexact P. + replace (rs2 # IR13) with sp by (subst rs2; Simpl). + replace (rs2 # IR14) with (rs0 # IR14) by (subst rs2; Simpl). + rewrite ATLR. eexact P. + set (rs4 := nextinstr rs3). assert (EXEC_PROLOGUE: exec_straight tge tf @@ -892,12 +931,17 @@ Opaque loadind. eapply exec_straight_trans with (rs2 := rs2) (m2 := m2'). apply exec_straight_one. unfold exec_instr. rewrite C. fold sp. - rewrite <- (sp_val _ _ _ AG). unfold Tptr, chunk_of_type, Archi.ptr64 in F. rewrite F. auto. - auto. + rewrite <- (sp_val _ _ _ AG). + unfold Tptr, chunk_of_quantity, Archi.ptr64, quantity_of_typ in F. + rewrite F. auto. + subst rs2. unfold nextinstr. rewrite Pregmap.gss. Simpl. + change (Preg.chunk_of PC) with (chunk_of_type Tany32). apply Val.load_result_same. + destruct (Pregmap.get PC rs0); simpl; auto. eapply exec_straight_trans with (rs2 := rs3) (m2 := m3'). eexact X. apply exec_straight_one. - simpl; reflexivity. reflexivity. + simpl; reflexivity. + subst rs4; Simpl. } (* After the function prologue is the code for the function body *) exploit exec_straight_steps_2; eauto using functions_transl. omega. constructor. @@ -911,9 +955,21 @@ Opaque loadind. apply agree_nextinstr. eapply agree_change_sp. apply agree_undef_regs with rs0; eauto. - intros; Simpl. +Local Transparent destroyed_at_function_entry. + unfold destroyed_at_function_entry; simpl; intros. + rewrite Pregmap.gso; intuition auto. congruence. - intros; apply Y; eauto with asmgen. + simpl; auto. + simpl; auto. + unfold destroyed_at_function_entry; simpl; intros. + apply Y; auto with asmgen. + intro. + unfold rs4, nextinstr. rewrite Pregmap.gso, Z; auto; try discriminate. + unfold rs2, nextinstr. rewrite Pregmap.gso, Pregmap.gso, Pregmap.gss; try discriminate. + change (Preg.chunk_of IR12) with (chunk_of_type Tany32). + apply Val.load_result_same. + apply Val.has_subtype with (ty1 := Tptr); simpl; auto. + eapply parent_sp_type; eauto. - (* external function *) exploit functions_translated; eauto. @@ -928,6 +984,21 @@ Opaque loadind. econstructor; eauto. unfold loc_external_result. apply agree_set_other; auto. apply agree_set_pair; auto. apply agree_undef_caller_save_regs; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + unfold Val.has_type_rpair in H3. + unfold loc_result, proj_sig_res in *. + destruct (sig_res (ef_sig ef)). destruct t0; simpl in *; auto. + destruct Archi.big_endian; destruct H3; simpl in *. + destruct res'; simpl in *; auto. + destruct res'; simpl in *; auto. + simpl in *; auto. + rewrite Pregmap.gss. + change (Preg.chunk_of PC) with (chunk_of_type Tany32). + rewrite ATLR. + apply Val.load_result_same. + apply Val.has_subtype with (ty1 := Tptr); auto. + eapply parent_ra_type; eauto. - (* return *) inv STACKS. simpl in *. @@ -948,7 +1019,10 @@ Proof. econstructor; eauto. constructor. apply Mem.extends_refl. - split. auto. simpl. unfold Vnullptr; simpl; congruence. intros. rewrite Regmap.gi. auto. + split. auto. simpl. unfold Vnullptr; simpl; congruence. simpl; auto. + intros. + replace (regmap_get r Regmap.init) with Vundef by (destruct r; compute; reflexivity). auto. + rewrite Pregmap.gso, Pregmap.gso, Pregmap.gss; auto; congruence. unfold Genv.symbol_address. rewrite (match_program_main TRANSF). rewrite symbols_preserved. diff --git a/arm/Asmgenproof1.v b/arm/Asmgenproof1.v index 98cd5eea03..7e1a1671a1 100644 --- a/arm/Asmgenproof1.v +++ b/arm/Asmgenproof1.v @@ -53,7 +53,11 @@ Hint Resolve ireg_of_not_R14': asmgen. Lemma nextinstr_nf_pc: forall rs, (nextinstr_nf rs)#PC = Val.offset_ptr rs#PC Ptrofs.one. Proof. - intros. reflexivity. + intros. + unfold nextinstr_nf, nextinstr, undef_flags. rewrite Pregmap.gss. + repeat rewrite Pregmap.gso by (compute; intuition congruence). + unfold Val.offset_ptr. + destruct (Pregmap.get PC rs); reflexivity. Qed. Definition if_preg (r: preg) : bool := @@ -74,18 +78,53 @@ Proof. intros; red; intros; subst; discriminate. Qed. -Hint Resolve data_if_preg if_preg_not_PC: asmgen. +Lemma if_preg_not_CR: forall r c, if_preg r = true -> r <> CR c. +Proof. + intros; red; intros; subst; discriminate. +Qed. + +Hint Resolve data_if_preg if_preg_not_PC if_preg_not_CR: asmgen. Lemma nextinstr_nf_inv: forall r rs, if_preg r = true -> (nextinstr_nf rs)#r = rs#r. Proof. - intros. destruct r; reflexivity || discriminate. + intros. + unfold nextinstr_nf, nextinstr, undef_flags. + rewrite !Pregmap.gso; destruct r; reflexivity || discriminate. Qed. Lemma nextinstr_nf_inv1: forall r rs, data_preg r = true -> (nextinstr_nf rs)#r = rs#r. Proof. - intros. destruct r; reflexivity || discriminate. + intros. + unfold nextinstr_nf, nextinstr, undef_flags. + rewrite !Pregmap.gso; destruct r; reflexivity || discriminate. +Qed. + +Lemma if_preg_nextinstr_nf_PC: + forall rs r v, + if_preg r = true -> + Pregmap.get PC (nextinstr_nf rs # r <- v) = + Val.offset_ptr (Pregmap.get PC rs) Ptrofs.one. +Proof. + intros. + unfold nextinstr_nf, nextinstr, undef_flags. rewrite Pregmap.gss. + repeat rewrite Pregmap.gso by (compute; intuition congruence). + rewrite Pregmap.gso. unfold Val.offset_ptr. destruct (Pregmap.get PC rs); reflexivity. + auto using not_eq_sym, if_preg_not_PC. +Qed. + +Lemma if_preg_nextinstr_PC: + forall rs r v, + if_preg r = true -> + Pregmap.get PC (nextinstr rs # r <- v) = + Val.offset_ptr (Pregmap.get PC rs) Ptrofs.one. +Proof. + intros. + unfold nextinstr_nf, nextinstr, undef_flags. rewrite Pregmap.gss. + repeat rewrite Pregmap.gso by (compute; intuition congruence). + rewrite Pregmap.gso. unfold Val.offset_ptr. destruct (Pregmap.get PC rs); reflexivity. + auto using not_eq_sym, if_preg_not_PC. Qed. (** Useful simplification tactic *) @@ -292,6 +331,7 @@ Qed. Lemma iterate_op_correct: forall op1 op2 (f: val -> int -> val) (rs: regset) (r: ireg) m v0 n k, + (forall (v: val) (i: int), Val.has_type (f v i) Tany32) -> (forall (rs:regset) n, exec_instr ge fn (op2 (SOimm n)) rs m = Next (nextinstr_nf (rs#r <- (f (rs#r) n))) m) -> @@ -303,23 +343,30 @@ Lemma iterate_op_correct: /\ rs'#r = List.fold_left f (decompose_int n) v0 /\ forall r': preg, r' <> r -> if_preg r' = true -> rs'#r' = rs#r'. Proof. - intros until k; intros SEM2 SEM1. + intros until k; intros TYP SEM2 SEM1. unfold iterate_op. destruct (decompose_int n) as [ | i tl] eqn:DI. unfold decompose_int in DI. destruct (decompose_int_base n); congruence. revert k. pattern tl. apply List.rev_ind. (* base case *) intros; simpl. econstructor. - split. apply exec_straight_one. rewrite SEM1. reflexivity. reflexivity. + split. apply exec_straight_one. rewrite SEM1. reflexivity. + intuition Simpl. intuition Simpl. + change (Preg.chunk_of r) with (chunk_of_type Tany32). + auto using Val.load_result_same. (* inductive case *) intros. rewrite List.map_app. simpl. rewrite app_ass. simpl. destruct (H (op2 (SOimm x) :: k)) as [rs' [A [B C]]]. econstructor. split. eapply exec_straight_trans. eexact A. apply exec_straight_one. - rewrite SEM2. reflexivity. reflexivity. - split. rewrite fold_left_app; simpl. Simpl. rewrite B. auto. + rewrite SEM2. reflexivity. + intuition Simpl. + split. rewrite fold_left_app; simpl. Simpl. + rewrite B. + change (Preg.chunk_of r) with (chunk_of_type Tany32). + auto using Val.load_result_same. intros; Simpl. Qed. @@ -337,20 +384,25 @@ Proof. set (l2 := length (decompose_int (Int.not n))). destruct (Nat.leb l1 1%nat). { (* single mov *) - econstructor; split. apply exec_straight_one. simpl; reflexivity. auto. + econstructor; split. apply exec_straight_one. simpl. reflexivity. + intuition Simpl. split; intros; Simpl. } destruct (Nat.leb l2 1%nat). { (* single movn *) econstructor; split. apply exec_straight_one. - simpl. rewrite Int.not_involutive. reflexivity. auto. + simpl. rewrite Int.not_involutive. reflexivity. + intuition Simpl. split; intros; Simpl. } destruct Archi.thumb2_support. { (* movw / movt *) unfold loadimm_word. destruct (Int.eq (Int.shru n (Int.repr 16)) Int.zero). econstructor; split. - apply exec_straight_one. simpl; eauto. auto. split; intros; Simpl. + apply exec_straight_one. simpl. reflexivity. + intuition Simpl. + split; intros; Simpl. econstructor; split. - eapply exec_straight_two. simpl; reflexivity. simpl; reflexivity. auto. auto. + eapply exec_straight_two. simpl; reflexivity. simpl. reflexivity. + intuition Simpl. intuition Simpl. split; intros; Simpl. simpl. f_equal. rewrite Int.zero_ext_and by omega. rewrite Int.and_assoc. change 65535 with (two_p 16 - 1). rewrite Int.and_idem. apply Int.same_bits_eq; intros. @@ -365,14 +417,16 @@ Proof. { (* mov - orr* *) replace (Vint n) with (List.fold_left (fun v i => Val.or v (Vint i)) (decompose_int n) Vzero). apply iterate_op_correct. - auto. + destruct v; simpl; auto. + intros; simpl. reflexivity. intros; simpl. rewrite Int.or_commut; rewrite Int.or_zero; auto. rewrite decompose_int_or. simpl. rewrite Int.or_commut; rewrite Int.or_zero; auto. } { (* mvn - bic* *) replace (Vint n) with (List.fold_left (fun v i => Val.and v (Vint (Int.not i))) (decompose_int (Int.not n)) (Vint Int.mone)). apply iterate_op_correct. - auto. + destruct v; simpl; auto. + intros; simpl. reflexivity. intros. simpl. rewrite Int.and_commut; rewrite Int.and_mone; auto. rewrite decompose_int_bic. simpl. rewrite Int.not_involutive. rewrite Int.and_commut. rewrite Int.and_mone; auto. } @@ -390,20 +444,26 @@ Proof. intros. unfold addimm. destruct (Int.ltu (Int.repr (-256)) n). (* sub *) - econstructor; split. apply exec_straight_one; simpl; auto. - split; intros; Simpl. apply Val.sub_opp_add. + econstructor; split. apply exec_straight_one; simpl; auto. Simpl. + split; intros; Simpl. + rewrite Val.sub_opp_add. + change (Preg.chunk_of r1) with (chunk_of_type Tany32). + rewrite Val.load_result_same; auto. + destruct (Pregmap.get r2 rs); simpl; auto. destruct (Nat.leb (length (decompose_int n)) (length (decompose_int (Int.neg n)))). (* add - add* *) - replace (Val.add (rs r2) (Vint n)) - with (List.fold_left (fun v i => Val.add v (Vint i)) (decompose_int n) (rs r2)). + replace (Val.add (rs # r2) (Vint n)) + with (List.fold_left (fun v i => Val.add v (Vint i)) (decompose_int n) (rs # r2)). apply iterate_op_correct. + destruct v; simpl; auto. auto. auto. apply decompose_int_add. (* sub - sub* *) - replace (Val.add (rs r2) (Vint n)) - with (List.fold_left (fun v i => Val.sub v (Vint i)) (decompose_int (Int.neg n)) (rs r2)). + replace (Val.add (rs # r2) (Vint n)) + with (List.fold_left (fun v i => Val.sub v (Vint i)) (decompose_int (Int.neg n)) (rs # r2)). apply iterate_op_correct. + destruct v; simpl; auto. auto. auto. rewrite decompose_int_sub. apply Val.sub_opp_add. @@ -421,12 +481,15 @@ Proof. intros. unfold andimm. destruct (is_immed_arith n). (* andi *) exists (nextinstr_nf (rs#r1 <- (Val.and rs#r2 (Vint n)))). - split. apply exec_straight_one; auto. split; intros; Simpl. + split. apply exec_straight_one; auto. Simpl. split; intros; Simpl. + change (Preg.chunk_of r1) with (chunk_of_type Tany32). + rewrite Val.load_result_same; auto. + destruct (Pregmap.get r2 rs); simpl; auto. (* bic - bic* *) - replace (Val.and (rs r2) (Vint n)) - with (List.fold_left (fun v i => Val.and v (Vint (Int.not i))) (decompose_int (Int.not n)) (rs r2)). + replace (Val.and (rs # r2) (Vint n)) + with (List.fold_left (fun v i => Val.and v (Vint (Int.not i))) (decompose_int (Int.not n)) (rs # r2)). apply iterate_op_correct. - auto. auto. + destruct v; simpl; auto. auto. auto. rewrite decompose_int_bic. rewrite Int.not_involutive. auto. Qed. @@ -441,14 +504,15 @@ Lemma rsubimm_correct: Proof. intros. unfold rsubimm. (* rsb - add* *) - replace (Val.sub (Vint n) (rs r2)) - with (List.fold_left (fun v i => Val.add v (Vint i)) (decompose_int n) (Val.neg (rs r2))). + replace (Val.sub (Vint n) (rs # r2)) + with (List.fold_left (fun v i => Val.add v (Vint i)) (decompose_int n) (Val.neg (rs # r2))). apply iterate_op_correct. + destruct v; simpl; auto. auto. - intros. simpl. destruct (rs r2); auto. simpl. rewrite Int.sub_add_opp. + intros. simpl. destruct (rs # r2); auto. simpl. rewrite Int.sub_add_opp. rewrite Int.add_commut; auto. rewrite decompose_int_add. - destruct (rs r2); simpl; auto. rewrite Int.sub_add_opp. rewrite Int.add_commut; auto. + destruct (rs # r2); simpl; auto. rewrite Int.sub_add_opp. rewrite Int.add_commut; auto. Qed. (** Or immediate *) @@ -462,9 +526,10 @@ Lemma orimm_correct: Proof. intros. unfold orimm. (* ori - ori* *) - replace (Val.or (rs r2) (Vint n)) - with (List.fold_left (fun v i => Val.or v (Vint i)) (decompose_int n) (rs r2)). + replace (Val.or (rs # r2) (Vint n)) + with (List.fold_left (fun v i => Val.or v (Vint i)) (decompose_int n) (rs # r2)). apply iterate_op_correct. + destruct v; simpl; auto. auto. auto. apply decompose_int_or. @@ -481,9 +546,10 @@ Lemma xorimm_correct: Proof. intros. unfold xorimm. (* xori - xori* *) - replace (Val.xor (rs r2) (Vint n)) - with (List.fold_left (fun v i => Val.xor v (Vint i)) (decompose_int n) (rs r2)). + replace (Val.xor (rs # r2) (Vint n)) + with (List.fold_left (fun v i => Val.xor v (Vint i)) (decompose_int n) (rs # r2)). apply iterate_op_correct. + destruct v; simpl; auto. auto. auto. apply decompose_int_xor. @@ -496,7 +562,7 @@ Lemma indexed_memory_access_correct: (mk_immed: int -> int) (base: ireg) n k (rs: regset) m m', (forall (r1: ireg) (rs1: regset) n1 k, Val.add rs1#r1 (Vint n1) = Val.add rs#base (Vint n) -> - (forall (r: preg), if_preg r = true -> r <> IR14 -> rs1 r = rs r) -> + (forall (r: preg), if_preg r = true -> r <> IR14 -> rs1 # r = rs # r) -> exists rs', exec_straight ge fn (mk_instr r1 n1 :: k) rs1 m k rs' m' /\ P rs') -> exists rs', @@ -522,97 +588,86 @@ Qed. Lemma loadind_int_correct: forall (base: ireg) ofs dst (rs: regset) m v k, - Mem.loadv Mint32 m (Val.offset_ptr rs#base ofs) = Some v -> + Mem.loadv Many32 m (Val.offset_ptr rs#base ofs) = Some v -> exists rs', exec_straight ge fn (loadind_int base ofs dst k) rs m k rs' m /\ rs'#dst = v /\ forall r, if_preg r = true -> r <> IR14 -> r <> dst -> rs'#r = rs#r. Proof. intros; unfold loadind_int. - assert (Val.offset_ptr (rs base) ofs = Val.add (rs base) (Vint (Ptrofs.to_int ofs))). - { destruct (rs base); try discriminate. simpl. f_equal; f_equal. symmetry; auto with ptrofs. } + assert (Val.offset_ptr (rs # base) ofs = Val.add (rs # base) (Vint (Ptrofs.to_int ofs))). + { destruct (rs # base); try discriminate. simpl. f_equal; f_equal. symmetry; auto with ptrofs. } apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_load. rewrite H1, <- H0, H. eauto. auto. + apply exec_straight_one. simpl. unfold exec_load. rewrite H1, <- H0, H. eauto. + Simpl. split; intros; Simpl. + apply Mem.loadv_type, Val.load_result_same in H. auto. Qed. Lemma loadind_correct: - forall (base: ireg) ofs ty dst k c (rs: regset) m v, - loadind base ofs ty dst k = OK c -> - Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v -> + forall (base: ireg) ofs q dst k c (rs: regset) m v, + loadind base ofs q dst k = OK c -> + Mem.loadv (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) = Some v -> exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of dst) = v /\ forall r, if_preg r = true -> r <> IR14 -> r <> preg_of dst -> rs'#r = rs#r. Proof. unfold loadind; intros. - assert (Val.offset_ptr (rs base) ofs = Val.add (rs base) (Vint (Ptrofs.to_int ofs))). - { destruct (rs base); try discriminate. simpl. f_equal; f_equal. symmetry; auto with ptrofs. } - destruct ty; destruct (preg_of dst); inv H; simpl in H0. -- (* int *) - apply loadind_int_correct; auto. -- (* float *) + assert (Val.offset_ptr (rs # base) ofs = Val.add (rs # base) (Vint (Ptrofs.to_int ofs))). + { destruct (rs # base); try discriminate. simpl. f_equal; f_equal. symmetry; auto with ptrofs. } + destruct q; destruct (preg_of dst); inv H; simpl in H0. +- (* any32 to general-purpose register *) apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. auto. + apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. Simpl. split; intros; Simpl. -- (* single *) + apply Mem.loadv_type in H0. simpl. destruct v; auto; inversion H0. +- (* any32 to single-precision register *) apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. auto. + apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. Simpl. split; intros; Simpl. -- (* any32 *) +- (* any64 to double-precision register *) apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. auto. - split; intros; Simpl. -- (* any64 *) - apply indexed_memory_access_correct; intros. - econstructor; split. - apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. auto. + apply exec_straight_one. simpl. unfold exec_load. rewrite H, <- H1, H0. eauto. Simpl. split; intros; Simpl. Qed. (** Indexed memory stores. *) Lemma storeind_correct: - forall (base: ireg) ofs ty src k c (rs: regset) m m', - storeind src base ofs ty k = OK c -> - Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' -> + forall (base: ireg) ofs q src k c (rs: regset) m m', + storeind src base ofs q k = OK c -> + Mem.storev (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' /\ forall r, if_preg r = true -> r <> IR14 -> rs'#r = rs#r. Proof. unfold storeind; intros. assert (DATA: data_preg (preg_of src) = true) by eauto with asmgen. - assert (Val.offset_ptr (rs base) ofs = Val.add (rs base) (Vint (Ptrofs.to_int ofs))). - { destruct (rs base); try discriminate. simpl. f_equal; f_equal. symmetry; auto with ptrofs. } - destruct ty; destruct (preg_of src); inv H; simpl in H0. -- (* int *) - apply indexed_memory_access_correct; intros. - econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. auto. - intros; Simpl. -- (* float *) - apply indexed_memory_access_correct; intros. - econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. auto. - intros; Simpl. -- (* single *) + assert (Val.offset_ptr (rs # base) ofs = Val.add (rs # base) (Vint (Ptrofs.to_int ofs))). + { destruct (rs # base); try discriminate. simpl. f_equal; f_equal. symmetry; auto with ptrofs. } + destruct q; destruct (preg_of src); inv H; simpl in H0. +- (* any32 from general-purpose register *) apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. + apply nextinstr_pc. intros; Simpl. -- (* any32 *) +- (* any32 from single-precision register *) apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. + apply nextinstr_pc. intros; Simpl. -- (* any64 *) +- (* any64 from double-precision register *) apply indexed_memory_access_correct; intros. econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H, <- H1, H2, H0 by auto with asmgen; eauto. + apply nextinstr_pc. intros; Simpl. Qed. @@ -620,7 +675,7 @@ Qed. Lemma save_lr_correct: forall ofs k (rs: regset) m m', - Mem.storev Mint32 m (Val.offset_ptr rs#IR13 ofs) (rs#IR14) = Some m' -> + Mem.storev Many32 m (Val.offset_ptr rs#IR13 ofs) (rs#IR14) = Some m' -> exists rs', exec_straight ge fn (save_lr ofs k) rs m k rs' m' /\ (forall r, if_preg r = true -> r <> IR12 -> rs'#r = rs#r) @@ -631,9 +686,10 @@ Proof. assert (EQ: Val.offset_ptr rs#IR13 ofs = Val.add rs#IR13 (Vint n)). { destruct rs#IR13; try discriminate. simpl. f_equal; f_equal. unfold n; symmetry; auto with ptrofs. } destruct (Int.eq n n1). -- econstructor; split. apply exec_straight_one. simpl; unfold exec_store. rewrite <- EQ, H; reflexivity. auto. +- econstructor; split. apply exec_straight_one. simpl; unfold exec_store. rewrite <- EQ, H; reflexivity. + apply nextinstr_pc. split. intros; Simpl. intros; Simpl. -- destruct (addimm_correct IR12 IR13 (Int.sub n n1) (Pstr IR14 IR12 (SOimm n1) :: k) rs m) +- destruct (addimm_correct IR12 IR13 (Int.sub n n1) (Pstr_a IR14 IR12 (SOimm n1) :: k) rs m) as (rs1 & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. @@ -643,7 +699,7 @@ Proof. rewrite (Int.add_commut (Int.neg n1)). rewrite Int.add_neg_zero. rewrite Int.add_zero. rewrite <- EQ. rewrite C by eauto with asmgen. rewrite H. reflexivity. - auto. + Simpl. split. intros; Simpl. congruence. Qed. @@ -658,6 +714,16 @@ Qed. (** Translation of conditions *) +Ltac decompose_compare := + match goal with + | [ |- Pregmap.get ?C (_ # ?C <- _) = _ ] => rewrite Pregmap.gss; decompose_compare + | [ |- Pregmap.get ?C (_ # _ <- _) = _ ] => rewrite Pregmap.gso; decompose_compare + | [ |- _ <> _ ] => compute; intuition congruence + | [ |- Val.load_result _ ?v = ?v ] => + change (Preg.chunk_of _) with (chunk_of_type Tany32); rewrite Val.load_result_same + | _ => idtac + end. + Lemma compare_int_spec: forall rs v1 v2 m, let rs1 := nextinstr (compare_int rs v1 v2 m) in @@ -666,7 +732,12 @@ Lemma compare_int_spec: /\ rs1#CC = Val.cmpu (Mem.valid_pointer m) Cge v1 v2 /\ rs1#CV = Val.sub_overflow v1 v2. Proof. - intros. unfold rs1. intuition. + intros. unfold rs1. Simpl. unfold compare_int. + repeat split; decompose_compare; auto. + destruct (Val.sub v1 v2); simpl; auto. + unfold Val.cmpu. destruct (Val.cmpu_bool _ _ v1 v2); try destruct b; simpl; auto. + unfold Val.cmpu. destruct (Val.cmpu_bool _ _ v1 v2); try destruct b; simpl; auto. + destruct v1, v2; simpl; auto. Qed. Lemma compare_int_inv: @@ -678,6 +749,14 @@ Proof. repeat Simplif. Qed. +Lemma compare_int_nextpc: + forall rs v1 v2 m, + (nextinstr (compare_int rs v1 v2 m)) # PC = Val.offset_ptr (rs # PC) Ptrofs.one. +Proof. + intros. unfold compare_int. Simplif. + destruct v1; destruct v2; repeat rewrite Pregmap.gso; auto; congruence. +Qed. + Lemma int_signed_eq: forall x y, Int.eq x y = zeq (Int.signed x) (Int.signed y). Proof. @@ -804,7 +883,8 @@ Lemma compare_float_spec: /\ rs1#CC = Val.of_bool (negb (Float.cmp Clt f1 f2)) /\ rs1#CV = Val.of_bool (negb (Float.cmp Ceq f1 f2 || Float.cmp Clt f1 f2 || Float.cmp Cgt f1 f2)). Proof. - intros. intuition. + intros. unfold rs1. Simpl. unfold compare_float. + repeat split; decompose_compare; auto; repeat destruct (Float.cmp _ _ _); simpl; auto. Qed. Lemma compare_float_inv: @@ -813,7 +893,7 @@ Lemma compare_float_inv: forall r', data_preg r' = true -> rs1#r' = rs#r'. Proof. intros. unfold rs1, compare_float. - assert (nextinstr (rs#CN <- Vundef #CZ <- Vundef #CC <- Vundef #CV <- Vundef) r' = rs r'). + assert ((nextinstr (rs#CN <- Vundef #CZ <- Vundef #CC <- Vundef #CV <- Vundef)) # r' = rs # r'). { repeat Simplif. } destruct v1; destruct v2; auto. repeat Simplif. @@ -821,9 +901,10 @@ Qed. Lemma compare_float_nextpc: forall rs v1 v2, - nextinstr (compare_float rs v1 v2) PC = Val.offset_ptr (rs PC) Ptrofs.one. + (nextinstr (compare_float rs v1 v2)) # PC = Val.offset_ptr (rs # PC) Ptrofs.one. Proof. - intros. unfold compare_float. destruct v1; destruct v2; reflexivity. + intros. unfold compare_float. Simplif. + destruct v1; destruct v2; repeat rewrite Pregmap.gso; auto; congruence. Qed. Lemma cond_for_float_cmp_correct: @@ -910,7 +991,8 @@ Lemma compare_float32_spec: /\ rs1#CC = Val.of_bool (negb (Float32.cmp Clt f1 f2)) /\ rs1#CV = Val.of_bool (negb (Float32.cmp Ceq f1 f2 || Float32.cmp Clt f1 f2 || Float32.cmp Cgt f1 f2)). Proof. - intros. intuition. + intros. unfold rs1, compare_float32. + repeat split; Simplif; decompose_compare; auto; repeat destruct (Float32.cmp _ _ _); simpl; auto. Qed. Lemma compare_float32_inv: @@ -919,7 +1001,7 @@ Lemma compare_float32_inv: forall r', data_preg r' = true -> rs1#r' = rs#r'. Proof. intros. unfold rs1, compare_float32. - assert (nextinstr (rs#CN <- Vundef #CZ <- Vundef #CC <- Vundef #CV <- Vundef) r' = rs r'). + assert ((nextinstr (rs#CN <- Vundef #CZ <- Vundef #CC <- Vundef #CV <- Vundef)) # r' = rs # r'). { repeat Simplif. } destruct v1; destruct v2; auto. repeat Simplif. @@ -927,9 +1009,10 @@ Qed. Lemma compare_float32_nextpc: forall rs v1 v2, - nextinstr (compare_float32 rs v1 v2) PC = Val.offset_ptr (rs PC) Ptrofs.one. + (nextinstr (compare_float32 rs v1 v2)) # PC = Val.offset_ptr (rs # PC) Ptrofs.one. Proof. - intros. unfold compare_float32. destruct v1; destruct v2; reflexivity. + intros. unfold compare_float32. Simplif. + destruct v1; destruct v2; repeat rewrite Pregmap.gso; auto; congruence. Qed. Lemma cond_for_float32_cmp_correct: @@ -1027,98 +1110,98 @@ Lemma transl_cond_correct: transl_cond cond args k = OK c -> exists rs', exec_straight ge fn c rs m k rs' m - /\ match eval_condition cond (map rs (map preg_of args)) m with + /\ match eval_condition cond (map (fun r => rs # r) (map preg_of args)) m with | Some b => eval_testcond (cond_for_cond cond) rs' = Some b /\ eval_testcond (cond_for_cond (negate_condition cond)) rs' = Some (negb b) | None => True end - /\ forall r, data_preg r = true -> rs'#r = rs r. + /\ forall r, data_preg r = true -> rs'#r = rs # r. Proof. intros until c; intros TR. unfold transl_cond in TR; destruct cond; ArgsInv. - (* Ccomp *) econstructor. - split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmp_bool c0 (rs x) (rs x0)) eqn:CMP; auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. + split. destruct (Val.cmp_bool c0 (rs # x) (rs # x0)) eqn:CMP; auto. split; apply cond_for_signed_cmp_correct; auto. rewrite Val.negate_cmp_bool, CMP; auto. apply compare_int_inv. - (* Ccompu *) econstructor. - split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (rs x0)) eqn:CMP; auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. + split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs # x) (rs # x0)) eqn:CMP; auto. split; apply cond_for_unsigned_cmp_correct; auto. rewrite Val.negate_cmpu_bool, CMP; auto. apply compare_int_inv. - (* Ccompshift *) econstructor. - split. apply exec_straight_one. simpl. eauto. auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. split. rewrite transl_shift_correct. - destruct (Val.cmp_bool c0 (rs x) (eval_shift s (rs x0))) eqn:CMP; auto. + destruct (Val.cmp_bool c0 (rs # x) (eval_shift s (rs # x0))) eqn:CMP; auto. split; apply cond_for_signed_cmp_correct; auto. rewrite Val.negate_cmp_bool, CMP; auto. apply compare_int_inv. - (* Ccompushift *) econstructor. - split. apply exec_straight_one. simpl. eauto. auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. split. rewrite transl_shift_correct. - destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (eval_shift s (rs x0))) eqn:CMP; auto. + destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs # x) (eval_shift s (rs # x0))) eqn:CMP; auto. split; apply cond_for_unsigned_cmp_correct; auto. rewrite Val.negate_cmpu_bool, CMP; auto. apply compare_int_inv. - (* Ccompimm *) destruct (is_immed_arith i). econstructor. - split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmp_bool c0 (rs x) (Vint i)) eqn:CMP; auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. + split. destruct (Val.cmp_bool c0 (rs # x) (Vint i)) eqn:CMP; auto. split; apply cond_for_signed_cmp_correct; auto. rewrite Val.negate_cmp_bool, CMP; auto. apply compare_int_inv. destruct (is_immed_arith (Int.neg i)). econstructor. - split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmp_bool c0 (rs x) (Vint i)) eqn:CMP; auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. + split. destruct (Val.cmp_bool c0 (rs # x) (Vint i)) eqn:CMP; auto. split; apply cond_for_signed_cmp_correct; rewrite Int.neg_involutive; auto. rewrite Val.negate_cmp_bool, CMP; auto. apply compare_int_inv. exploit (loadimm_correct IR14). intros [rs' [P [Q R]]]. econstructor. split. eapply exec_straight_trans. eexact P. apply exec_straight_one. simpl. - rewrite Q. rewrite R; eauto with asmgen. auto. + rewrite Q. rewrite R; eauto with asmgen. auto using compare_int_nextpc. split. rewrite <- R by (eauto with asmgen). - destruct (Val.cmp_bool c0 (rs' x) (Vint i)) eqn:CMP; auto. + destruct (Val.cmp_bool c0 (rs' # x) (Vint i)) eqn:CMP; auto. split; apply cond_for_signed_cmp_correct; auto. rewrite Val.negate_cmp_bool, CMP; auto. intros. rewrite compare_int_inv by auto. auto with asmgen. - (* Ccompuimm *) destruct (is_immed_arith i). econstructor. - split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (Vint i)) eqn:CMP; auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. + split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs # x) (Vint i)) eqn:CMP; auto. split; apply cond_for_unsigned_cmp_correct; auto. rewrite Val.negate_cmpu_bool, CMP; auto. apply compare_int_inv. destruct (is_immed_arith (Int.neg i)). econstructor. - split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (Vint i)) eqn:CMP; auto. + split. apply exec_straight_one. simpl. eauto. auto using compare_int_nextpc. + split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs # x) (Vint i)) eqn:CMP; auto. split; apply cond_for_unsigned_cmp_correct; rewrite Int.neg_involutive; auto. rewrite Val.negate_cmpu_bool, CMP; auto. apply compare_int_inv. exploit (loadimm_correct IR14). intros [rs' [P [Q R]]]. econstructor. split. eapply exec_straight_trans. eexact P. apply exec_straight_one. simpl. - rewrite Q. rewrite R; eauto with asmgen. auto. + rewrite Q. rewrite R; eauto with asmgen. auto using compare_int_nextpc. split. rewrite <- R by (eauto with asmgen). - destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs' x) (Vint i)) eqn:CMP; auto. + destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs' # x) (Vint i)) eqn:CMP; auto. split; apply cond_for_unsigned_cmp_correct; auto. rewrite Val.negate_cmpu_bool, CMP; auto. intros. rewrite compare_int_inv by auto. auto with asmgen. - (* Ccompf *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float_nextpc. - split. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) eqn:CMP; auto. - destruct (rs x); try discriminate. destruct (rs x0); try discriminate. + split. destruct (Val.cmpf_bool c0 (rs # x) (rs # x0)) eqn:CMP; auto. + destruct (rs # x); try discriminate. destruct (rs # x0); try discriminate. simpl in CMP. inv CMP. split. apply cond_for_float_cmp_correct. apply cond_for_float_not_cmp_correct. apply compare_float_inv. - (* Cnotcompf *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float_nextpc. - split. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) eqn:CMP; auto. - destruct (rs x); try discriminate. destruct (rs x0); try discriminate. + split. destruct (Val.cmpf_bool c0 (rs # x) (rs # x0)) eqn:CMP; auto. + destruct (rs # x); try discriminate. destruct (rs # x0); try discriminate. simpl in CMP. inv CMP. Local Opaque compare_float. simpl. split. apply cond_for_float_not_cmp_correct. rewrite negb_involutive. apply cond_for_float_cmp_correct. @@ -1127,16 +1210,16 @@ Local Opaque compare_float. simpl. - (* Ccompfzero *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float_nextpc. - split. destruct (Val.cmpf_bool c0 (rs x) (Vfloat Float.zero)) eqn:CMP; auto. - destruct (rs x); try discriminate. + split. destruct (Val.cmpf_bool c0 (rs # x) (Vfloat Float.zero)) eqn:CMP; auto. + destruct (rs # x); try discriminate. simpl in CMP. inv CMP. split. apply cond_for_float_cmp_correct. apply cond_for_float_not_cmp_correct. apply compare_float_inv. - (* Cnotcompfzero *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float_nextpc. - split. destruct (Val.cmpf_bool c0 (rs x) (Vfloat Float.zero)) eqn:CMP; auto. - destruct (rs x); try discriminate. simpl in CMP. inv CMP. + split. destruct (Val.cmpf_bool c0 (rs # x) (Vfloat Float.zero)) eqn:CMP; auto. + destruct (rs # x); try discriminate. simpl in CMP. inv CMP. Local Opaque compare_float. simpl. split. apply cond_for_float_not_cmp_correct. rewrite negb_involutive. apply cond_for_float_cmp_correct. exact I. @@ -1144,16 +1227,16 @@ Local Opaque compare_float. simpl. - (* Ccompfs *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float32_nextpc. - split. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) eqn:CMP; auto. - destruct (rs x); try discriminate. destruct (rs x0); try discriminate. + split. destruct (Val.cmpfs_bool c0 (rs # x) (rs # x0)) eqn:CMP; auto. + destruct (rs # x); try discriminate. destruct (rs # x0); try discriminate. simpl in CMP. inv CMP. split. apply cond_for_float32_cmp_correct. apply cond_for_float32_not_cmp_correct. apply compare_float32_inv. - (* Cnotcompfs *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float32_nextpc. - split. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) eqn:CMP; auto. - destruct (rs x); try discriminate. destruct (rs x0); try discriminate. + split. destruct (Val.cmpfs_bool c0 (rs # x) (rs # x0)) eqn:CMP; auto. + destruct (rs # x); try discriminate. destruct (rs # x0); try discriminate. simpl in CMP. inv CMP. Local Opaque compare_float32. simpl. split. apply cond_for_float32_not_cmp_correct. rewrite negb_involutive. apply cond_for_float32_cmp_correct. @@ -1162,16 +1245,16 @@ Local Opaque compare_float32. simpl. - (* Ccompfszero *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float32_nextpc. - split. destruct (Val.cmpfs_bool c0 (rs x) (Vsingle Float32.zero)) eqn:CMP; auto. - destruct (rs x); try discriminate. + split. destruct (Val.cmpfs_bool c0 (rs # x) (Vsingle Float32.zero)) eqn:CMP; auto. + destruct (rs # x); try discriminate. simpl in CMP. inv CMP. split. apply cond_for_float32_cmp_correct. apply cond_for_float32_not_cmp_correct. apply compare_float32_inv. - (* Cnotcompfzero *) econstructor. split. apply exec_straight_one. simpl. eauto. apply compare_float32_nextpc. - split. destruct (Val.cmpfs_bool c0 (rs x) (Vsingle Float32.zero)) eqn:CMP; auto. - destruct (rs x); try discriminate. simpl in CMP. inv CMP. + split. destruct (Val.cmpfs_bool c0 (rs # x) (Vsingle Float32.zero)) eqn:CMP; auto. + destruct (rs # x); try discriminate. simpl in CMP. inv CMP. simpl. split. apply cond_for_float32_not_cmp_correct. rewrite negb_involutive. apply cond_for_float32_cmp_correct. exact I. apply compare_float32_inv. @@ -1179,15 +1262,40 @@ Qed. (** Translation of arithmetic operations. *) +Ltac invert_load_result := + match goal with + | [ |- Val.load_result ?c (?f ?arg1 ?arg2 ?arg3) = ?f ?arg1 ?arg2 ?arg3 ] => + destruct arg1, arg2, arg3; simpl; auto + | [ |- Val.load_result ?c (?f ?arg1 ?arg2) = ?f ?arg1 ?arg2 ] => + destruct arg1, arg2; simpl; auto; + match goal with + | [ |- match (if ?cond then _ else _) with _ => _ end = if ?cond then _ else _ ] => + destruct cond; auto + | _ => + idtac + end + | [ |- Val.load_result ?c (?f ?arg1) = ?f ?arg1 ] => + destruct arg1; simpl; auto + end. + +Ltac ResolveNextinstrPC := + match goal with + | [ |- Pregmap.get PC (nextinstr _) = Val.offset_ptr (Pregmap.get PC _) _ ] => + rewrite if_preg_nextinstr_PC; Simpl; reflexivity + | [ |- Pregmap.get PC (nextinstr_nf _) = Val.offset_ptr (Pregmap.get PC _) _ ] => + rewrite if_preg_nextinstr_nf_PC; Simpl; reflexivity + | _ => reflexivity + end. + Ltac TranslOpSimpl := econstructor; split; - [ apply exec_straight_one; [simpl; eauto | reflexivity ] + [ apply exec_straight_one; [simpl; eauto | ResolveNextinstrPC ] | split; [try rewrite transl_shift_correct; repeat Simpl | intros; repeat Simpl] ]. Lemma transl_op_correct_same: forall op args res k c (rs: regset) m v, transl_op op args res k = OK c -> - eval_operation ge rs#IR13 op (map rs (map preg_of args)) m = Some v -> + eval_operation ge rs#IR13 op (map (fun r => rs # r) (map preg_of args)) m = Some v -> match op with Ocmp _ => False | Oaddrstack _ => False | _ => True end -> exists rs', exec_straight ge fn c rs m k rs' m @@ -1195,40 +1303,53 @@ Lemma transl_op_correct_same: /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs'#r = rs#r. Proof. intros until v; intros TR EV NOCMP. - unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; inv EV; try (TranslOpSimpl; fail). + unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; inv EV; try (TranslOpSimpl; try invert_load_result; fail). (* Omove *) destruct (preg_of res) eqn:RES; try discriminate; destruct (preg_of m0) eqn:ARG; inv TR. - econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. - econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. + (* ireg *) + econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. intuition Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + change (Preg.type i) with (Preg.type i0). + apply Pregmap.get_has_type. + (* freg *) + econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. intuition Simpl. (* Ointconst *) generalize (loadimm_correct x i k rs m). intros [rs' [A [B C]]]. exists rs'; auto with asmgen. + (* Oaddrsymbol *) + econstructor; split. apply exec_straight_one; simpl; eauto using if_preg_nextinstr_PC. + intuition Simpl. + change (Preg.chunk_of x) with (chunk_of_type Tany32). + apply Val.load_result_same; auto. + unfold Genv.symbol_address. destruct (Genv.find_symbol ge i); simpl; auto. (* Oaddrstack *) contradiction. (* Ocast8signed *) destruct Archi.thumb2_support. - econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. - destruct (rs x0); auto; simpl. rewrite Int.shru_zero. reflexivity. + econstructor; split. apply exec_straight_one; simpl; eauto using if_preg_nextinstr_PC. + intuition Simpl. + destruct (rs # x0); auto; simpl. rewrite Int.shru_zero. reflexivity. set (rs1 := nextinstr_nf (rs#x <- (Val.shl rs#x0 (Vint (Int.repr 24))))). set (rs2 := nextinstr_nf (rs1#x <- (Val.shr rs1#x (Vint (Int.repr 24))))). exists rs2. - split. apply exec_straight_two with rs1 m; auto. + split. apply exec_straight_two with rs1 m; eauto using if_preg_nextinstr_nf_PC. split. unfold rs2; Simpl. unfold rs1; Simpl. - unfold Val.shr, Val.shl; destruct (rs x0); auto. + unfold Val.shr, Val.shl; destruct (rs # x0); auto. change (Int.ltu (Int.repr 24) Int.iwordsize) with true; simpl. f_equal. symmetry. apply (Int.sign_ext_shr_shl 8). compute; auto. intros. unfold rs2, rs1; Simpl. (* Ocast16signed *) destruct Archi.thumb2_support. - econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. - destruct (rs x0); auto; simpl. rewrite Int.shru_zero. reflexivity. + econstructor; split. apply exec_straight_one; simpl; eauto using if_preg_nextinstr_PC. + intuition Simpl. + destruct (rs # x0); auto; simpl. rewrite Int.shru_zero. reflexivity. set (rs1 := nextinstr_nf (rs#x <- (Val.shl rs#x0 (Vint (Int.repr 16))))). set (rs2 := nextinstr_nf (rs1#x <- (Val.shr rs1#x (Vint (Int.repr 16))))). exists rs2. - split. apply exec_straight_two with rs1 m; auto. + split. apply exec_straight_two with rs1 m; eauto using if_preg_nextinstr_nf_PC. split. unfold rs2; Simpl. unfold rs1; Simpl. - unfold Val.shr, Val.shl; destruct (rs x0); auto. + unfold Val.shr, Val.shl; destruct (rs # x0); auto. change (Int.ltu (Int.repr 16) Int.iwordsize) with true; simpl. f_equal. symmetry. apply (Int.sign_ext_shr_shl 16). compute; auto. intros. unfold rs2, rs1; Simpl. @@ -1240,13 +1361,24 @@ Proof. generalize (rsubimm_correct x x0 i k rs m). intros [rs' [A [B C]]]. exists rs'; auto with asmgen. + (* divs *) Local Transparent destroyed_by_op. - econstructor. split. apply exec_straight_one. simpl. rewrite H0. reflexivity. auto. - split. Simpl. simpl; intros. intuition Simpl. + econstructor. split. apply exec_straight_one. simpl. rewrite H0. reflexivity. + unfold nextinstr. intuition Simpl; invert_load_result. + split. simpl. intuition Simpl. + destruct (rs#IR0), (rs#IR1); inv H0. + destruct (_ || _ && _); inv H1; auto. + intros. Simpl. simpl in H2; decompose [and] H2. + rewrite !Pregmap.gso; auto. (* divu *) - econstructor. split. apply exec_straight_one. simpl. rewrite H0. reflexivity. auto. - split. Simpl. simpl; intros. intuition Simpl. + econstructor. split. apply exec_straight_one. simpl. rewrite H0. reflexivity. + unfold nextinstr. intuition Simpl; invert_load_result. + split. simpl. intuition Simpl. + destruct (rs#IR0), (rs#IR1); inv H0. + destruct (Int.eq i0 Int.zero); inv H1; auto. + intros. Simpl. simpl in H2; decompose [and] H2. + rewrite !Pregmap.gso; auto. (* Oandimm *) generalize (andimm_correct x x0 i k rs m). intros [rs' [A [B C]]]. @@ -1260,14 +1392,15 @@ Local Transparent destroyed_by_op. intros [rs' [A [B C]]]. exists rs'; auto with asmgen. (* Oshrximm *) - destruct (rs x0) eqn: X0; simpl in H0; try discriminate. + destruct (rs # x0) eqn: X0; simpl in H0; try discriminate. destruct (Int.ltu i (Int.repr 31)) eqn: LTU; inv H0. revert EQ2. predSpec Int.eq Int.eq_spec i Int.zero; intros EQ2. (* i = 0 *) inv EQ2. econstructor. - split. apply exec_straight_one. simpl. reflexivity. auto. + split. apply exec_straight_one. simpl. reflexivity. Simpl. split. Simpl. unfold Int.shrx. rewrite Int.shl_zero. unfold Int.divs. - change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed. auto. + change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed. + rewrite X0; reflexivity. intros. Simpl. (* i <> 0 *) inv EQ2. @@ -1296,40 +1429,49 @@ Local Transparent destroyed_by_op. exists rs3; split. apply exec_straight_three with rs1 m rs2 m. simpl. rewrite X0; reflexivity. - simpl. f_equal. Simpl. replace (rs1 x0) with (rs x0). rewrite X0; reflexivity. + simpl. f_equal. Simpl. replace (rs1 # x0) with (rs # x0). rewrite X0; reflexivity. unfold rs1; Simpl. reflexivity. - auto. auto. auto. + eauto using if_preg_nextinstr_nf_PC. + eauto using if_preg_nextinstr_nf_PC. + eauto using if_preg_nextinstr_nf_PC. split. unfold rs3; Simpl. unfold rs2; Simpl. unfold rs1; Simpl. simpl. change (Int.ltu (Int.repr 31) Int.iwordsize) with true. simpl. rewrite LTU'; simpl. rewrite LTU''; simpl. f_equal. symmetry. apply Int.shrx_shr_2. assumption. intros. unfold rs3; Simpl. unfold rs2; Simpl. unfold rs1; Simpl. (* intoffloat *) - econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. + econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. Simpl. Transparent destroyed_by_op. - simpl. intuition Simpl. + split. Simpl. destruct (rs#x0); inv H0. destruct (Float.to_int f); inv H1; auto. + simpl; intros. decompose [and] H2. Simpl. (* intuoffloat *) econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. simpl. intuition Simpl. + split. Simpl. destruct (rs#x0); inv H0. destruct (Float.to_intu f); inv H1; auto. + simpl; intros. decompose [and] H2. Simpl. (* floatofint *) - econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. + econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. Simpl. intuition Simpl. (* floatofintu *) - econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. + econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. Simpl. intuition Simpl. (* intofsingle *) econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. simpl. intuition Simpl. + split. Simpl. destruct (rs#x0); inv H0. destruct (Float32.to_int f); inv H1; auto. + simpl; intros. decompose [and] H2. Simpl. (* intuofsingle *) econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. simpl. intuition Simpl. + split. Simpl. destruct (rs#x0); inv H0. destruct (Float32.to_intu f); inv H1; auto. + simpl; intros. decompose [and] H2. Simpl. (* singleofint *) - econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. - intuition Simpl. + econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. Simpl. + split. Simpl. intuition Simpl. (* singleofintu *) - econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. auto. - intuition Simpl. + econstructor; split. apply exec_straight_one; simpl. rewrite H0; simpl. eauto. Simpl. + split. Simpl. intuition Simpl. (* Ocmp *) contradiction. Qed. @@ -1337,7 +1479,7 @@ Qed. Lemma transl_op_correct: forall op args res k c (rs: regset) m v, transl_op op args res k = OK c -> - eval_operation ge rs#IR13 op (map rs (map preg_of args)) m = Some v -> + eval_operation ge rs#IR13 op (map (fun r => rs#r) (map preg_of args)) m = Some v -> exists rs', exec_straight ge fn c rs m k rs' m /\ Val.lessdef v rs'#(preg_of res) @@ -1357,16 +1499,17 @@ Proof. clear SAME; simpl in *; ArgsInv. destruct (addimm_correct x IR13 (Ptrofs.to_int i) k rs m) as [rs' [EX [RES OTH]]]. exists rs'; split. auto. split. - rewrite RES; inv H0. destruct (rs IR13); simpl; auto. rewrite Ptrofs.of_int_to_int; auto. + rewrite RES; inv H0. destruct (rs # IR13); simpl; auto. rewrite Ptrofs.of_int_to_int; auto. intros; apply OTH; eauto with asmgen. - (* Ocmp *) clear SAME. simpl in H. monadInv H. simpl in H0. inv H0. rewrite (ireg_of_eq _ _ EQ). exploit transl_cond_correct; eauto. instantiate (1 := rs). instantiate (1 := m). intros [rs1 [A [B C]]]. econstructor; split. - eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto. + eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. + auto using if_preg_nextinstr_PC. split; intros; Simpl. - destruct (eval_condition c0 rs ## (preg_of ## args) m) as [b|]; simpl; auto. + destruct (eval_condition c0 (map (fun r => rs#r) (map preg_of args)) m) as [b|]; simpl; auto. destruct B as [B1 B2]; rewrite B1. destruct b; auto. Qed. @@ -1387,11 +1530,11 @@ Lemma transl_memory_access_correct: (mk_immed: int -> int) addr args k c (rs: regset) a m m', transl_memory_access mk_instr_imm mk_instr_gen mk_immed addr args k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> match a with Vptr _ _ => True | _ => False end -> (forall (r1: ireg) (rs1: regset) n k, Val.add rs1#r1 (Vint n) = a -> - (forall (r: preg), if_preg r = true -> r <> IR14 -> rs1 r = rs r) -> + (forall (r: preg), if_preg r = true -> r <> IR14 -> rs1 # r = rs # r) -> exists rs', exec_straight ge fn (mk_instr_imm r1 n :: k) rs1 m k rs' m' /\ P rs') -> match mk_instr_gen with @@ -1417,14 +1560,32 @@ Proof. erewrite ! ireg_of_eq; eauto. rewrite transl_shift_correct. auto. (* Ainstack *) inv TR. apply indexed_memory_access_correct. intros. eapply MK1; eauto. - rewrite H. destruct (rs IR13); try contradiction. simpl. f_equal; f_equal. auto with ptrofs. + rewrite H. destruct (rs # IR13); try contradiction. simpl. f_equal; f_equal. auto with ptrofs. +Qed. + +Lemma loadv_Many32_aux: + forall chunk m a v x, + quantity_chunk chunk = Q32 -> + Mem.loadv chunk m a = Some v -> + Preg.type x = Tany32 -> + Val.load_result (Preg.chunk_of x) v = v. +Proof. + intros. + apply Mem.loadv_type in H0. + assert (T: Val.has_type v Tany32). + { apply Val.has_subtype with (ty1 := type_of_chunk chunk); auto. + destruct chunk; simpl in *; congruence. } + replace (Preg.chunk_of x) with (chunk_of_type Tany32). + apply Val.load_result_same; auto. + destruct x; auto. inversion H1. Qed. Lemma transl_load_int_correct: forall mk_instr is_immed dst addr args k c (rs: regset) a chunk m v, transl_memory_access_int mk_instr is_immed dst addr args k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> + quantity_chunk chunk = Q32 -> (forall (r1 r2: ireg) (sa: shift_op) (rs1: regset), exec_instr ge fn (mk_instr r1 r2 sa) rs1 m = exec_load chunk (Val.add rs1#r2 (eval_shift_op sa rs1)) r1 rs1 m) -> @@ -1437,18 +1598,20 @@ Proof. eapply transl_memory_access_correct; eauto. destruct a; discriminate || trivial. intros; simpl. econstructor; split. apply exec_straight_one. - rewrite H2. unfold exec_load. simpl eval_shift_op. rewrite H. rewrite H1. eauto. auto. - split. Simpl. intros; Simpl. + rewrite H3. unfold exec_load. simpl eval_shift_op. rewrite H. rewrite H1. eauto. Simpl. + split. Simpl. eauto using loadv_Many32_aux. + intros; Simpl. simpl; intros. econstructor; split. apply exec_straight_one. - rewrite H2. unfold exec_load. rewrite H. rewrite H1. eauto. auto. - split. Simpl. intros; Simpl. + rewrite H3. unfold exec_load. rewrite H. rewrite H1. eauto. Simpl. + split. Simpl. eauto using loadv_Many32_aux. + intros; Simpl. Qed. Lemma transl_load_float_correct: forall mk_instr is_immed dst addr args k c (rs: regset) a chunk m v, transl_memory_access_float mk_instr is_immed dst addr args k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> (forall (r1: freg) (r2: ireg) (n: int) (rs1: regset), exec_instr ge fn (mk_instr r1 r2 n) rs1 m = @@ -1462,7 +1625,7 @@ Proof. eapply transl_memory_access_correct; eauto. destruct a; discriminate || trivial. intros; simpl. econstructor; split. apply exec_straight_one. - rewrite H2. unfold exec_load. rewrite H. rewrite H1. eauto. auto. + rewrite H2. unfold exec_load. rewrite H. rewrite H1. eauto. Simpl. split. Simpl. intros; Simpl. simpl; auto. Qed. @@ -1470,7 +1633,7 @@ Qed. Lemma transl_store_int_correct: forall mr mk_instr is_immed src addr args k c (rs: regset) a chunk m m', transl_memory_access_int mk_instr is_immed src addr args k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.storev chunk m a rs#(preg_of src) = Some m' -> (forall (r1 r2: ireg) (sa: shift_op) (rs1: regset), exec_instr ge fn (mk_instr r1 r2 sa) rs1 m = @@ -1487,16 +1650,16 @@ Proof. rewrite H2. unfold exec_store. simpl eval_shift_op. rewrite H. rewrite H3; eauto with asmgen. rewrite H1. eauto. auto. intros; Simpl. - simpl; intros. + simpl; intros. Simpl. econstructor; split. apply exec_straight_one. - rewrite H2. unfold exec_store. rewrite H. rewrite H1. eauto. auto. + rewrite H2. unfold exec_store. rewrite H. rewrite H1. eauto. Simpl. intros; Simpl. Qed. Lemma transl_store_float_correct: forall mr mk_instr is_immed src addr args k c (rs: regset) a chunk m m', transl_memory_access_float mk_instr is_immed src addr args k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.storev chunk m a rs#(preg_of src) = Some m' -> (forall (r1: freg) (r2: ireg) (n: int) (rs1: regset), exec_instr ge fn (mk_instr r1 r2 n) rs1 m = @@ -1510,7 +1673,7 @@ Proof. eapply transl_memory_access_correct; eauto. destruct a; discriminate || trivial. intros; simpl. econstructor; split. apply exec_straight_one. - rewrite H2. unfold exec_store. rewrite H. rewrite H3; auto with asmgen. rewrite H1. eauto. auto. + rewrite H2. unfold exec_store. rewrite H. rewrite H3; auto with asmgen. rewrite H1. eauto. Simpl. intros; Simpl. simpl; auto. Qed. @@ -1518,7 +1681,7 @@ Qed. Lemma transl_load_correct: forall chunk addr args dst k c (rs: regset) a m v, transl_load chunk addr args dst k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> exists rs', exec_straight ge fn c rs m k rs' m @@ -1541,18 +1704,18 @@ Qed. Lemma transl_store_correct: forall chunk addr args src k c (rs: regset) a m m', transl_store chunk addr args src k = OK c -> - eval_addressing ge (rs#SP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#SP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.storev chunk m a rs#(preg_of src) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' /\ forall r, data_preg r = true -> preg_notin r (destroyed_by_store chunk addr) -> rs'#r = rs#r. Proof. intros. destruct chunk; simpl in H. -- assert (Mem.storev Mint8unsigned m a (rs (preg_of src)) = Some m'). +- assert (Mem.storev Mint8unsigned m a (rs # (preg_of src)) = Some m'). rewrite <- H1. destruct a; simpl; auto. symmetry. apply Mem.store_signed_unsigned_8. clear H1. eapply transl_store_int_correct; eauto. - eapply transl_store_int_correct; eauto. -- assert (Mem.storev Mint16unsigned m a (rs (preg_of src)) = Some m'). +- assert (Mem.storev Mint16unsigned m a (rs # (preg_of src)) = Some m'). rewrite <- H1. destruct a; simpl; auto. symmetry. apply Mem.store_signed_unsigned_16. clear H1. eapply transl_store_int_correct; eauto. - eapply transl_store_int_correct; eauto. diff --git a/arm/Conventions1.v b/arm/Conventions1.v index c5277e8dca..423d795523 100644 --- a/arm/Conventions1.v +++ b/arm/Conventions1.v @@ -16,6 +16,7 @@ Require Import Coqlib. Require Import Decidableplus. Require Import AST. +Require Import Values. Require Import Events. Require Import Locations. Require Archi. @@ -122,6 +123,17 @@ Proof. intros. unfold proj_sig_res, loc_result. destruct (sig_res sig) as [[]|]; destruct Archi.big_endian; auto. Qed. +Lemma loc_result_has_type: + forall res sig, + Val.has_type res (proj_sig_res sig) -> + Val.has_type_rpair res (loc_result sig) Val.loword Val.hiword mreg_type. +Proof. + intros. unfold Val.has_type_rpair, loc_result, proj_sig_res in *. + destruct sig; simpl. destruct sig_res; simpl in H. + destruct t, res, Archi.big_endian; simpl; auto. + destruct res; simpl; auto. +Qed. + (** The result locations are caller-save registers *) Lemma loc_result_caller_save: @@ -195,16 +207,16 @@ Fixpoint loc_arguments_hf | (Tint | Tany32) as ty :: tys => if zlt ir 4 then One (R (ireg_param ir)) :: loc_arguments_hf tys (ir + 1) fr ofs - else One (S Outgoing ofs ty) :: loc_arguments_hf tys ir fr (ofs + 1) + else One (S Outgoing ofs Q32) :: loc_arguments_hf tys ir fr (ofs + 1) | (Tfloat | Tany64) as ty :: tys => if zlt fr 8 then One (R (freg_param fr)) :: loc_arguments_hf tys ir (fr + 1) ofs else let ofs := align ofs 2 in - One (S Outgoing ofs ty) :: loc_arguments_hf tys ir fr (ofs + 2) + One (S Outgoing ofs Q64) :: loc_arguments_hf tys ir fr (ofs + 2) | Tsingle :: tys => if zlt fr 8 then One (R (freg_param fr)) :: loc_arguments_hf tys ir (fr + 1) ofs - else One (S Outgoing ofs Tsingle) :: loc_arguments_hf tys ir fr (ofs + 1) + else One (S Outgoing ofs Q32) :: loc_arguments_hf tys ir fr (ofs + 1) | Tlong :: tys => let ohi := if Archi.big_endian then 0 else 1 in let olo := if Archi.big_endian then 1 else 0 in @@ -212,7 +224,7 @@ Fixpoint loc_arguments_hf if zlt ir 4 then Twolong (R (ireg_param (ir + ohi))) (R (ireg_param (ir + olo))) :: loc_arguments_hf tys (ir + 2) fr ofs else let ofs := align ofs 2 in - Twolong (S Outgoing (ofs + ohi) Tint) (S Outgoing (ofs + olo) Tint) :: loc_arguments_hf tys ir fr (ofs + 2) + Twolong (S Outgoing (ofs + ohi) Q32) (S Outgoing (ofs + olo) Q32) :: loc_arguments_hf tys ir fr (ofs + 2) end. (** For the "softfloat" configuration, as well as for variable-argument functions @@ -240,21 +252,21 @@ Fixpoint loc_arguments_sf match tyl with | nil => nil | (Tint|Tany32) as ty :: tys => - One (if zlt ofs 0 then R (ireg_param (ofs + 4)) else S Outgoing ofs ty) + One (if zlt ofs 0 then R (ireg_param (ofs + 4)) else S Outgoing ofs Q32) :: loc_arguments_sf tys (ofs + 1) | (Tfloat|Tany64) as ty :: tys => let ofs := align ofs 2 in - One (if zlt ofs 0 then R (freg_param (ofs + 4)) else S Outgoing ofs ty) + One (if zlt ofs 0 then R (freg_param (ofs + 4)) else S Outgoing ofs Q64) :: loc_arguments_sf tys (ofs + 2) | Tsingle :: tys => - One (if zlt ofs 0 then R (freg_param (ofs + 4)) else S Outgoing ofs Tsingle) + One (if zlt ofs 0 then R (freg_param (ofs + 4)) else S Outgoing ofs Q32) :: loc_arguments_sf tys (ofs + 1) | Tlong :: tys => let ohi := if Archi.big_endian then 0 else 1 in let olo := if Archi.big_endian then 1 else 0 in let ofs := align ofs 2 in - Twolong (if zlt ofs 0 then R (ireg_param (ofs+ohi+4)) else S Outgoing (ofs+ohi) Tint) - (if zlt ofs 0 then R (ireg_param (ofs+olo+4)) else S Outgoing (ofs+olo) Tint) + Twolong (if zlt ofs 0 then R (ireg_param (ofs+ohi+4)) else S Outgoing (ofs+ohi) Q32) + (if zlt ofs 0 then R (ireg_param (ofs+olo+4)) else S Outgoing (ofs+olo) Q32) :: loc_arguments_sf tys (ofs + 2) end. @@ -319,14 +331,14 @@ Definition size_arguments (s: signature) : Z := Definition loc_argument_acceptable (l: loc) : Prop := match l with | R r => is_callee_save r = false - | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs) + | S Outgoing ofs q => ofs >= 0 /\ (quantity_align q | ofs) | _ => False end. Definition loc_argument_charact (ofs: Z) (l: loc) : Prop := match l with | R r => is_callee_save r = false - | S Outgoing ofs' ty => ofs' >= ofs /\ typealign ty = 1 + | S Outgoing ofs' q => ofs' >= ofs /\ quantity_align q = 1 | _ => False end. @@ -525,9 +537,9 @@ Proof. 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. + forall ofs q tyl ir fr ofs0, + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_hf tyl ir fr ofs0)) -> + ofs + typesize (typ_of_quantity q) <= size_arguments_hf tyl ir fr ofs0. Proof. induction tyl; simpl; intros. elim H. @@ -579,9 +591,9 @@ Proof. 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. + forall ofs q tyl ofs0, + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_sf tyl ofs0)) -> + Z.max 0 (ofs + typesize (typ_of_quantity q)) <= size_arguments_sf tyl ofs0. Proof. induction tyl; simpl; intros. elim H. @@ -623,16 +635,16 @@ Proof. 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. + forall (s: signature) (ofs: Z) (q: quantity), + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments s)) -> + ofs + typesize (typ_of_quantity q) <= 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)). + assert (In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_sf (sig_args s) (-4))) -> + ofs + typesize (typ_of_quantity q) <= 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). + assert (In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_hf (sig_args s) 0 0 0)) -> + ofs + typesize (typ_of_quantity q) <= 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. diff --git a/arm/Machregs.v b/arm/Machregs.v index ae0ff6bfa5..c7ec2dca0b 100644 --- a/arm/Machregs.v +++ b/arm/Machregs.v @@ -16,6 +16,7 @@ Require Import Decidableplus. Require Import Maps. Require Import AST. Require Import Op. +Require Import Memdata. (** ** Machine registers *) @@ -72,6 +73,11 @@ Definition mreg_type (r: mreg): typ := | F8 | F9 | F10 | F11 | F12 | F13 | F14 | F15 => Tany64 end. +Lemma mreg_type_cases: forall r, mreg_type r = Tany32 \/ mreg_type r = Tany64. +Proof. + destruct r; simpl; auto. +Qed. + Open Scope positive_scope. Module IndexedMreg <: INDEXED_TYPE. @@ -79,20 +85,39 @@ Module IndexedMreg <: INDEXED_TYPE. Definition eq := mreg_eq. Definition index (r: mreg): positive := match r with - | R0 => 1 | R1 => 2 | R2 => 3 | R3 => 4 - | R4 => 5 | R5 => 6 | R6 => 7 | R7 => 8 - | R8 => 9 | R9 => 10 | R10 => 11 | R11 => 12 - | R12 => 13 - | F0 => 14 | F1 => 15 | F2 => 16 | F3 => 17 - | F4 => 18 | F5 => 19 | F6 => 20 | F7 => 21 - | F8 => 22 | F9 => 23 | F10 => 24 | F11 => 25 - | F12 => 26 | F13 => 27 | F14 => 28 | F15 => 29 + | R0 => 2 | R1 => 4 | R2 => 6 | R3 => 8 + | R4 => 10 | R5 => 12 | R6 => 14 | R7 => 16 + | R8 => 18 | R9 => 20 | R10 => 22 | R11 => 24 + | R12 => 26 + | F0 => 28 | F1 => 30 | F2 => 32 | F3 => 34 + | F4 => 36 | F5 => 38 | F6 => 40 | F7 => 42 + | F8 => 44 | F9 => 46 | F10 => 48 | F11 => 50 + | F12 => 52 | F13 => 54 | F14 => 56 | F15 => 58 end. Lemma index_inj: forall r1 r2, index r1 = index r2 -> r1 = r2. Proof. decide_goal. Qed. + + Open Scope Z_scope. + + Lemma scaled_index_with_size_aux: + forall r1 r2, Zpos (index r1) < Zpos (index r2) -> Zpos (index r1) + 2 <= Zpos (index r2). + Proof. + decide_goal. + Qed. + + Lemma scaled_index_with_size: + forall r1 r2, + Zpos (index r1) < Zpos (index r2) -> + Zpos (index r1) * 4 + AST.typesize (mreg_type r1) <= Zpos (index r2) * 4. + Proof. + intros. + generalize (scaled_index_with_size_aux r1 r2); intro. + assert (AST.typesize (mreg_type r1) <= 8) by (destruct (mreg_type r1); simpl; omega). + omega. + Qed. End IndexedMreg. Definition is_stack_reg (r: mreg) : bool := false. @@ -156,7 +181,7 @@ Definition destroyed_by_builtin (ef: external_function): list mreg := | _ => nil end. -Definition destroyed_by_setstack (ty: typ): list mreg := nil. +Definition destroyed_by_setstack (q: quantity): list mreg := nil. Definition destroyed_at_function_entry: list mreg := R12 :: nil. diff --git a/arm/Op.v b/arm/Op.v index 60c214d082..b22dc688a0 100644 --- a/arm/Op.v +++ b/arm/Op.v @@ -557,6 +557,17 @@ Proof. intros; discriminate. Qed. +Lemma is_not_move_operation: + forall (F V A: Type) (genv: Genv.t F V) (sp: val) + (op: operation) (f: A -> val) (args: list A) (m: mem) (v: val), + eval_operation genv sp op (map f args) m = Some v -> + is_move_operation op args = None -> + op <> Omove. +Proof. + intros. destruct (eq_operation op Omove); auto. + subst. destruct args; try destruct args; simpl in *; congruence. +Qed. + (** [negate_condition cond] returns a condition that is logically equivalent to the negation of [cond]. *) diff --git a/arm/TargetPrinter.ml b/arm/TargetPrinter.ml index 52d2ada67f..917deeea0c 100644 --- a/arm/TargetPrinter.ml +++ b/arm/TargetPrinter.ml @@ -421,11 +421,11 @@ struct fprintf oc " vcvt.f64.f32 %a, %a\n" freg r1 freg_single r2 | Pfldd(r1, r2, n) | Pfldd_a(r1, r2, n) -> fprintf oc " vldr %a, [%a, #%a]\n" freg r1 ireg r2 coqint n - | Pflds(r1, r2, n) -> + | Pflds(r1, r2, n) | Pflds_a(r1, r2, n) -> fprintf oc " vldr %a, [%a, #%a]\n" freg_single r1 ireg r2 coqint n | Pfstd(r1, r2, n) | Pfstd_a(r1, r2, n) -> fprintf oc " vstr %a, [%a, #%a]\n" freg r1 ireg r2 coqint n - | Pfsts(r1, r2, n) -> + | Pfsts(r1, r2, n) | Pfsts_a(r1, r2, n) -> fprintf oc " vstr %a, [%a, #%a]\n" freg_single r1 ireg r2 coqint n (* Pseudo-instructions *) | Pallocframe(sz, ofs) -> diff --git a/backend/Allocation.v b/backend/Allocation.v index cf62295dde..c8cbc87d65 100644 --- a/backend/Allocation.v +++ b/backend/Allocation.v @@ -16,7 +16,7 @@ Require Import FSets FSetAVLplus. Require Import Coqlib Ordered Maps Errors Integers Floats. Require Import AST Lattice Kildall Memdata. Require Archi. -Require Import Op Registers RTL Locations Conventions RTLtyping LTL. +Require Import Op Registers RTL Locations Conventions RTLtyping LTL LTLtyping. (** The validation algorithm used here is described in "Validating register allocation and spilling", @@ -847,8 +847,10 @@ Definition remove_equations_builtin_res (env: regenv) (res: builtin_res reg) (res': builtin_res mreg) (e: eqs) : option eqs := match res, res' with | BR r, BR r' => Some (remove_equation (Eq Full r (R r')) e) - | BR r, BR_splitlong (BR rhi) (BR rlo) => + | BR r, BR_splitlong rhi rlo => assertion (typ_eq (env r) Tlong); + assertion (subtype Tint (mreg_type rhi)); + assertion (subtype Tint (mreg_type rlo)); if mreg_eq rhi rlo then None else Some (remove_equation (Eq Low r (R rlo)) (remove_equation (Eq High r (R rhi)) e)) @@ -932,12 +934,6 @@ Definition destroyed_by_move (src dst: loc) := | _, _ => destroyed_by_op Omove end. -Definition well_typed_move (env: regenv) (dst: loc) (e: eqs) : bool := - match dst with - | R r => true - | S sl ofs ty => loc_type_compat env dst e - end. - (** Simulate the effect of a sequence of moves [mv] on a set of equations [e]. The set [e] is the equations that must hold after the sequence of moves. Return the set of equations that @@ -950,7 +946,7 @@ Fixpoint track_moves (env: regenv) (mv: moves) (e: eqs) : option eqs := | MV src dst :: mv => do e1 <- track_moves env mv e; assertion (can_undef_except dst (destroyed_by_move src dst)) e1; - assertion (well_typed_move env dst e1); + assertion (loc_type_compat env dst e1); subst_loc dst src e1 | MVmakelong src1 src2 dst :: mv => assertion (negb Archi.ptr64); @@ -1359,7 +1355,11 @@ Definition transf_function (f: RTL.function) : res LTL.function := | OK env => match regalloc f with | Error m => Error m - | OK tf => do x <- check_function f tf env; OK tf + | OK tf => + if wt_function tf then + do x <- check_function f tf env; OK tf + else + Error (msg "Ill-typed LTL code") end end. diff --git a/backend/Allocproof.v b/backend/Allocproof.v index 1804f46beb..a8c076c7fe 100644 --- a/backend/Allocproof.v +++ b/backend/Allocproof.v @@ -19,6 +19,7 @@ Require Import Coqlib Ordered Maps Errors Integers Floats. Require Import AST Linking Lattice Kildall. Require Import Values Memory Globalenvs Events Smallstep. Require Archi. +Require Import LTLtyping. Require Import Op Registers RTL Locations Conventions RTLtyping LTL. Require Import Allocation. @@ -215,7 +216,8 @@ Proof. Qed. Lemma extract_moves_sound: - forall b mv b', + forall f b mv b', + wt_bblock f b = true -> extract_moves nil b = (mv, b') -> wf_moves mv /\ b = expand_moves mv b'. Proof. @@ -226,18 +228,19 @@ Proof. { intros; split; auto. unfold wf_moves in *; rewrite Forall_forall in *. intros. apply H. rewrite <- in_rev in H0; auto. } - assert (IND: forall b accu mv b', + assert (IND: forall f b accu mv b', + wt_bblock f b = true -> extract_moves accu b = (mv, b') -> wf_moves accu -> wf_moves mv /\ expand_moves (List.rev accu) b = expand_moves mv b'). { induction b; simpl; intros. - - inv H. auto. - - destruct a; try (inv H; apply BASE; auto; fail). + - inv H0. auto. + - destruct a; try (inv H0; apply BASE; auto; fail); simpl in H; InvBooleans. + destruct (is_move_operation op args) as [arg|] eqn:E. exploit is_move_operation_correct; eauto. intros [A B]; subst. (* reg-reg move *) exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto. - inv H; apply BASE; auto. + inv H0; apply BASE; auto. + (* stack-reg move *) exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto. + (* reg-stack move *) @@ -247,7 +250,8 @@ Proof. Qed. Lemma extract_moves_ext_sound: - forall b mv b', + forall f b mv b', + wt_bblock f b = true -> extract_moves_ext nil b = (mv, b') -> wf_moves mv /\ b = expand_moves mv b'. Proof. @@ -258,13 +262,14 @@ Proof. { intros; split; auto. unfold wf_moves in *; rewrite Forall_forall in *. intros. apply H. rewrite <- in_rev in H0; auto. } - assert (IND: forall b accu mv b', + assert (IND: forall f b accu mv b', + wt_bblock f b = true -> extract_moves_ext accu b = (mv, b') -> wf_moves accu -> wf_moves mv /\ expand_moves (List.rev accu) b = expand_moves mv b'). { induction b; simpl; intros. - - inv H. auto. - - destruct a; try (inv H; apply BASE; auto; fail). + - inv H0. auto. + - destruct a; try (inv H0; apply BASE; auto; fail); simpl in H; InvBooleans. + destruct (classify_operation op args). * (* reg-reg move *) exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto. @@ -275,7 +280,7 @@ Proof. * (* highlong *) exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto. * (* default *) - inv H; apply BASE; auto. + inv H0; apply BASE; auto. + (* stack-reg move *) exploit IHb; eauto. constructor; auto. exact I. rewrite expand_moves_cons; auto. + (* reg-stack move *) @@ -292,12 +297,36 @@ Proof. destruct (peq s s0); simpl in H; inv H. exists b; auto. Qed. +Lemma wt_bblock_expand_moves_head: + forall f m i b, + wt_bblock f (expand_moves m (i :: b)) = true -> LTLtyping.wt_instr f i = true. +Proof. + intros. unfold expand_moves, wt_bblock in H. + rewrite forallb_app in H; InvBooleans. + simpl in H1; InvBooleans. auto. +Qed. + +Lemma wt_bblock_expand_moves_cons: + forall f m i b, + wt_bblock f (expand_moves m (i :: b)) = true -> wt_bblock f b = true. +Proof. + induction b; intros; auto. + simpl. unfold expand_moves, wt_bblock in H. + rewrite forallb_app in H; InvBooleans. + simpl in H1; InvBooleans. unfold wt_bblock. + rewrite H1, H3; auto. +Qed. + Ltac UseParsingLemmas := match goal with - | [ H: extract_moves nil _ = (_, _) |- _ ] => - destruct (extract_moves_sound _ _ _ H); clear H; subst; UseParsingLemmas - | [ H: extract_moves_ext nil _ = (_, _) |- _ ] => - destruct (extract_moves_ext_sound _ _ _ H); clear H; subst; UseParsingLemmas + | [ H: extract_moves nil ?b = (_, _), WT: wt_bblock _ ?b = true |- _ ] => + destruct (extract_moves_sound _ _ _ _ WT H); clear H; subst; UseParsingLemmas + | [ H: extract_moves nil ?b = (_, _), WT: wt_bblock _ (expand_moves _ (_ :: ?b)) = true |- _ ] => + apply wt_bblock_expand_moves_cons in WT; UseParsingLemmas + | [ H: extract_moves_ext nil ?b = (_, _), WT: wt_bblock _ ?b = true |- _ ] => + destruct (extract_moves_ext_sound _ _ _ _ WT H); clear H; subst; UseParsingLemmas + | [ H: extract_moves_ext nil ?b = (_, _), WT: wt_bblock _ (expand_moves _ (_ :: ?b)) = true |- _ ] => + apply wt_bblock_expand_moves_cons in WT; UseParsingLemmas | [ H: check_succ _ _ = true |- _ ] => try (discriminate H); destruct (check_succ_sound _ _ H); clear H; subst; UseParsingLemmas @@ -305,19 +334,21 @@ Ltac UseParsingLemmas := end. Lemma pair_instr_block_sound: - forall i b bsh, + forall i f b bsh, + wt_bblock f b = true -> pair_instr_block i b = Some bsh -> expand_block_shape bsh i b. Proof. - assert (OP: forall op args res s b bsh, + assert (OP: forall op args res s f b bsh, + wt_bblock f b = true -> pair_Iop_block op args res s b = Some bsh -> expand_block_shape bsh (Iop op args res s) b). { unfold pair_Iop_block; intros. MonadInv. destruct b0. MonadInv; UseParsingLemmas. destruct i; MonadInv; UseParsingLemmas. eapply ebs_op; eauto. - inv H0. eapply ebs_op_dead; eauto. } + inv H1. eapply ebs_op_dead; eauto. } - intros; destruct i; simpl in H; MonadInv; UseParsingLemmas. + intros i f b bsh WT; intros; destruct i; simpl in H; MonadInv; UseParsingLemmas. - (* nop *) econstructor; eauto. - (* op *) @@ -372,11 +403,14 @@ Lemma matching_instr_block: forall f1 f2 pc bsh i, (pair_codes f1 f2)!pc = Some bsh -> (RTL.fn_code f1)!pc = Some i -> + LTLtyping.wt_function f2 = true -> exists b, (LTL.fn_code f2)!pc = Some b /\ expand_block_shape bsh i b. Proof. intros. unfold pair_codes in H. rewrite PTree.gcombine in H; auto. rewrite H0 in H. - destruct (LTL.fn_code f2)!pc as [b|]. - exists b; split; auto. apply pair_instr_block_sound; auto. + destruct (LTL.fn_code f2)!pc as [b|] eqn:B. + exists b; split; auto. + eapply wt_function_wt_bblock in H1; eauto. + eapply pair_instr_block_sound; eauto. discriminate. Qed. @@ -399,7 +433,7 @@ Definition sel_val (k: equation_kind) (v: val) : val := is less defined than [ls l] (the value of [l] in the LTL code). *) Definition satisf (rs: regset) (ls: locset) (e: eqs) : Prop := - forall q, EqSet.In q e -> Val.lessdef (sel_val (ekind q) rs#(ereg q)) (ls (eloc q)). + forall q, EqSet.In q e -> Val.lessdef (sel_val (ekind q) rs#(ereg q)) (ls @ (eloc q)). Lemma empty_eqs_satisf: forall rs ls, satisf rs ls empty_eqs. @@ -425,7 +459,7 @@ Qed. Lemma add_equation_lessdef: forall rs ls q e, - satisf rs ls (add_equation q e) -> Val.lessdef (sel_val (ekind q) rs#(ereg q)) (ls (eloc q)). + satisf rs ls (add_equation q e) -> Val.lessdef (sel_val (ekind q) rs#(ereg q)) (ls @ (eloc q)). Proof. intros. apply H. unfold add_equation. simpl. apply EqSet.add_1. auto. Qed. @@ -686,11 +720,12 @@ Lemma loc_unconstrained_satisf: satisf rs ls (remove_equation (Eq k r l) e) -> loc_unconstrained (R mr) (remove_equation (Eq k r l) e) = true -> Val.lessdef (sel_val k rs#r) v -> + Val.has_type v (mreg_type mr) -> satisf rs (Locmap.set l v ls) e. Proof. intros; red; intros. destruct (OrderedEquation.eq_dec q (Eq k r l)). - subst q; simpl. unfold l; rewrite Locmap.gss. auto. + subst q. unfold l; rewrite Locmap.gss_typed; auto. assert (EqSet.In q (remove_equation (Eq k r l) e)). simpl. ESD.fsetdec. rewrite Locmap.gso. apply H; auto. eapply loc_unconstrained_sound; eauto. @@ -710,13 +745,14 @@ Lemma parallel_assignment_satisf: forall k r mr e rs ls v v', let l := R mr in Val.lessdef (sel_val k v) v' -> + Val.has_type v' (mreg_type mr) -> reg_loc_unconstrained r (R mr) (remove_equation (Eq k r l) e) = true -> satisf rs ls (remove_equation (Eq k r l) e) -> satisf (rs#r <- v) (Locmap.set l v' ls) e. Proof. intros; red; intros. destruct (OrderedEquation.eq_dec q (Eq k r l)). - subst q; simpl. unfold l; rewrite Regmap.gss; rewrite Locmap.gss; auto. + subst q. unfold l; rewrite Regmap.gss, Locmap.gss_typed; auto. assert (EqSet.In q (remove_equation {| ekind := k; ereg := r; eloc := l |} e)). simpl. ESD.fsetdec. exploit reg_loc_unconstrained_sound; eauto. intros [A B]. @@ -730,24 +766,25 @@ Lemma parallel_assignment_satisf_2: reg_unconstrained res e' = true -> forallb (fun l => loc_unconstrained l e') (map R (regs_of_rpair res')) = true -> Val.lessdef v v' -> + Val.has_type_rpair v' res' Val.loword Val.hiword mreg_type -> satisf (rs#res <- v) (Locmap.setpair res' v' ls) e. Proof. intros. functional inversion H. - (* One location *) subst. simpl in H2. InvBooleans. simpl. apply parallel_assignment_satisf with Full; auto. - unfold reg_loc_unconstrained. rewrite H1, H4. auto. + unfold reg_loc_unconstrained. rewrite H1, H5. auto. - (* Two 32-bit halves *) - subst. + subst. destruct H4. set (e' := remove_equation {| ekind := Low; ereg := res; eloc := R mr2 |} (remove_equation {| ekind := High; ereg := res; eloc := R mr1 |} e)) in *. - simpl in H2. InvBooleans. simpl. - red; intros. + simpl in H2. InvBooleans. + red; intros. unfold Locmap.setpair. destruct (OrderedEquation.eq_dec q (Eq Low res (R mr2))). - subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gss. + subst q. rewrite Regmap.gss, Locmap.gss_typed; auto. apply Val.loword_lessdef; auto. destruct (OrderedEquation.eq_dec q (Eq High res (R mr1))). - subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gso by auto. rewrite Locmap.gss. + subst q. rewrite Regmap.gss. rewrite Locmap.gso, Locmap.gss_typed by auto. apply Val.hiword_lessdef; auto. assert (EqSet.In q e'). unfold e', remove_equation; simpl; ESD.fsetdec. rewrite Regmap.gso. rewrite ! Locmap.gso. auto. @@ -1008,19 +1045,14 @@ Proof. exact (select_loc_h_monotone l). Qed. -Lemma well_typed_move_charact: +Lemma loc_type_compat_well_typed: forall env l e k r rs, - well_typed_move env l e = true -> + loc_type_compat env l e = true -> EqSet.In (Eq k r l) e -> wt_regset env rs -> - match l with - | R mr => True - | S sl ofs ty => Val.has_type (sel_val k rs#r) ty - end. + Val.has_type (sel_val k rs#r) (Loc.type l). Proof. - unfold well_typed_move; intros. - destruct l as [mr | sl ofs ty]. - auto. + intros. exploit loc_type_compat_charact; eauto. intros [A | A]. simpl in A. eapply Val.has_subtype; eauto. generalize (H1 r). destruct k; simpl; intros. @@ -1041,19 +1073,18 @@ Qed. Lemma subst_loc_satisf: forall env src dst rs ls e e', subst_loc dst src e = Some e' -> - well_typed_move env dst e = true -> + loc_type_compat env dst e = true -> wt_regset env rs -> + Val.has_type (ls @ src) (Loc.type dst) -> satisf rs ls e' -> - satisf rs (Locmap.set dst (ls src) ls) e. + satisf rs (Locmap.set dst (ls @ src) ls) e. Proof. intros; red; intros. exploit in_subst_loc; eauto. intros [[A B] | [A B]]. - subst dst. rewrite Locmap.gss. + subst dst. rewrite Locmap.gss_typed; auto. destruct q as [k r l]; simpl in *. - exploit well_typed_move_charact; eauto. - destruct l as [mr | sl ofs ty]; intros. - apply (H2 _ B). - apply val_lessdef_normalize; auto. apply (H2 _ B). + exploit loc_type_compat_well_typed; eauto. + intro. apply (H3 _ B). rewrite Locmap.gso; auto. Qed. @@ -1108,24 +1139,28 @@ Qed. Lemma subst_loc_part_satisf_lowlong: forall src dst rs ls e e', subst_loc_part (R dst) (R src) Low e = Some e' -> + Val.has_type (Val.loword (ls @ (R src))) (mreg_type dst) -> satisf rs ls e' -> - satisf rs (Locmap.set (R dst) (Val.loword (ls (R src))) ls) e. + satisf rs (Locmap.set (R dst) (Val.loword (ls @ (R src))) ls) e. Proof. intros; red; intros. exploit in_subst_loc_part; eauto. intros [[A [B C]] | [A B]]. - rewrite A, B. apply H0 in C. rewrite Locmap.gss. apply Val.loword_lessdef. exact C. + rewrite A, B. apply H1 in C. rewrite Locmap.gss_typed by auto. + apply Val.loword_lessdef. exact C. rewrite Locmap.gso; auto. Qed. Lemma subst_loc_part_satisf_highlong: forall src dst rs ls e e', subst_loc_part (R dst) (R src) High e = Some e' -> + Val.has_type (Val.hiword (ls @ (R src))) (mreg_type dst) -> satisf rs ls e' -> - satisf rs (Locmap.set (R dst) (Val.hiword (ls (R src))) ls) e. + satisf rs (Locmap.set (R dst) (Val.hiword (ls @ (R src))) ls) e. Proof. intros; red; intros. exploit in_subst_loc_part; eauto. intros [[A [B C]] | [A B]]. - rewrite A, B. apply H0 in C. rewrite Locmap.gss. apply Val.hiword_lessdef. exact C. + rewrite A, B. apply H1 in C. rewrite Locmap.gss_typed by auto. + apply Val.hiword_lessdef. exact C. rewrite Locmap.gso; auto. Qed. @@ -1207,7 +1242,8 @@ Lemma subst_loc_pair_satisf_makelong: wt_regset env rs -> satisf rs ls e' -> Archi.ptr64 = false -> - satisf rs (Locmap.set (R dst) (Val.longofwords (ls (R src1)) (ls (R src2))) ls) e. + Val.has_type (Val.longofwords (ls @ (R src1)) (ls @ (R src2))) (mreg_type dst) -> + satisf rs (Locmap.set (R dst) (Val.longofwords (ls @ (R src1)) (ls @ (R src2))) ls) e. Proof. intros; red; intros. exploit in_subst_loc_pair; eauto. intros [[A [B [C D]]] | [A B]]. @@ -1215,7 +1251,8 @@ Proof. assert (subtype (env (ereg q)) Tlong = true). { exploit long_type_compat_charact; eauto. intros [P|P]; auto. eelim Loc.diff_not_eq; eauto. } - rewrite Locmap.gss. simpl. rewrite <- (val_longofwords_eq_1 rs#(ereg q)). + rewrite Locmap.gss_typed by auto. + simpl. rewrite <- (val_longofwords_eq_1 rs#(ereg q)). apply Val.longofwords_lessdef. exact C. exact D. eapply Val.has_subtype; eauto. assumption. @@ -1235,10 +1272,13 @@ Qed. Lemma undef_regs_outside: forall ml ls l, - Loc.notin l (map R ml) -> undef_regs ml ls l = ls l. + Loc.notin l (map R ml) -> (undef_regs ml ls) @ l = ls @ l. Proof. - induction ml; simpl; intros. auto. - rewrite Locmap.gso. apply IHml. tauto. apply Loc.diff_sym. tauto. + induction ml; destruct ls; simpl; intros. auto. + generalize (IHml (t, t0)); intro IH. + rewrite LTL_undef_regs_Regfile_undef_regs in *. + fold (Locmap.set (R a) Vundef (Regfile.undef_regs ml t, t0)). + rewrite Locmap.gso. apply IH. tauto. apply Loc.diff_sym. tauto. Qed. Lemma can_undef_satisf: @@ -1267,19 +1307,18 @@ Qed. Lemma subst_loc_undef_satisf: forall env src dst rs ls ml e e', subst_loc dst src e = Some e' -> - well_typed_move env dst e = true -> + loc_type_compat env dst e = true -> can_undef_except dst ml e = true -> wt_regset env rs -> satisf rs ls e' -> - satisf rs (Locmap.set dst (ls src) (undef_regs ml ls)) e. + satisf rs (Locmap.set dst (ls @ src) (undef_regs ml ls)) e. Proof. intros; red; intros. exploit in_subst_loc; eauto. intros [[A B] | [A B]]. subst dst. rewrite Locmap.gss. destruct q as [k r l]; simpl in *. - exploit well_typed_move_charact; eauto. - destruct l as [mr | sl ofs ty]; intros. - apply (H3 _ B). + exploit loc_type_compat_well_typed; eauto. intros. + rewrite <- Locmap.chunk_of_loc_charact. apply val_lessdef_normalize; auto. apply (H3 _ B). rewrite Locmap.gso; auto. rewrite undef_regs_outside. eauto. eapply can_undef_except_sound; eauto. apply Loc.diff_sym; auto. @@ -1290,7 +1329,7 @@ Lemma transfer_use_def_satisf: transfer_use_def args res args' res' und e = Some e' -> satisf rs ls e' -> Val.lessdef_list rs##args (reglist ls args') /\ - (forall v v', Val.lessdef v v' -> + (forall v v', Val.lessdef v v' -> Val.has_type v' (mreg_type res') -> satisf (rs#res <- v) (Locmap.set (R res') v' (undef_regs und ls)) e). Proof. unfold transfer_use_def; intros. MonadInv. @@ -1321,7 +1360,8 @@ Lemma return_regs_agree_callee_save: forall caller callee, agree_callee_save caller (return_regs caller callee). Proof. - intros; red; intros. unfold return_regs. red in H. + intros; red; intros. + rewrite return_regs_correct. unfold Locmap.get, return_regs_spec. red in H. destruct l. rewrite H; auto. destruct sl; auto || congruence. @@ -1448,9 +1488,9 @@ Proof. red; intros. rewrite Regmap.gsspec. destruct (peq (ereg q) r1). exploit compat_left2_sound; eauto. intros [[A B] | [A B]]; rewrite A; rewrite B; simpl. - apply Val.lessdef_trans with (Val.hiword (Val.longofwords (ls l1) (ls l2))). + apply Val.lessdef_trans with (Val.hiword (Val.longofwords (ls @ l1) (ls @ l2))). apply Val.hiword_lessdef; auto. apply val_hiword_longofwords. - apply Val.lessdef_trans with (Val.loword (Val.longofwords (ls l1) (ls l2))). + apply Val.lessdef_trans with (Val.loword (Val.longofwords (ls @ l1) (ls @ l2))). apply Val.loword_lessdef; auto. apply val_loword_longofwords. eapply IHb; eauto. - (* error case *) @@ -1464,8 +1504,9 @@ Lemma call_regs_param_values: Proof. intros. unfold loc_parameters. rewrite list_map_compose. apply list_map_exten; intros. symmetry. - assert (A: forall l, loc_argument_acceptable l -> call_regs ls (parameter_of_argument l) = ls l). - { destruct l as [r | [] ofs ty]; simpl; auto; contradiction. } + assert (A: forall l, loc_argument_acceptable l -> (call_regs ls) @ (parameter_of_argument l) = ls @ l). + { intro l. rewrite call_regs_correct. + destruct l as [r | [] ofs ty]; simpl; auto; contradiction. } exploit loc_arguments_acceptable; eauto. destruct x; simpl; intros. - auto. - destruct H0; f_equal; auto. @@ -1482,6 +1523,7 @@ Proof. apply list_map_exten; intros. apply Locmap.getpair_exten; intros. assert (In l (regs_of_rpairs (loc_arguments sg))) by (eapply in_regs_of_rpairs; eauto). + rewrite return_regs_correct. exploit loc_arguments_acceptable_2; eauto. exploit H; eauto. destruct l; simpl; intros; try contradiction. rewrite H4; auto. Qed. @@ -1493,7 +1535,8 @@ Lemma find_function_tailcall: Proof. unfold ros_compatible_tailcall, find_function; intros. destruct ros as [r|id]; auto. - unfold return_regs. destruct (is_callee_save r). discriminate. auto. + rewrite return_regs_correct. + unfold Locmap.get, return_regs_spec. destruct (is_callee_save r). discriminate. auto. Qed. Lemma loadv_int64_split: @@ -1529,15 +1572,15 @@ Lemma add_equations_builtin_arg_lessdef: add_equations_builtin_arg env arg arg' e = Some e' -> satisf rs ls e' -> wt_regset env rs -> - exists v', eval_builtin_arg ge ls sp m arg' v' /\ Val.lessdef v v'. + exists v', eval_builtin_arg ge (Locmap.read ls) sp m arg' v' /\ Val.lessdef v v'. Proof. induction 1; simpl; intros e e' arg' AE SAT WT; destruct arg'; MonadInv. - exploit add_equation_lessdef; eauto. simpl; intros. - exists (ls x0); auto with barg. + exists (ls @ x0); econstructor; auto with barg; econstructor. - destruct arg'1; MonadInv. destruct arg'2; MonadInv. exploit add_equation_lessdef. eauto. simpl; intros LD1. exploit add_equation_lessdef. eapply add_equation_satisf. eauto. simpl; intros LD2. - exists (Val.longofwords (ls x0) (ls x1)); split; auto with barg. + exists (Val.longofwords (ls @ x0) (ls @ x1)); split; auto with barg. econstructor; econstructor. rewrite <- (val_longofwords_eq_2 rs#x); auto. apply Val.longofwords_lessdef; auto. rewrite <- e0; apply WT. - econstructor; eauto with barg. @@ -1575,14 +1618,14 @@ Lemma add_equations_builtin_args_lessdef: satisf rs ls e' -> wt_regset env rs -> Mem.extends m tm -> - exists vl', eval_builtin_args ge ls sp tm arg' vl' /\ Val.lessdef_list vl vl'. + exists vl', eval_builtin_args ge (Locmap.read ls) sp tm arg' vl' /\ Val.lessdef_list vl vl'. Proof. induction 1; simpl; intros; destruct arg'; MonadInv. - exists (@nil val); split; constructor. - exploit IHlist_forall2; eauto. intros (vl' & A & B). exploit add_equations_builtin_arg_lessdef; eauto. eapply add_equations_builtin_args_satisf; eauto. intros (v1' & C & D). - exploit (@eval_builtin_arg_lessdef _ ge ls ls); eauto. intros (v1'' & E & F). + exploit (@eval_builtin_arg_lessdef _ ge (Locmap.read ls) (Locmap.read ls)); eauto. intros (v1'' & E & F). exists (v1'' :: vl'); split; constructor; auto. eapply Val.lessdef_trans; eauto. Qed. @@ -1604,7 +1647,7 @@ Lemma add_equations_debug_args_eval: satisf rs ls e' -> wt_regset env rs -> Mem.extends m tm -> - exists vl', eval_builtin_args ge ls sp tm arg' vl'. + exists vl', eval_builtin_args ge (Locmap.read ls) sp tm arg' vl'. Proof. induction 1; simpl; intros; destruct arg'; MonadInv. - exists (@nil val); constructor. @@ -1613,7 +1656,7 @@ Proof. + exploit IHlist_forall2; eauto. intros (vl' & B). exploit add_equations_builtin_arg_lessdef; eauto. eapply add_equations_debug_args_satisf; eauto. intros (v1' & C & D). - exploit (@eval_builtin_arg_lessdef _ ge ls ls); eauto. intros (v1'' & E & F). + exploit (@eval_builtin_arg_lessdef _ ge (Locmap.read ls) (Locmap.read ls)); eauto. intros (v1'' & E & F). exists (v1'' :: vl'); constructor; auto. + eauto. Qed. @@ -1631,7 +1674,7 @@ Lemma add_equations_builtin_eval: external_call ef ge vargs m1 t vres m2 -> satisf rs ls e1 /\ exists vargs' vres' m2', - eval_builtin_args ge ls sp m1' args' vargs' + eval_builtin_args ge (Locmap.read ls) sp m1' args' vargs' /\ external_call ef ge vargs' m1' t vres' m2' /\ Val.lessdef vres vres' /\ Mem.extends m2 m2'. @@ -1640,7 +1683,7 @@ Proof. assert (DEFAULT: add_equations_builtin_args env args args' e1 = Some e2 -> satisf rs ls e1 /\ exists vargs' vres' m2', - eval_builtin_args ge ls sp m1' args' vargs' + eval_builtin_args ge (Locmap.read ls) sp m1' args' vargs' /\ external_call ef ge vargs' m1' t vres' m2' /\ Val.lessdef vres vres' /\ Mem.extends m2 m2'). @@ -1667,22 +1710,30 @@ Lemma parallel_set_builtin_res_satisf: forallb (fun mr => loc_unconstrained (R mr) e1) (params_of_builtin_res res') = true -> satisf rs ls e1 -> Val.lessdef v v' -> + Val.has_type_builtin_res v' res' Val.loword Val.hiword mreg_type -> satisf (regmap_setres res v rs) (Locmap.setres res' v' ls) e0. Proof. intros. rewrite forallb_forall in *. destruct res, res'; simpl in *; inv H. - apply parallel_assignment_satisf with (k := Full); auto. unfold reg_loc_unconstrained. rewrite H0 by auto. rewrite H1 by auto. auto. -- destruct res'1; try discriminate. destruct res'2; try discriminate. - rename x0 into hi; rename x1 into lo. MonadInv. destruct (mreg_eq hi lo); inv H5. - set (e' := remove_equation {| ekind := High; ereg := x; eloc := R hi |} e0) in *. +- set (e' := remove_equation {| ekind := High; ereg := x; eloc := R hi |} e0) in *. set (e'' := remove_equation {| ekind := Low; ereg := x; eloc := R lo |} e') in *. - simpl in *. red; intros. + red; intros. + assert (subtype Tint (mreg_type hi) = true /\ + subtype Tint (mreg_type lo) = true /\ + lo <> hi /\ e'' = e1). + { destruct (typ_eq (env x) Tlong), (mreg_eq hi lo), (mreg_type hi), (mreg_type lo); + try inversion H6; auto. } + decompose [and] H4. decompose [and] H5; subst. + fold (Locmap.set (R hi) (Val.hiword v') ls). + fold (Locmap.set (R lo) (Val.loword v') (Locmap.set (R hi) (Val.hiword v') ls)). destruct (OrderedEquation.eq_dec q (Eq Low x (R lo))). - subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gss. apply Val.loword_lessdef; auto. + subst q. rewrite Regmap.gss. + rewrite Locmap.gss, Val.load_result_same; auto. apply Val.loword_lessdef; auto. destruct (OrderedEquation.eq_dec q (Eq High x (R hi))). - subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gso by (red; auto). - rewrite Locmap.gss. apply Val.hiword_lessdef; auto. + subst q. rewrite Regmap.gss. rewrite Locmap.gso by (red; auto). + rewrite Locmap.gss, Val.load_result_same; auto. apply Val.hiword_lessdef; auto. assert (EqSet.In q e''). { unfold e'', e', remove_equation; simpl; ESD.fsetdec. } rewrite Regmap.gso. rewrite ! Locmap.gso. auto. @@ -1746,12 +1797,14 @@ Proof. unfold transf_function; intros. destruct (type_function f) as [env|] eqn:TY; try discriminate. destruct (regalloc f); try discriminate. + destruct (LTLtyping.wt_function f0) eqn:WT; try discriminate. destruct (check_function f f0 env) as [] eqn:?; inv H. unfold check_function in Heqr. destruct (analyze f env (pair_codes f tf)) as [an|] eqn:?; try discriminate. monadInv Heqr. destruct (check_entrypoints_aux f tf env x) as [y|] eqn:?; try discriminate. unfold check_entrypoints_aux, pair_entrypoints in Heqo0. MonadInv. + eapply wt_function_wt_bblock in WT; eauto. exploit extract_moves_ext_sound; eauto. intros [A B]. subst b. exploit check_succ_sound; eauto. intros [k EQ1]. subst b0. econstructor; eauto. eapply type_function_correct; eauto. congruence. @@ -1760,6 +1813,7 @@ Qed. Lemma invert_code: forall f env tf pc i opte e, wt_function f env -> + LTLtyping.wt_function tf = true -> (RTL.fn_code f)!pc = Some i -> transfer f env (pair_codes f tf) pc opte = OK e -> exists eafter, exists bsh, exists bb, @@ -1770,10 +1824,10 @@ Lemma invert_code: transfer_aux f env bsh eafter = Some e /\ wt_instr f env i. Proof. - intros. destruct opte as [eafter|]; simpl in H1; try discriminate. exists eafter. + intros. destruct opte as [eafter|]; simpl in H2; try discriminate. exists eafter. destruct (pair_codes f tf)!pc as [bsh|] eqn:?; try discriminate. exists bsh. exploit matching_instr_block; eauto. intros [bb [A B]]. - destruct (transfer_aux f env bsh eafter) as [e1|] eqn:?; inv H1. + destruct (transfer_aux f env bsh eafter) as [e1|] eqn:?; inv H2. exists bb. exploit wt_instr_at; eauto. tauto. Qed. @@ -1833,7 +1887,7 @@ Proof. destruct ros as [r|id]; destruct ros' as [r'|id']; simpl in H0; MonadInv. (* two regs *) exploit add_equation_lessdef; eauto. intros LD. inv LD. - eapply functions_translated; eauto. + fold (ls @ (R r')) in H3. eapply functions_translated; eauto. congruence. rewrite <- H2 in H. simpl in H. congruence. (* two symbols *) rewrite symbols_preserved. rewrite Heqo. @@ -1846,15 +1900,16 @@ Lemma exec_moves: wf_moves mv -> satisf rs ls e' -> wt_regset env rs -> + LTLtyping.wt_bblock f (expand_moves mv bb) = true -> exists ls', - star step tge (Block s f sp (expand_moves mv bb) ls m) - E0 (Block s f sp bb ls' m) + star step tge (Block s f sp (expand_moves mv bb) ls m) + E0 (Block s f sp bb ls' m) /\ satisf rs ls' e. Proof. Opaque destroyed_by_op. induction mv; simpl; intros. (* base *) -- unfold expand_moves; simpl. inv H. exists ls; split. apply star_refl. auto. +- unfold expand_moves; simpl. inv H. exists ls; split. apply star_refl. tauto. (* step *) - assert (wf_moves mv) by (inv H0; auto). destruct a; unfold expand_moves; simpl; MonadInv. @@ -1862,30 +1917,55 @@ Opaque destroyed_by_op. destruct src as [rsrc | ssrc]; destruct dst as [rdst | sdst]. * (* reg-reg *) exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto. + InvBooleans. eauto. intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto. econstructor. simpl. eauto. auto. auto. * (* reg->stack *) exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto. + InvBooleans. eauto. intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto. econstructor. simpl. eauto. auto. * (* stack->reg *) simpl in Heqb. exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto. + InvBooleans. eauto. intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto. econstructor. auto. auto. * (* stack->stack *) - inv H0. simpl in H6. contradiction. + inv H0. simpl in H8. contradiction. + (* makelong *) + assert (HT: Val.has_type (Val.longofwords (ls @ (R src1)) (ls @ (R src2))) (mreg_type dst)). + { + InvBooleans. simpl in H5. + destruct (mreg_type dst) eqn:T; auto; inversion H5. + destruct (ls @ (R src1)), (ls @ (R src2)); simpl; auto. + destruct (ls @ (R src1)), (ls @ (R src2)); simpl; auto. + } exploit IHmv; eauto. eapply subst_loc_pair_satisf_makelong; eauto. + InvBooleans. eauto. intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto. - econstructor. simpl; eauto. reflexivity. traceEq. + econstructor. simpl; eauto. destruct ls; reflexivity. traceEq. + (* lowlong *) + assert (HT: Val.has_type (Val.loword (ls @ (R src))) (mreg_type dst)). + { + simpl in H3. InvBooleans. + destruct (mreg_type dst) eqn:T; auto; inversion H5; + destruct (ls @ (R src)); simpl; auto. + } exploit IHmv; eauto. eapply subst_loc_part_satisf_lowlong; eauto. + InvBooleans. eauto. intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto. - econstructor. simpl; eauto. reflexivity. traceEq. + econstructor. simpl; eauto. destruct ls; reflexivity. traceEq. + (* highlong *) + assert (HT: Val.has_type (Val.hiword (ls @ (R src))) (mreg_type dst)). + { + simpl in H3. InvBooleans. + destruct (mreg_type dst) eqn:T; auto; inversion H5; + destruct (ls @ (R src)); simpl; auto. + } exploit IHmv; eauto. eapply subst_loc_part_satisf_highlong; eauto. + InvBooleans. eauto. intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto. - econstructor. simpl; eauto. reflexivity. traceEq. + econstructor. simpl; eauto. destruct ls; reflexivity. traceEq. Qed. (** The simulation relation *) @@ -1964,11 +2044,25 @@ Qed. Ltac UseShape := match goal with - | [ WT: wt_function _ _, CODE: (RTL.fn_code _)!_ = Some _, EQ: transfer _ _ _ _ _ = OK _ |- _ ] => - destruct (invert_code _ _ _ _ _ _ _ WT CODE EQ) as (eafter & bsh & bb & AFTER & BSH & TCODE & EBS & TR & WTI); + | [ WT: wt_function _ _, + WTTF: LTLtyping.wt_function _ = true, + CODE: (RTL.fn_code _)!_ = Some _, + EQ: transfer _ _ _ _ _ = OK _ |- _ ] => + destruct (invert_code _ _ _ _ _ _ _ WT WTTF CODE EQ) as (eafter & bsh & bb & AFTER & BSH & TCODE & EBS & TR & WTI); inv EBS; unfold transfer_aux in TR; MonadInv end. +Ltac WellTypedBlock := + match goal with + | [ T: (fn_code ?tf) ! _ = Some _ |- wt_bblock _ _ = true ] => + apply wt_function_wt_bblock in T; try eassumption; + unfold wt_bblock, expand_moves in *; try eassumption; WellTypedBlock + | [ T: forallb (LTLtyping.wt_instr _) (_ ++ _) = true |- forallb _ _ = true ] => + rewrite forallb_app in T; simpl in T; + InvBooleans; try eassumption; WellTypedBlock + | _ => idtac + end. + Remark addressing_not_long: forall env f addr args dst s r, wt_instr f env (Iload Mint64 addr args dst s) -> Archi.splitlong = true -> @@ -1984,41 +2078,104 @@ Proof. red; intros; subst r. rewrite C in H8; discriminate. Qed. +Lemma wt_transf_function: + forall (f: RTL.function) (tf: LTL.function), + transf_function f = OK tf -> + LTLtyping.wt_function tf = true. +Proof. + intros. unfold transf_function in H. + destruct (type_function f); try congruence. + destruct (regalloc f); try congruence. + destruct (LTLtyping.wt_function f0) eqn:WT; try congruence. + monadInv H; auto. +Qed. + +Lemma wt_transf_fundef: + forall (fd: RTL.fundef) (tfd: LTL.fundef), + transf_fundef fd = OK tfd -> + LTLtyping.wt_fundef tfd. +Proof. + intros. + destruct fd, tfd; simpl in *; auto; try congruence. + monadInv H. simpl; eauto using wt_transf_function. +Qed. + +Lemma wt_prog: wt_program prog. +Proof. + red; intros. + exploit list_forall2_in_left. eexact (proj1 TRANSF). eauto. + intros ([i' gd] & A & B & C). simpl in *; subst i'. + inv C. destruct f; simpl in *. +- monadInv H2. + unfold transf_function in EQ. + destruct (type_function f) as [env|] eqn:TF; try discriminate. + econstructor. eapply type_function_correct; eauto. +- constructor. +Qed. + +Corollary wt_tprog: LTLtyping.wt_program tprog. +Proof. + generalize wt_prog; unfold wt_program; intros. + unfold LTLtyping.wt_program; intros. + unfold match_prog, match_program, match_program_gen in TRANSF. + decompose [and] TRANSF. + exploit list_forall2_in_right; eauto. + intros ([i' gd] & A & B & C). + inversion C. eapply wt_transf_fundef. eauto. +Qed. + +Lemma star_step_type_preservation: + forall S1 t S2, + LTLtyping.wt_state S1 -> + star LTL.step tge S1 t S2 -> + LTLtyping.wt_state S2. +Proof. + intros. induction H0; auto. + apply IHstar. eapply LTLtyping.step_type_preservation; eauto. + eapply wt_tprog. +Qed. + (** The proof of semantic preservation is a simulation argument of the "plus" kind. *) Lemma step_simulation: forall S1 t S2, RTL.step ge S1 t S2 -> wt_state S1 -> - forall S1', match_states S1 S1' -> - exists S2', plus LTL.step tge S1' t S2' /\ match_states S2 S2'. + forall S1', match_states S1 S1' -> LTLtyping.wt_state S1' -> + exists S2', plus LTL.step tge S1' t S2' /\ LTLtyping.wt_state S2' /\ match_states S2 S2'. Proof. - induction 1; intros WT S1' MS; inv MS; try UseShape. + induction 1; intros WT S1' MS WT'; inv MS; + try assert (WTTF: LTLtyping.wt_function tf = true) by (inversion WT'; auto); + inversion WT'; subst; + try UseShape. (* nop *) -- exploit exec_moves; eauto. intros [ls1 [X Y]]. +- exploit exec_moves; eauto using wt_function_wt_bblock. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* op move *) - generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'. simpl in H0. inv H0. - exploit (exec_moves mv); eauto. intros [ls1 [X Y]]. + exploit (exec_moves mv); eauto using wt_function_wt_bblock. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. eapply subst_reg_satisf; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* op makelong *) - generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'. simpl in H0. inv H0. - exploit (exec_moves mv); eauto. intros [ls1 [X Y]]. + exploit (exec_moves mv); eauto using wt_function_wt_bblock. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. @@ -2026,12 +2183,13 @@ Proof. exploit satisf_successors; eauto. simpl; eauto. eapply subst_reg_kind_satisf_makelong. eauto. eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* op lowlong *) - generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'. simpl in H0. inv H0. - exploit (exec_moves mv); eauto. intros [ls1 [X Y]]. + exploit (exec_moves mv); eauto using wt_function_wt_bblock. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. @@ -2039,12 +2197,13 @@ Proof. exploit satisf_successors; eauto. simpl; eauto. eapply subst_reg_kind_satisf_lowlong. eauto. eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* op highlong *) - generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'. simpl in H0. inv H0. - exploit (exec_moves mv); eauto. intros [ls1 [X Y]]. + exploit (exec_moves mv); eauto using wt_function_wt_bblock. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. @@ -2052,14 +2211,29 @@ Proof. exploit satisf_successors; eauto. simpl; eauto. eapply subst_reg_kind_satisf_highlong. eauto. eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* op regular *) - generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'. - exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv1); eauto using wt_function_wt_bblock. + intros [ls1 [A1 B1]]. exploit transfer_use_def_satisf; eauto. intros [X Y]. exploit eval_operation_lessdef; eauto. intros [v' [F G]]. - exploit (exec_moves mv2); eauto. intros [ls2 [A2 B2]]. + assert (RES_TYPE: Val.has_type v' (mreg_type res')). + { + generalize (wt_function_wt_bblock tf pc _ WTTF TCODE); intro WTBB. + apply wt_bblock_expand_moves_head in WTBB. + simpl in WTBB. + destruct (is_move_operation op args') eqn:MOVE. + - apply is_move_operation_correct in MOVE. destruct MOVE; subst. simpl in *. + inversion F. fold (ls1 @ (R m0)). eapply Val.has_subtype; eauto. apply well_typed_locset. + - apply Val.has_subtype with (ty1 := snd (type_of_operation op)). + destruct (type_of_operation op); simpl; auto. + generalize (is_not_move_operation _ _ _ _ _ F MOVE); intros. + eapply type_of_operation_sound; eauto. + } + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls2 [A2 B2]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_trans. eexact A1. @@ -2067,11 +2241,13 @@ Proof. apply eval_operation_preserved. exact symbols_preserved. eauto. eapply star_right. eexact A2. constructor. eauto. eauto. eauto. traceEq. - exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]]. - econstructor; eauto. + exploit satisf_successors; eauto. simpl; eauto. + intros [enext [U V]]. + split; econstructor; eauto. (* op dead *) -- exploit exec_moves; eauto. intros [ls1 [X Y]]. +- exploit exec_moves; eauto using wt_function_wt_bblock. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. @@ -2079,16 +2255,27 @@ Proof. exploit satisf_successors. eauto. eauto. simpl; eauto. eauto. eapply reg_unconstrained_satisf; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. eapply wt_exec_Iop; eauto. (* load regular *) - generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'. - exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv1); eauto using wt_function_wt_bblock. + intros [ls1 [A1 B1]]. exploit transfer_use_def_satisf; eauto. intros [X Y]. exploit eval_addressing_lessdef; eauto. intros [a' [F G]]. exploit Mem.loadv_extends; eauto. intros [v' [P Q]]. - exploit (exec_moves mv2); eauto. intros [ls2 [A2 B2]]. + assert (DST_TYPE: Val.has_type v' (mreg_type dst')). + { + generalize (wt_function_wt_bblock tf pc _ WTTF TCODE); intro WTBB. + apply wt_bblock_expand_moves_head in WTBB. + simpl in WTBB. + eapply Val.has_subtype; eauto. + unfold Mem.loadv in P; destruct a'; try inversion P. + eapply Mem.load_type; eauto. + } + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls2 [A2 B2]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_trans. eexact A1. @@ -2097,18 +2284,27 @@ Proof. eapply star_right. eexact A2. constructor. eauto. eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* load pair *) - generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'. exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2). set (v2' := if Archi.big_endian then v2 else v1) in *. set (v1' := if Archi.big_endian then v1 else v2) in *. - exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv1); eauto using wt_function_wt_bblock. + intros [ls1 [A1 B1]]. assert (LD1: Val.lessdef_list rs##args (reglist ls1 args1')). { eapply add_equations_lessdef; eauto. } exploit eval_addressing_lessdef. eexact LD1. eauto. intros [a1' [F1 G1]]. exploit Mem.loadv_extends. eauto. eexact LOAD1. eexact G1. intros (v1'' & LOAD1' & LD2). + assert (DST_TYPE1: Val.has_type v1'' (mreg_type dst1')). + { + generalize (wt_function_wt_bblock tf pc _ WTTF TCODE); intro WTBB. + apply wt_bblock_expand_moves_head in WTBB. + unfold Mem.loadv in LOAD1'; destruct a1'; inversion LOAD1'. + eapply Val.has_subtype; eauto. + eapply Mem.load_type; eauto. + } set (ls2 := Locmap.set (R dst1') v1'' (undef_regs (destroyed_by_load Mint32 addr) ls1)). assert (SAT2: satisf (rs#dst <- v) ls2 e2). { eapply loc_unconstrained_satisf. eapply can_undef_satisf; eauto. @@ -2116,8 +2312,10 @@ Proof. eapply add_equations_satisf; eauto. assumption. rewrite Regmap.gss. apply Val.lessdef_trans with v1'; unfold sel_val; unfold kind_first_word; unfold v1'; destruct Archi.big_endian; auto. + auto. } - exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]]. + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls3 [A3 B3]]. assert (LD3: Val.lessdef_list rs##args (reglist ls3 args2')). { replace (rs##args) with ((rs#dst<-v)##args). eapply add_equations_lessdef; eauto. @@ -2130,21 +2328,33 @@ Proof. assert (LOADX: exists v2'', Mem.loadv Mint32 m' a2' = Some v2'' /\ Val.lessdef v2' v2''). { discriminate || (eapply Mem.loadv_extends; [eauto|eexact LOAD2|eexact G2]). } destruct LOADX as (v2'' & LOAD2' & LD4). + assert (DST_TYPE2: Val.has_type v2'' (mreg_type dst2')). + { + generalize (wt_function_wt_bblock tf pc _ WTTF TCODE); intro WTBB. + apply wt_bblock_expand_moves_cons in WTBB. + apply wt_bblock_expand_moves_head in WTBB. + unfold Mem.loadv in LOAD2'; destruct a2'; inversion LOAD2'. + eapply Val.has_subtype; eauto. + eapply Mem.load_type; eauto. + } set (ls4 := Locmap.set (R dst2') v2'' (undef_regs (destroyed_by_load Mint32 addr2) ls3)). assert (SAT4: satisf (rs#dst <- v) ls4 e0). { eapply loc_unconstrained_satisf. eapply can_undef_satisf; eauto. eapply add_equations_satisf; eauto. assumption. rewrite Regmap.gss. apply Val.lessdef_trans with v2'; unfold sel_val; unfold kind_second_word; unfold v2'; destruct Archi.big_endian; auto. + auto. } - exploit (exec_moves mv3); eauto. intros [ls5 [A5 B5]]. + exploit (exec_moves mv3); WellTypedBlock; eauto. + intros [ls5 [A5 B5]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_trans. eexact A1. eapply star_left. econstructor. instantiate (1 := a1'). rewrite <- F1. apply eval_addressing_preserved. exact symbols_preserved. eexact LOAD1'. instantiate (1 := ls2); auto. - eapply star_trans. eexact A3. + eapply star_trans. + eexact A3. eapply star_left. econstructor. instantiate (1 := a2'). rewrite <- F2. apply eval_addressing_preserved. exact symbols_preserved. eexact LOAD2'. instantiate (1 := ls4); auto. @@ -2152,18 +2362,29 @@ Proof. constructor. eauto. eauto. eauto. eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]]. - econstructor; eauto. + split; econstructor; eauto. (* load first word of a pair *) - generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'. exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2). set (v2' := if Archi.big_endian then v2 else v1) in *. set (v1' := if Archi.big_endian then v1 else v2) in *. - exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv1); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. assert (LD1: Val.lessdef_list rs##args (reglist ls1 args')). { eapply add_equations_lessdef; eauto. } exploit eval_addressing_lessdef. eexact LD1. eauto. intros [a1' [F1 G1]]. exploit Mem.loadv_extends. eauto. eexact LOAD1. eexact G1. intros (v1'' & LOAD1' & LD2). + assert (DST_TYPE: Val.has_type v1'' (mreg_type dst')). + { + generalize (wt_function_wt_bblock tf pc _ WTTF TCODE); intro WTBB. + apply wt_bblock_expand_moves_head in WTBB. simpl in WTBB. + unfold Mem.loadv in LOAD1'; destruct a1'; try inversion LOAD1'. + apply Mem.load_type in LOAD1'. + destruct (mreg_type dst'); auto; try congruence. + apply Val.has_subtype with (ty1 := Tint); auto. + apply Val.has_subtype with (ty1 := Tint); auto. + } set (ls2 := Locmap.set (R dst') v1'' (undef_regs (destroyed_by_load Mint32 addr) ls1)). assert (SAT2: satisf (rs#dst <- v) ls2 e0). { eapply parallel_assignment_satisf; eauto. @@ -2171,7 +2392,8 @@ Proof. unfold sel_val; unfold kind_first_word; unfold v1'; destruct Archi.big_endian; auto. eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto. } - exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]]. + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls3 [A3 B3]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_trans. eexact A1. @@ -2182,14 +2404,15 @@ Proof. constructor. eauto. eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]]. - econstructor; eauto. + split; econstructor; eauto. (* load second word of a pair *) - generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'. exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2). set (v2' := if Archi.big_endian then v2 else v1) in *. set (v1' := if Archi.big_endian then v1 else v2) in *. - exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv1); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. assert (LD1: Val.lessdef_list rs##args (reglist ls1 args')). { eapply add_equations_lessdef; eauto. } exploit eval_addressing_lessdef. eexact LD1. @@ -2198,13 +2421,24 @@ Proof. assert (LOADX: exists v2'', Mem.loadv Mint32 m' a1' = Some v2'' /\ Val.lessdef v2' v2''). { discriminate || (eapply Mem.loadv_extends; [eauto|eexact LOAD2|eexact G1]). } destruct LOADX as (v2'' & LOAD2' & LD2). + assert (DST_TYPE: Val.has_type v2'' (mreg_type dst')). + { + generalize (wt_function_wt_bblock tf pc _ WTTF TCODE); intro WTBB. + apply wt_bblock_expand_moves_head in WTBB. simpl in WTBB. + unfold Mem.loadv in LOAD2'; destruct a1'; try inversion LOAD2'. + apply Mem.load_type in LOAD2'. + destruct (mreg_type dst'); auto; try congruence. + apply Val.has_subtype with (ty1 := Tint); auto. + apply Val.has_subtype with (ty1 := Tint); auto. + } set (ls2 := Locmap.set (R dst') v2'' (undef_regs (destroyed_by_load Mint32 addr2) ls1)). assert (SAT2: satisf (rs#dst <- v) ls2 e0). { eapply parallel_assignment_satisf; eauto. apply Val.lessdef_trans with v2'; unfold sel_val; unfold kind_second_word; unfold v2'; destruct Archi.big_endian; auto. eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto. } - exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]]. + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls3 [A3 B3]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_trans. eexact A1. @@ -2215,10 +2449,11 @@ Proof. constructor. eauto. eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]]. - econstructor; eauto. + split; econstructor; eauto. (* load dead *) -- exploit exec_moves; eauto. intros [ls1 [X Y]]. +- exploit exec_moves; WellTypedBlock; eauto. + intros [ls1 [X Y]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact X. econstructor; eauto. @@ -2226,11 +2461,12 @@ Proof. exploit satisf_successors. eauto. eauto. simpl; eauto. eauto. eapply reg_unconstrained_satisf; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. eapply wt_exec_Iload; eauto. (* store *) -- exploit exec_moves; eauto. intros [ls1 [X Y]]. +- exploit exec_moves; WellTypedBlock; eauto. + intros [ls1 [X Y]]. exploit add_equations_lessdef; eauto. intros LD. simpl in LD. inv LD. exploit eval_addressing_lessdef; eauto. intros [a' [F G]]. exploit Mem.storev_extends; eauto. intros [m'' [P Q]]. @@ -2242,7 +2478,7 @@ Proof. constructor. eauto. eauto. traceEq. exploit satisf_successors; eauto. simpl; eauto. eapply can_undef_satisf; eauto. eapply add_equations_satisf; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* store 2 *) - assert (SF: Archi.ptr64 = false) by (apply Archi.splitlong_ptr32; auto). @@ -2254,7 +2490,8 @@ Proof. with (sel_val kind_second_word rs#src) by (unfold kind_second_word; destruct Archi.big_endian; reflexivity). intros [m1 [STORE1 STORE2]]. - exploit (exec_moves mv1); eauto. intros [ls1 [X Y]]. + exploit (exec_moves mv1); WellTypedBlock; eauto. + intros [ls1 [X Y]]. exploit add_equations_lessdef. eexact Heqo1. eexact Y. intros LD1. exploit add_equation_lessdef. eapply add_equations_satisf. eexact Heqo1. eexact Y. simpl. intros LD2. @@ -2267,7 +2504,8 @@ Proof. rewrite <- F1. apply eval_addressing_preserved. exact symbols_preserved. exploit Mem.storev_extends. eauto. eexact STORE1. eexact G1. eauto. intros [m1' [STORE1' EXT1]]. - exploit (exec_moves mv2); eauto. intros [ls3 [U V]]. + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls3 [U V]]. exploit add_equations_lessdef. eexact Heqo. eexact V. intros LD3. exploit add_equation_lessdef. eapply add_equations_satisf. eexact Heqo. eexact V. simpl. intros LD4. @@ -2275,7 +2513,7 @@ Proof. assert (F2': eval_addressing tge sp addr (reglist ls3 args2') = Some a2'). rewrite <- F2. apply eval_addressing_preserved. exact symbols_preserved. exploit (eval_offset_addressing tge); eauto. intros F2''. - assert (STOREX: exists m2', Mem.storev Mint32 m1' (Val.add a2' (Vint (Int.repr 4))) (ls3 (R src2')) = Some m2' /\ Mem.extends m' m2'). + assert (STOREX: exists m2', Mem.storev Mint32 m1' (Val.add a2' (Vint (Int.repr 4))) (ls3 @ (R src2')) = Some m2' /\ Mem.extends m' m2'). { try discriminate; (eapply Mem.storev_extends; [eexact EXT1 | eexact STORE2 | apply Val.add_lessdef; [eexact G2|eauto] | eauto]). } @@ -2293,13 +2531,14 @@ Proof. eapply can_undef_satisf. eauto. eapply add_equation_satisf. eapply add_equations_satisf; eauto. intros [enext [P Q]]. - econstructor; eauto. + split; econstructor; eauto. (* call *) - set (sg := RTL.funsig fd) in *. set (args' := loc_arguments sg) in *. set (res' := loc_result sg) in *. - exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv1); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. exploit find_function_translated. eauto. eauto. eapply add_equations_args_satisf; eauto. intros [tfd [E F]]. assert (SIG: funsig tfd = sg). eapply sig_function_translated; eauto. @@ -2308,6 +2547,8 @@ Proof. eapply star_right. eexact A1. econstructor; eauto. eauto. traceEq. exploit analyze_successors; eauto. simpl. left; eauto. intros [enext [U V]]. + split. apply wt_call_state; auto. constructor; auto. WellTypedBlock. + eauto using wt_transf_fundef. econstructor; eauto. econstructor; eauto. inv WTI. congruence. @@ -2317,6 +2558,7 @@ Proof. eapply add_equations_args_satisf; eauto. congruence. apply wt_regset_assign; auto. + WellTypedBlock. intros [ls2 [A2 B2]]. exists ls2; split. eapply star_right. eexact A2. constructor. traceEq. @@ -2330,7 +2572,8 @@ Proof. - set (sg := RTL.funsig fd) in *. set (args' := loc_arguments sg) in *. exploit Mem.free_parallel_extends; eauto. intros [m'' [P Q]]. - exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. exploit find_function_translated. eauto. eauto. eapply add_equations_args_satisf; eauto. intros [tfd [E F]]. assert (SIG: funsig tfd = sg). eapply sig_function_translated; eauto. @@ -2341,6 +2584,8 @@ Proof. replace (fn_stacksize tf) with (RTL.fn_stacksize f); eauto. destruct (transf_function_inv _ _ FUN); auto. eauto. traceEq. + split. + constructor. auto. eapply wt_transf_fundef; eauto. econstructor; eauto. eapply match_stackframes_change_sig; eauto. rewrite SIG. rewrite e0. decEq. destruct (transf_function_inv _ _ FUN); auto. @@ -2350,15 +2595,30 @@ Proof. rewrite SIG. inv WTI. rewrite <- H6. apply wt_regset_list; auto. (* builtin *) -- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]]. +- exploit (exec_moves mv1); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. exploit add_equations_builtin_eval; eauto. intros (C & vargs' & vres' & m'' & D & E & F & G). assert (WTRS': wt_regset env (regmap_setres res vres rs)) by (eapply wt_exec_Ibuiltin; eauto). + assert (WTBR: wt_builtin_res (proj_sig_res (ef_sig ef)) res' = true). + { + apply wt_function_wt_bblock in TCODE; auto. + unfold wt_bblock, expand_moves in *. rewrite forallb_app in TCODE. + simpl in TCODE. InvBooleans. auto. + } + exploit external_call_well_typed; eauto; intros. set (ls2 := Locmap.setres res' vres' (undef_regs (destroyed_by_builtin ef) ls1)). assert (satisf (regmap_setres res vres rs) ls2 e0). { eapply parallel_set_builtin_res_satisf; eauto. - eapply can_undef_satisf; eauto. } - exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]]. + eapply can_undef_satisf; eauto. + unfold Val.has_type_builtin_res; unfold wt_builtin_res in WTBR. + destruct res'; try eapply Val.has_subtype; eauto. + InvBooleans. split. + eapply Val.has_subtype; eauto. destruct vres'; simpl; eauto. + eapply Val.has_subtype; eauto. destruct vres'; simpl; eauto. + } + exploit (exec_moves mv2); WellTypedBlock; eauto. + intros [ls3 [A3 B3]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_trans. eexact A1. @@ -2371,10 +2631,11 @@ Proof. reflexivity. reflexivity. reflexivity. traceEq. exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* cond *) -- exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]]. +- exploit (exec_moves mv); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact A1. @@ -2384,11 +2645,12 @@ Proof. instantiate (1 := if b then ifso else ifnot). simpl. destruct b; auto. eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* jumptable *) -- exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]]. - assert (Val.lessdef (Vint n) (ls1 (R arg'))). +- exploit (exec_moves mv); WellTypedBlock; eauto. + intros [ls1 [A1 B1]]. + assert (Val.lessdef (Vint n) (ls1 @ (R arg'))). rewrite <- H0. eapply add_equation_lessdef with (q := Eq Full arg (R arg')); eauto. inv H2. econstructor; split. @@ -2399,28 +2661,32 @@ Proof. instantiate (1 := pc'). simpl. eapply list_nth_z_in; eauto. eapply can_undef_satisf. eauto. eapply add_equation_satisf; eauto. intros [enext [U V]]. - econstructor; eauto. + split; econstructor; eauto. (* return *) - destruct (transf_function_inv _ _ FUN). exploit Mem.free_parallel_extends; eauto. rewrite H10. intros [m'' [P Q]]. inv WTI; MonadInv. + (* without an argument *) - exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv); eauto using wt_function_wt_bblock. + intros [ls1 [A1 B1]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact A1. econstructor. eauto. eauto. traceEq. - simpl. econstructor; eauto. + simpl. split. econstructor; eauto. + econstructor; eauto. apply return_regs_agree_callee_save. constructor. + (* with an argument *) - exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]]. + exploit (exec_moves mv); eauto using wt_function_wt_bblock. + intros [ls1 [A1 B1]]. econstructor; split. eapply plus_left. econstructor; eauto. eapply star_right. eexact A1. econstructor. eauto. eauto. traceEq. - simpl. econstructor; eauto. rewrite <- H11. + simpl. split. econstructor; eauto. + econstructor; eauto. rewrite <- H11. replace (Locmap.getpair (map_rpair R (loc_result (RTL.fn_sig f))) (return_regs (parent_locset ts) ls1)) with (Locmap.getpair (map_rpair R (loc_result (RTL.fn_sig f))) ls1). @@ -2428,8 +2694,12 @@ Proof. rewrite H13. apply WTRS. generalize (loc_result_caller_save (RTL.fn_sig f)). destruct (loc_result (RTL.fn_sig f)); simpl. - intros A; rewrite A; auto. - intros [A B]; rewrite A, B; auto. + fold ((return_regs (parent_locset ts) ls1) @ (R r)). rewrite return_regs_correct. + unfold Locmap.read, return_regs_spec. intros A; rewrite A; auto. + fold ((return_regs (parent_locset ts) ls1) @ (R rhi)). + fold ((return_regs (parent_locset ts) ls1) @ (R rlo)). + rewrite !return_regs_correct. + unfold Locmap.read, return_regs_spec. intros [A B]; rewrite A, B; auto. apply return_regs_agree_callee_save. unfold proj_sig_res. rewrite <- H11; rewrite H13. apply WTRS. @@ -2444,6 +2714,7 @@ Proof. eapply can_undef_satisf; eauto. eapply compat_entry_satisf; eauto. rewrite call_regs_param_values. eexact ARGS. exact WTRS. + eapply wt_function_wt_bblock; eauto. intros [ls1 [A B]]. econstructor; split. eapply plus_left. econstructor; eauto. @@ -2451,7 +2722,7 @@ Proof. eapply star_right. eexact A. econstructor; eauto. eauto. eauto. traceEq. - econstructor; eauto. + split; econstructor; eauto. (* external function *) - exploit external_call_mem_extends; eauto. intros [v' [m'' [F [G [J K]]]]]. @@ -2459,16 +2730,30 @@ Proof. econstructor; split. apply plus_one. econstructor; eauto. eapply external_call_symbols_preserved with (ge1 := ge); eauto. apply senv_preserved. + exploit external_call_well_typed; eauto; intro WTRES. + split. econstructor; eauto. econstructor; eauto. - simpl. destruct (loc_result (ef_sig ef)) eqn:RES; simpl. + simpl. destruct (loc_result (ef_sig ef)) eqn:RES. + unfold Locmap.getpair, map_rpair, Locmap.setpair. rewrite Locmap.gss; auto. + generalize (loc_result_type (ef_sig ef)); intro SUBTYP. + rewrite RES in SUBTYP; simpl in SUBTYP. + exploit Val.has_subtype; eauto; intros. + rewrite Val.load_result_same; auto. generalize (loc_result_pair (ef_sig ef)); rewrite RES; intros (A & B & C & D & E). exploit external_call_well_typed; eauto. unfold proj_sig_res; rewrite B. intros WTRES'. + unfold Locmap.getpair, map_rpair, Locmap.setpair. rewrite Locmap.gss. rewrite Locmap.gso by (red; auto). rewrite Locmap.gss. + generalize (loc_result_type (ef_sig ef)); intro SUBTYP. + rewrite RES in SUBTYP; simpl in SUBTYP. + rewrite !Val.load_result_same. rewrite val_longofwords_eq_1 by auto. auto. + eapply Val.has_subtype; eauto. destruct v'; simpl; auto. + eapply Val.has_subtype; eauto. destruct v'; simpl; auto. red; intros. rewrite (AG l H0). rewrite locmap_get_set_loc_result_callee_save by auto. - unfold undef_caller_save_regs. destruct l; simpl in H0. + rewrite undef_caller_save_regs_correct. + unfold Locmap.get, undef_caller_save_regs_spec. destruct l; simpl in H0. rewrite H0; auto. destruct sl; auto; congruence. eapply external_call_well_typed; eauto. @@ -2478,6 +2763,9 @@ Proof. exploit STEPS; eauto. rewrite WTRES0; auto. intros [ls2 [A B]]. econstructor; split. eapply plus_left. constructor. eexact A. traceEq. + split. + apply star_step_type_preservation in A; eauto. + inversion WTSTK. econstructor; eauto. econstructor; eauto. apply wt_regset_assign; auto. rewrite WTRES0; auto. Qed. @@ -2489,7 +2777,7 @@ Proof. intros. inv H. exploit function_ptr_translated; eauto. intros [tf [FIND TR]]. exploit sig_function_translated; eauto. intros SIG. - exists (LTL.Callstate nil tf (Locmap.init Vundef) m0); split. + exists (LTL.Callstate nil tf Locmap.init m0); split. econstructor; eauto. eapply (Genv.init_mem_transf_partial TRANSF); eauto. rewrite symbols_preserved. @@ -2512,30 +2800,19 @@ Proof. rewrite H; auto. Qed. -Lemma wt_prog: wt_program prog. -Proof. - red; intros. - exploit list_forall2_in_left. eexact (proj1 TRANSF). eauto. - intros ([i' gd] & A & B & C). simpl in *; subst i'. - inv C. destruct f; simpl in *. -- monadInv H2. - unfold transf_function in EQ. - destruct (type_function f) as [env|] eqn:TF; try discriminate. - econstructor. eapply type_function_correct; eauto. -- constructor. -Qed. - Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (LTL.semantics tprog). Proof. - set (ms := fun s s' => wt_state s /\ match_states s s'). + set (ms := fun s s' => wt_state s /\ LTLtyping.wt_state s' /\ match_states s s'). eapply forward_simulation_plus with (match_states := ms). - apply senv_preserved. - intros. exploit initial_states_simulation; eauto. intros [st2 [A B]]. exists st2; split; auto. split; auto. apply wt_initial_state with (p := prog); auto. exact wt_prog. -- intros. destruct H. eapply final_states_simulation; eauto. -- intros. destruct H0. + split; auto. + apply LTLtyping.wt_initial_state with (prog := tprog); auto. exact wt_tprog. +- intros. destruct H as [H1 [H2 H3]]. eapply final_states_simulation; eauto. +- intros. destruct H0 as [H1 [H2 H3]]. exploit step_simulation; eauto. intros [s2' [A B]]. exists s2'; split. exact A. split. eapply subject_reduction; eauto. eexact wt_prog. eexact H. diff --git a/backend/Asmgenproof0.v b/backend/Asmgenproof0.v index 70c4323c5c..db8f6f1c60 100644 --- a/backend/Asmgenproof0.v +++ b/backend/Asmgenproof0.v @@ -80,10 +80,23 @@ Qed. Hint Resolve preg_of_not_SP preg_of_not_PC: asmgen. +Lemma load_result_offset_pc: + forall ofs rs, + Val.load_result (Preg.chunk_of PC) (Val.offset_ptr (Pregmap.get PC rs) ofs) = + Val.offset_ptr (Pregmap.get PC rs) ofs. +Proof. + intros. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tptr). apply pc_type. + destruct (Pregmap.get PC rs); simpl; auto. + unfold Tptr. destruct Archi.ptr64; auto. +Qed. + Lemma nextinstr_pc: forall rs, (nextinstr rs)#PC = Val.offset_ptr rs#PC Ptrofs.one. Proof. - intros. apply Pregmap.gss. + intros. unfold nextinstr. rewrite Pregmap.gss. + apply load_result_offset_pc. Qed. Lemma nextinstr_inv: @@ -102,17 +115,9 @@ Lemma nextinstr_set_preg: forall rs m v, (nextinstr (rs#(preg_of m) <- v))#PC = Val.offset_ptr rs#PC Ptrofs.one. Proof. - intros. unfold nextinstr. rewrite Pregmap.gss. - rewrite Pregmap.gso. auto. apply not_eq_sym. apply preg_of_not_PC. -Qed. - -Lemma undef_regs_other: - forall r rl rs, - (forall r', In r' rl -> r <> r') -> - undef_regs rl rs r = rs r. -Proof. - induction rl; simpl; intros. auto. - rewrite IHrl by auto. rewrite Pregmap.gso; auto. + intros. unfold nextinstr. rewrite Pregmap.gss, Pregmap.gso. + apply load_result_offset_pc. + auto using preg_of_not_PC. Qed. Fixpoint preg_notin (r: preg) (rl: list mreg) : Prop := @@ -138,11 +143,12 @@ Qed. Lemma undef_regs_other_2: forall r rl rs, preg_notin r rl -> - undef_regs (map preg_of rl) rs r = rs r. + (undef_regs (map preg_of rl) rs) # r = rs # r. Proof. - intros. apply undef_regs_other. intros. + intros. apply undef_regs_other. + unfold not; intro. exploit list_in_map_inv; eauto. intros [mr [A B]]. subst. - rewrite preg_notin_charact in H. auto. + rewrite preg_notin_charact in H. eapply H; eauto. Qed. (** * Agreement between Mach registers and processor registers *) @@ -150,18 +156,19 @@ Qed. Record agree (ms: Mach.regset) (sp: val) (rs: Asm.regset) : Prop := mkagree { agree_sp: rs#SP = sp; agree_sp_def: sp <> Vundef; - agree_mregs: forall r: mreg, Val.lessdef (ms r) (rs#(preg_of r)) + agree_sp_type: Val.has_type sp Tptr; + agree_mregs: forall r: mreg, Val.lessdef (regmap_get r ms) (rs#(preg_of r)) }. Lemma preg_val: - forall ms sp rs r, agree ms sp rs -> Val.lessdef (ms r) rs#(preg_of r). + forall ms sp rs r, agree ms sp rs -> Val.lessdef (regmap_get r ms) rs#(preg_of r). Proof. intros. destruct H. auto. Qed. Lemma preg_vals: forall ms sp rs, agree ms sp rs -> - forall l, Val.lessdef_list (map ms l) (map rs (map preg_of l)). + forall l, Val.lessdef_list (map (regmap_read ms) l) (map (pregmap_read rs) (map preg_of l)). Proof. induction l; simpl. constructor. constructor. eapply preg_val; eauto. auto. Qed. @@ -176,7 +183,7 @@ Lemma ireg_val: forall ms sp rs r r', agree ms sp rs -> ireg_of r = OK r' -> - Val.lessdef (ms r) rs#r'. + Val.lessdef (regmap_get r ms) rs#r'. Proof. intros. rewrite <- (ireg_of_eq _ _ H0). eapply preg_val; eauto. Qed. @@ -185,7 +192,7 @@ Lemma freg_val: forall ms sp rs r r', agree ms sp rs -> freg_of r = OK r' -> - Val.lessdef (ms r) (rs#r'). + Val.lessdef (regmap_get r ms) (rs#r'). Proof. intros. rewrite <- (freg_of_eq _ _ H0). eapply preg_val; eauto. Qed. @@ -203,27 +210,54 @@ Qed. (** Preservation of register agreement under various assignments. *) +Remark preg_of_type: + forall r, Preg.type (preg_of r) = Mreg.type r. +Proof. + destruct r; auto. +Qed. + +Remark has_type_lessdef: + forall v v' t, + Val.lessdef v v' -> + Val.has_type v' t -> + Val.has_type v t. +Proof. + intros. inversion H; subst; simpl; auto. +Qed. + Lemma agree_set_mreg: forall ms sp rs r v rs', agree ms sp rs -> Val.lessdef v (rs'#(preg_of r)) -> - (forall r', data_preg r' = true -> r' <> preg_of r -> rs'#r' = rs#r') -> + (forall r', data_preg r' = true -> r' <> (preg_of r) -> rs'#r' = rs#r') -> agree (Regmap.set r v ms) sp rs'. Proof. intros. destruct H. split; auto. - rewrite H1; auto. apply not_eq_sym. apply preg_of_not_SP. - intros. unfold Regmap.set. destruct (RegEq.eq r0 r). congruence. - rewrite H1. auto. apply preg_of_data. - red; intros; elim n. eapply preg_of_injective; eauto. + rewrite H1; auto using preg_of_not_SP. + intros. destruct (Mreg.eq_dec r0 r). + - subst. rewrite Regmap.gss. + assert (Val.has_type v (Mreg.type r)). + { generalize (Pregmap.get_has_type (preg_of r) rs'); intro. + rewrite <- preg_of_type. + eapply has_type_lessdef; eauto. } + rewrite Mreg.chunk_of_reg_type. + rewrite Val.load_result_same; auto. + - rewrite Regmap.gso by auto. + rewrite H1; auto using preg_of_injective, preg_of_data. Qed. Corollary agree_set_mreg_parallel: forall ms sp rs r v v', agree ms sp rs -> Val.lessdef v v' -> + Val.has_type v' (mreg_type r) -> agree (Regmap.set r v ms) sp (Pregmap.set (preg_of r) v' rs). Proof. - intros. eapply agree_set_mreg; eauto. rewrite Pregmap.gss; auto. intros; apply Pregmap.gso; auto. + intros. eapply agree_set_mreg; eauto. + rewrite Pregmap.gss; auto. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + destruct r; auto. + intros; apply Pregmap.gso; auto. Qed. Lemma agree_set_other: @@ -233,7 +267,8 @@ Lemma agree_set_other: agree ms sp (rs#r <- v). Proof. intros. apply agree_exten with rs. auto. - intros. apply Pregmap.gso. congruence. + intros. apply Pregmap.gso. + auto using data_diff. Qed. Lemma agree_nextinstr: @@ -247,12 +282,14 @@ Lemma agree_set_pair: forall sp p v v' ms rs, agree ms sp rs -> Val.lessdef v v' -> + Val.has_type v' (typ_rpair mreg_type p) -> agree (Mach.set_pair p v ms) sp (set_pair (map_rpair preg_of p) v' rs). Proof. intros. destruct p; simpl. - apply agree_set_mreg_parallel; auto. -- apply agree_set_mreg_parallel. apply agree_set_mreg_parallel; auto. - apply Val.hiword_lessdef; auto. apply Val.loword_lessdef; auto. +- repeat apply agree_set_mreg_parallel; auto using Val.hiword_lessdef, Val.loword_lessdef. + destruct (mreg_type_cases rhi) as [T | T], v'; rewrite T; simpl; auto. + destruct (mreg_type_cases rlo) as [T | T], v'; rewrite T; simpl; auto. Qed. Lemma agree_undef_nondata_regs: @@ -262,10 +299,14 @@ Lemma agree_undef_nondata_regs: agree ms sp (undef_regs rl rs). Proof. induction rl; simpl; intros. auto. - apply IHrl. apply agree_exten with rs; auto. - intros. apply Pregmap.gso. red; intros; subst. - assert (data_preg a = false) by auto. congruence. - intros. apply H0; auto. + apply agree_exten with rs; auto. + intros. + assert (data_preg a = false) by auto. + rewrite Pregmap.gso by congruence. + rewrite undef_regs_other; auto. + red; intros. + assert (data_preg r = false) by auto. + congruence. Qed. Lemma agree_undef_regs: @@ -309,10 +350,13 @@ Lemma agree_set_undef_mreg: (forall r', data_preg r' = true -> r' <> preg_of r -> preg_notin r' rl -> rs'#r' = rs#r') -> agree (Regmap.set r v (Mach.undef_regs rl ms)) sp rs'. Proof. - intros. apply agree_set_mreg with (rs'#(preg_of r) <- (rs#(preg_of r))); auto. - apply agree_undef_regs with rs; auto. - intros. unfold Pregmap.set. destruct (PregEq.eq r' (preg_of r)). - congruence. auto. + intros. + apply agree_set_mreg with (rs'#(preg_of r) <- (rs#(preg_of r))); auto. + apply agree_undef_regs with rs; auto. intros. + destruct (Preg.eq_dec r' (preg_of r)). + subst. rewrite Pregmap.gss, Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Pregmap.get_has_type. + rewrite Pregmap.gso; auto. intros. rewrite Pregmap.gso; auto. Qed. @@ -321,24 +365,33 @@ Lemma agree_undef_caller_save_regs: agree ms sp rs -> agree (Mach.undef_caller_save_regs ms) sp (Asm.undef_caller_save_regs rs). Proof. - intros. destruct H. unfold Mach.undef_caller_save_regs, Asm.undef_caller_save_regs; split. -- unfold proj_sumbool; rewrite dec_eq_true. auto. + intros. destruct H. split. +- rewrite undef_caller_save_regs_correct. unfold undef_caller_save_regs_spec. + unfold proj_sumbool; rewrite dec_eq_true. auto. - auto. -- intros. unfold proj_sumbool. rewrite dec_eq_false by (apply preg_of_not_SP). - destruct (in_dec preg_eq (preg_of r) (List.map preg_of (List.filter is_callee_save all_mregs))); simpl. +- auto. +- intros. + rewrite Mach.undef_caller_save_regs_correct. unfold Mach.undef_caller_save_regs_spec. + rewrite undef_caller_save_regs_correct. unfold undef_caller_save_regs_spec. + unfold proj_sumbool. rewrite dec_eq_false by (apply preg_of_not_SP). + destruct (in_dec preg_eq (preg_of r) (List.map preg_of (List.filter is_callee_save all_mregs))). + apply list_in_map_inv in i. destruct i as (mr & A & B). assert (r = mr) by (apply preg_of_injective; auto). subst mr; clear A. - apply List.filter_In in B. destruct B as [C D]. rewrite D. auto. + apply List.filter_In in B. destruct B as [C D]. rewrite D. simpl. auto. + destruct (is_callee_save r) eqn:CS; auto. elim n. apply List.in_map. apply List.filter_In. auto using all_mregs_complete. Qed. Lemma agree_change_sp: forall ms sp rs sp', - agree ms sp rs -> sp' <> Vundef -> + agree ms sp rs -> + sp' <> Vundef -> + Val.has_type sp' Tptr -> + Val.has_type sp' (Preg.type SP) -> agree ms sp' (rs#SP <- sp'). Proof. intros. inv H. split; auto. + rewrite Pregmap.gss, Preg.chunk_of_reg_type, Val.load_result_same; auto. intros. rewrite Pregmap.gso; auto with asmgen. Qed. @@ -401,22 +454,22 @@ Qed. (** Translation of arguments and results to builtins. *) Remark builtin_arg_match: - forall ge (rs: regset) sp m a v, - eval_builtin_arg ge (fun r => rs (preg_of r)) sp m a v -> - eval_builtin_arg ge rs sp m (map_builtin_arg preg_of a) v. + forall ge rs sp m a v, + eval_builtin_arg ge (fun r => rs # (preg_of r)) sp m a v -> + eval_builtin_arg ge (pregmap_read rs) sp m (map_builtin_arg preg_of a) v. Proof. - induction 1; simpl; eauto with barg. + induction 1; simpl; eauto with barg. econstructor. Qed. Lemma builtin_args_match: forall ge ms sp rs m m', agree ms sp rs -> Mem.extends m m' -> - forall al vl, eval_builtin_args ge ms sp m al vl -> - exists vl', eval_builtin_args ge rs sp m' (map (map_builtin_arg preg_of) al) vl' + forall al vl, eval_builtin_args ge (fun r => Regmap.get r ms) sp m al vl -> + exists vl', eval_builtin_args ge (pregmap_read rs) sp m' (map (map_builtin_arg preg_of) al) vl' /\ Val.lessdef_list vl vl'. Proof. induction 3; intros; simpl. exists (@nil val); split; constructor. - exploit (@eval_builtin_arg_lessdef _ ge ms (fun r => rs (preg_of r))); eauto. + exploit (@eval_builtin_arg_lessdef _ ge (fun r => Regmap.get r ms) (fun r => rs # (preg_of r))); eauto. intros; eapply preg_val; eauto. intros (v1' & A & B). destruct IHlist_forall2 as [vl' [C D]]. @@ -427,26 +480,37 @@ Lemma agree_set_res: forall res ms sp rs v v', agree ms sp rs -> Val.lessdef v v' -> + Val.has_type_builtin_res v' res Val.loword Val.hiword mreg_type -> agree (Mach.set_res res v ms) sp (Asm.set_res (map_builtin_res preg_of res) v' rs). Proof. induction res; simpl; intros. -- eapply agree_set_mreg; eauto. rewrite Pregmap.gss. auto. +- eapply agree_set_mreg; eauto. + rewrite Pregmap.gss, Preg.chunk_of_reg_type, Val.load_result_same; auto. + rewrite preg_of_type; auto. intros. apply Pregmap.gso; auto. - auto. -- apply IHres2. apply IHres1. auto. - apply Val.hiword_lessdef; auto. - apply Val.loword_lessdef; auto. +- apply agree_set_mreg with (rs := rs # (preg_of hi) <- (Val.hiword v')). + apply agree_set_mreg with (rs := rs). + auto. + rewrite Pregmap.gss, Preg.chunk_of_reg_type, Val.load_result_same; auto using Val.hiword_lessdef. + rewrite preg_of_type; tauto. + intros. apply Pregmap.gso; auto. + rewrite Pregmap.gss, Preg.chunk_of_reg_type, Val.load_result_same; auto using Val.loword_lessdef. + rewrite preg_of_type; tauto. + intros. apply Pregmap.gso; auto. Qed. Lemma set_res_other: forall r res v rs, data_preg r = false -> - set_res (map_builtin_res preg_of res) v rs r = rs r. + (set_res (map_builtin_res preg_of res) v rs) # r = rs # r. Proof. induction res; simpl; intros. - apply Pregmap.gso. red; intros; subst r. rewrite preg_of_data in H; discriminate. - auto. -- rewrite IHres2, IHres1; auto. +- rewrite !Pregmap.gso; auto. + red; intros; subst r. rewrite preg_of_data in H; discriminate. + red; intros; subst r. rewrite preg_of_data in H; discriminate. Qed. (** * Correspondence between Mach code and Asm code *) @@ -912,6 +976,7 @@ Inductive match_stack: list Mach.stackframe -> Prop := Genv.find_funct_ptr ge fb = Some (Internal f) -> transl_code_at_pc ge ra fb f c false tf tc -> sp <> Vundef -> + Val.has_type sp Tptr -> match_stack s -> match_stack (Stackframe fb sp ra c :: s). @@ -922,6 +987,13 @@ Proof. auto. Qed. +Lemma parent_sp_type: forall s, match_stack s -> Val.has_type (parent_sp s) Tptr. +Proof. + induction 1; simpl. + unfold Vnullptr, Tptr; destruct Archi.ptr64; simpl; auto. + auto. +Qed. + Lemma parent_ra_def: forall s, match_stack s -> parent_ra s <> Vundef. Proof. induction 1; simpl. @@ -929,6 +1001,13 @@ Proof. inv H0. congruence. Qed. +Lemma parent_ra_type: forall s, match_stack s -> Val.has_type (parent_ra s) Tptr. +Proof. + induction 1; simpl. + unfold Vnullptr, Tptr; destruct Archi.ptr64; simpl; auto. + inv H0. apply Val.Vptr_has_type. +Qed. + Lemma lessdef_parent_sp: forall s v, match_stack s -> Val.lessdef (parent_sp s) v -> v = parent_sp s. diff --git a/backend/Bounds.v b/backend/Bounds.v index fa69523434..9f65a30767 100644 --- a/backend/Bounds.v +++ b/backend/Bounds.v @@ -55,23 +55,23 @@ Variable b: bounds. Definition mreg_within_bounds (r: mreg) := is_callee_save r = true -> In r (used_callee_save b). -Definition slot_within_bounds (sl: slot) (ofs: Z) (ty: typ) := +Definition slot_within_bounds (sl: slot) (ofs: Z) (q: quantity) := match sl with - | Local => ofs + typesize ty <= bound_local b - | Outgoing => ofs + typesize ty <= bound_outgoing b + | Local => ofs + typesize (typ_of_quantity q) <= bound_local b + | Outgoing => ofs + typesize (typ_of_quantity q) <= bound_outgoing b | Incoming => True end. Definition instr_within_bounds (i: instruction) := match i with - | Lgetstack sl ofs ty r => slot_within_bounds sl ofs ty /\ mreg_within_bounds r - | Lsetstack r sl ofs ty => slot_within_bounds sl ofs ty + | Lgetstack sl ofs q r => slot_within_bounds sl ofs q /\ mreg_within_bounds r + | Lsetstack r sl ofs q => slot_within_bounds sl ofs q | Lop op args res => mreg_within_bounds res | Lload chunk addr args dst => mreg_within_bounds dst | Lcall sig ros => size_arguments sig <= bound_outgoing b | Lbuiltin ef args res => (forall r, In r (params_of_builtin_res res) \/ In r (destroyed_by_builtin ef) -> mreg_within_bounds r) - /\ (forall sl ofs ty, In (S sl ofs ty) (params_of_builtin_args args) -> slot_within_bounds sl ofs ty) + /\ (forall sl ofs q, In (S sl ofs q) (params_of_builtin_args args) -> slot_within_bounds sl ofs q) | _ => True end. @@ -101,8 +101,8 @@ Definition record_regs (u: RegSet.t) (rl: list mreg) : RegSet.t := Definition record_regs_of_instr (u: RegSet.t) (i: instruction) : RegSet.t := match i with - | Lgetstack sl ofs ty r => record_reg u r - | Lsetstack r sl ofs ty => record_reg u r + | Lgetstack sl ofs q r => record_reg u r + | Lsetstack r sl ofs q => record_reg u r | Lop op args res => record_reg u res | Lload chunk addr args dst => record_reg u dst | Lstore chunk addr args src => u @@ -120,17 +120,17 @@ Definition record_regs_of_instr (u: RegSet.t) (i: instruction) : RegSet.t := Definition record_regs_of_function : RegSet.t := fold_left record_regs_of_instr f.(fn_code) RegSet.empty. -Fixpoint slots_of_locs (l: list loc) : list (slot * Z * typ) := +Fixpoint slots_of_locs (l: list loc) : list (slot * Z * quantity) := match l with | nil => nil - | S sl ofs ty :: l' => (sl, ofs, ty) :: slots_of_locs l' + | S sl ofs q :: l' => (sl, ofs, q) :: slots_of_locs l' | R r :: l' => slots_of_locs l' end. -Definition slots_of_instr (i: instruction) : list (slot * Z * typ) := +Definition slots_of_instr (i: instruction) : list (slot * Z * quantity) := match i with - | Lgetstack sl ofs ty r => (sl, ofs, ty) :: nil - | Lsetstack r sl ofs ty => (sl, ofs, ty) :: nil + | Lgetstack sl ofs q r => (sl, ofs, q) :: nil + | Lsetstack r sl ofs q => (sl, ofs, q) :: nil | Lbuiltin ef args res => slots_of_locs (params_of_builtin_args args) | _ => nil end. @@ -141,17 +141,17 @@ Definition max_over_list {A: Type} (valu: A -> Z) (l: list A) : Z := Definition max_over_instrs (valu: instruction -> Z) : Z := max_over_list valu f.(fn_code). -Definition max_over_slots_of_instr (valu: slot * Z * typ -> Z) (i: instruction) : Z := +Definition max_over_slots_of_instr (valu: slot * Z * quantity -> Z) (i: instruction) : Z := max_over_list valu (slots_of_instr i). -Definition max_over_slots_of_funct (valu: slot * Z * typ -> Z) : Z := +Definition max_over_slots_of_funct (valu: slot * Z * quantity -> Z) : Z := max_over_instrs (max_over_slots_of_instr valu). -Definition local_slot (s: slot * Z * typ) := - match s with (Local, ofs, ty) => ofs + typesize ty | _ => 0 end. +Definition local_slot (s: slot * Z * quantity) := + match s with (Local, ofs, q) => ofs + typesize (typ_of_quantity q) | _ => 0 end. -Definition outgoing_slot (s: slot * Z * typ) := - match s with (Outgoing, ofs, ty) => ofs + typesize ty | _ => 0 end. +Definition outgoing_slot (s: slot * Z * quantity) := + match s with (Outgoing, ofs, q) => ofs + typesize (typ_of_quantity q) | _ => 0 end. Definition outgoing_space (i: instruction) := match i with Lcall sig _ => size_arguments sig | _ => 0 end. @@ -168,7 +168,7 @@ Proof. Qed. Lemma max_over_slots_of_funct_pos: - forall (valu: slot * Z * typ -> Z), max_over_slots_of_funct valu >= 0. + forall (valu: slot * Z * quantity -> Z), max_over_slots_of_funct valu >= 0. Proof. intros. unfold max_over_slots_of_funct. unfold max_over_instrs. apply max_over_list_pos. @@ -278,7 +278,7 @@ Qed. Definition defined_by_instr (r': mreg) (i: instruction) := match i with - | Lgetstack sl ofs ty r => r' = r + | Lgetstack sl ofs q r => r' = r | Lop op args res => r' = res | Lload chunk addr args dst => r' = dst | Lbuiltin ef args res => In r' (params_of_builtin_res res) \/ In r' (destroyed_by_builtin ef) @@ -324,7 +324,7 @@ Proof. Qed. Lemma max_over_slots_of_funct_bound: - forall (valu: slot * Z * typ -> Z) i s, + forall (valu: slot * Z * quantity -> Z) i s, In i f.(fn_code) -> In s (slots_of_instr i) -> valu s <= max_over_slots_of_funct valu. Proof. @@ -335,22 +335,22 @@ Proof. Qed. Lemma local_slot_bound: - forall i ofs ty, - In i f.(fn_code) -> In (Local, ofs, ty) (slots_of_instr i) -> - ofs + typesize ty <= bound_local function_bounds. + forall i ofs q, + In i f.(fn_code) -> In (Local, ofs, q) (slots_of_instr i) -> + ofs + typesize (typ_of_quantity q) <= bound_local function_bounds. Proof. intros. unfold function_bounds, bound_local. - change (ofs + typesize ty) with (local_slot (Local, ofs, ty)). + change (ofs + typesize (typ_of_quantity q)) with (local_slot (Local, ofs, q)). eapply max_over_slots_of_funct_bound; eauto. Qed. Lemma outgoing_slot_bound: - forall i ofs ty, - In i f.(fn_code) -> In (Outgoing, ofs, ty) (slots_of_instr i) -> - ofs + typesize ty <= bound_outgoing function_bounds. + forall i ofs q, + In i f.(fn_code) -> In (Outgoing, ofs, q) (slots_of_instr i) -> + ofs + typesize (typ_of_quantity q) <= bound_outgoing function_bounds. Proof. - intros. change (ofs + typesize ty) with (outgoing_slot (Outgoing, ofs, ty)). + intros. change (ofs + typesize (typ_of_quantity q)) with (outgoing_slot (Outgoing, ofs, q)). unfold function_bounds, bound_outgoing. apply Zmax_bound_r. eapply max_over_slots_of_funct_bound; eauto. Qed. @@ -381,8 +381,8 @@ Qed. Lemma slot_is_within_bounds: forall i, In i f.(fn_code) -> - forall sl ty ofs, In (sl, ofs, ty) (slots_of_instr i) -> - slot_within_bounds function_bounds sl ofs ty. + forall sl q ofs, In (sl, ofs, q) (slots_of_instr i) -> + slot_within_bounds function_bounds sl ofs q. Proof. intros. unfold slot_within_bounds. destruct sl. @@ -392,7 +392,7 @@ Proof. Qed. Lemma slots_of_locs_charact: - forall sl ofs ty l, In (sl, ofs, ty) (slots_of_locs l) <-> In (S sl ofs ty) l. + forall sl ofs q l, In (sl, ofs, q) (slots_of_locs l) <-> In (S sl ofs q) l. Proof. induction l; simpl; intros. tauto. diff --git a/backend/Conventions.v b/backend/Conventions.v index 989bfa05dd..ebbf41b8ef 100644 --- a/backend/Conventions.v +++ b/backend/Conventions.v @@ -101,7 +101,7 @@ Proof. unfold loc_argument_acceptable. destruct l; intros. auto. destruct sl; try contradiction. destruct H1. generalize (loc_arguments_bounded _ _ _ H0). - generalize (typesize_pos ty). omega. + generalize (typesize_pos (typ_of_quantity q)). omega. Qed. @@ -131,7 +131,7 @@ Definition callee_save_loc (l: loc) := Hint Unfold callee_save_loc. Definition agree_callee_save (ls1 ls2: Locmap.t) : Prop := - forall l, callee_save_loc l -> ls1 l = ls2 l. + forall l, callee_save_loc l -> Locmap.get l ls1 = Locmap.get l ls2. (** * Assigning result locations *) @@ -140,7 +140,7 @@ Definition agree_callee_save (ls1 ls2: Locmap.t) : Prop := Lemma locmap_get_set_loc_result: forall sg v rs l, match l with R r => is_callee_save r = true | S _ _ _ => True end -> - Locmap.setpair (loc_result sg) v rs l = rs l. + Locmap.get l (Locmap.setpair (loc_result sg) v rs) = Locmap.get l rs. Proof. intros. apply Locmap.gpo. assert (X: forall r, is_callee_save r = false -> Loc.diff l (R r)). @@ -151,7 +151,7 @@ Qed. Lemma locmap_get_set_loc_result_callee_save: forall sg v rs l, callee_save_loc l -> - Locmap.setpair (loc_result sg) v rs l = rs l. + Locmap.get l (Locmap.setpair (loc_result sg) v rs) = Locmap.get l rs. Proof. intros. apply locmap_get_set_loc_result. red in H; destruct l; auto. diff --git a/backend/Debugvar.v b/backend/Debugvar.v index 1f36103017..0cb33a8860 100644 --- a/backend/Debugvar.v +++ b/backend/Debugvar.v @@ -111,11 +111,11 @@ Fixpoint arg_no_overlap (a: builtin_arg loc) (l: loc) : bool := Definition kill (l: loc) (s: avail) : avail := List.filter (fun vi => arg_no_overlap (proj1_sig (snd vi)) l) s. -Fixpoint kill_res (r: builtin_res mreg) (s: avail) : avail := +Definition kill_res (r: builtin_res mreg) (s: avail) : avail := match r with | BR r => kill (R r) s | BR_none => s - | BR_splitlong hi lo => kill_res hi (kill_res lo s) + | BR_splitlong hi lo => kill (R hi) (kill (R lo) s) end. (** Likewise when a function call takes place. *) diff --git a/backend/Debugvarproof.v b/backend/Debugvarproof.v index d31c63ec74..5e1c6bd761 100644 --- a/backend/Debugvarproof.v +++ b/backend/Debugvarproof.v @@ -344,7 +344,7 @@ Qed. Lemma can_eval_safe_arg: forall (rs: locset) sp m (a: builtin_arg loc), - safe_builtin_arg a -> exists v, eval_builtin_arg tge rs sp m a v. + safe_builtin_arg a -> exists v, eval_builtin_arg tge (Locmap.read rs) sp m a v. Proof. induction a; simpl; intros; try contradiction; try (econstructor; now eauto with barg). @@ -369,13 +369,13 @@ Proof. econstructor. constructor. eexact E1. constructor. simpl; constructor. - simpl; auto. + destruct rs; simpl; auto. traceEq. - eapply star_step; eauto. econstructor. constructor. simpl; constructor. - simpl; auto. + destruct rs; simpl; auto. traceEq. Qed. @@ -533,7 +533,7 @@ Lemma transf_initial_states: Proof. intros. inversion H. exploit function_ptr_translated; eauto. intros [tf [A B]]. - exists (Callstate nil tf (Locmap.init Vundef) m0); split. + exists (Callstate nil tf Locmap.init m0); split. econstructor; eauto. eapply (Genv.init_mem_transf_partial TRANSF); eauto. rewrite (match_program_main TRANSF), symbols_preserved. auto. rewrite <- H3. apply sig_preserved. auto. diff --git a/backend/FragmentBlock.v b/backend/FragmentBlock.v new file mode 100644 index 0000000000..81fffac880 --- /dev/null +++ b/backend/FragmentBlock.v @@ -0,0 +1,343 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Gergö Barany, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +Require Import OrderedType. +Require Import Coqlib. +Require Import Decidableplus. +Require Import Maps. +Require Import Ordered. +Require Import AST. +Require Import Values. +Require Import Memdata. +Require Import Memory. +Require Export Machregs. + +Open Scope Z_scope. + +(** * Memory-like blocks storing value fragments *) + +(** The register file and the stack are represented as instances of the module + below, which is a basis for defining mappings from locations to values. This + basic module implements a word-addressed memory block that can be read in + words or doublewords (that is, as Q32 or Q64 quantities). Accesses use a word + offset into the block, which is internally scaled to byte addresses. *) + +Module FragBlock. + + Definition t := ZMap.t memval. + + Definition init := ZMap.init Undef. + + (* A register's address: The index of its first byte. *) + Definition addr (ofs: Z) : Z := ofs * 4. + + (* The size of a quantity in bytes. *) + Definition quantity_size (q: quantity) : Z := Z.of_nat (size_quantity_nat q). + + (* The address one byte past the end of a value at [ofs] of kind [q]. The next + nonoverlapping value may start here. *) + Definition next_addr (ofs: Z) (q: quantity) : Z := addr ofs + quantity_size q. + + Definition get_bytes (ofs: Z) (q: quantity) (fb: t) : list memval := + Mem.getN (size_quantity_nat q) (addr ofs) fb. + + Definition get (ofs: Z) (q: quantity) (fb: t) : val := + decode_val (chunk_of_quantity q) (get_bytes ofs q fb). + + Definition set_bytes (ofs: Z) (q: quantity) (bytes: list memval) (fb: t) : t := + Mem.setN (firstn (size_quantity_nat q) bytes) (addr ofs) fb. + + Definition set (ofs: Z) (q: quantity) (v: val) (fb: t) : t := + set_bytes ofs q (encode_val (chunk_of_quantity q) v) fb. + + (* Update the [old] block with copies of the fragments from [new] at the + locations in [locs] specified as pairs of an offset and a quantity. *) + Fixpoint override (locs: list (Z * quantity)) (new old: t) : t := + match locs with + | nil => old + | (ofs, q) :: locs' => set_bytes ofs q (get_bytes ofs q new) (override locs' new old) + end. + + Fixpoint undef_locs (locs: list (Z * quantity)) (fb: t) : t := + match locs with + | nil => fb + | (ofs, q) :: locs' => set ofs q Vundef (undef_locs locs' fb) + end. + + Lemma quantity_size_lt: + forall q q', quantity_size q < quantity_size q' -> q = Q32 /\ q' = Q64. + Proof. + intros. destruct q, q'; try inv H. auto. + Qed. + + Lemma gss_bytes: + forall ofs q bs fb, + length bs = size_quantity_nat q -> + get_bytes ofs q (set_bytes ofs q bs fb) = bs. + Proof. + intros. unfold get_bytes, set_bytes. + rewrite <- H, firstn_all, Mem.getN_setN_same. reflexivity. + Qed. + + Lemma gss: + forall ofs q v fb, + get ofs q (set ofs q v fb) = Val.load_result (chunk_of_quantity q) v. + Proof. + intros. unfold get, set. + rewrite gss_bytes; auto using encode_val_quantity_size. + apply decode_encode_val_q. + Qed. + + Lemma gso_bytes: + forall ofs q ofs' q' bs fb, + next_addr ofs q <= addr ofs' \/ next_addr ofs' q' <= addr ofs -> + length bs = size_quantity_nat q' -> + get_bytes ofs q (set_bytes ofs' q' bs fb) = get_bytes ofs q fb. + Proof. + intros until fb. intros H LEN. unfold get_bytes, set_bytes. + rewrite Mem.getN_setN_outside; auto. + unfold next_addr, addr, quantity_size in *. + destruct H. + - left; omega. + - right. rewrite <- LEN, firstn_all. omega. + Qed. + + Lemma gso: + forall ofs q ofs' q' v fb, + next_addr ofs q <= addr ofs' \/ next_addr ofs' q' <= addr ofs -> + get ofs q (set ofs' q' v fb) = get ofs q fb. + Proof. + intros. unfold get, set. + rewrite gso_bytes; auto using encode_val_quantity_size. + Qed. + + Lemma gu_set_inside: + forall ofs q ofs' q' v fb, + addr ofs <= addr ofs' /\ addr ofs' < next_addr ofs q -> + quantity_size q' < quantity_size q -> + get ofs q (set ofs' q' v fb) = Vundef. + Proof. + intros until fb. intros (LB & UB) QLT. + exploit quantity_size_lt; eauto; intros [Q' Q]; subst. + + assert (OFS: ofs = ofs' \/ ofs + 1 = ofs'). + { destruct (Zle_lt_or_eq _ _ LB). + - right. unfold next_addr, addr, quantity_size in *. + simpl in *; omega. + - left. unfold addr in H. omega. } + + unfold get, set, get_bytes, set_bytes. + simpl size_quantity_nat. + change 8%nat with (4 + 4)%nat. + rewrite Mem.getN_concat. + replace 4%nat with (length (encode_val Many32 v)) by apply encode_val_length. + rewrite firstn_all. + simpl chunk_of_quantity. rewrite encode_val_any32. + + destruct OFS as [EQ | LT]. + - subst ofs. + + rewrite Mem.getN_setN_same. unfold inj_value. + destruct (Val.eq v Vundef). simpl; auto. + destruct (fits_quantity_dec v Q32). simpl. + erewrite decode_val_diff_q with (q := Q32); simpl; eauto. congruence. + simpl. auto. + + - subst ofs'. + assert (A: addr ofs + Z.of_nat (length (inj_value Q32 v)) = addr (ofs + 1)). + { rewrite inj_value_length. unfold addr. simpl. omega. } + rewrite A, Mem.getN_setN_same. + unfold inj_value. + destruct (Val.eq v Vundef). + rewrite decode_val_undef. auto. apply in_app. right. simpl; auto. + destruct (fits_quantity_dec v Q32). + erewrite decode_val_diff_q with (q := Q32); simpl; eauto. congruence. + intuition eauto. + rewrite decode_val_undef. auto. apply in_app. right. simpl; auto. + Qed. + + Lemma gu_partial_overlap: + forall ofs q ofs' q' v fb, + addr ofs < addr ofs' /\ addr ofs' < next_addr ofs q -> + get ofs q (set ofs' q' v fb) = Vundef. + Proof. + intros until fb. intros (LB & UB). + + assert (AUX: q = Q64 /\ ofs + 1 = ofs'). + { unfold next_addr, addr, quantity_size in *. + destruct q; simpl in *; split; auto; omega. } + destruct AUX as [Q OFS]. subst q. subst ofs'. + + destruct q'. + rewrite gu_set_inside. auto. omega. unfold quantity_size. simpl. omega. + + unfold get, set, get_bytes, set_bytes. + simpl size_quantity_nat. rewrite encode_val_any64. + change 8%nat with (4 + 4)%nat. + rewrite Mem.getN_concat. + rewrite Mem.getN_setN_outside by (left; unfold addr; simpl; omega). + replace (4 + 4)%nat with (length (inj_value Q64 v)) by apply inj_value_length. + rewrite firstn_all. + + assert (A: addr ofs + Z.of_nat 4 = addr (ofs + 1)). + { unfold addr. simpl. omega. } + rewrite A, Mem.getN_setN_prefix. + + simpl chunk_of_quantity. unfold decode_val. + unfold inj_value. + destruct (Val.eq v Vundef). + + rewrite proj_bytes_undef, proj_value_undef; simpl; intuition auto using in_app. + rewrite pred_dec_true by apply any_fits_quantity_Q64. + erewrite proj_bytes_fragment by (simpl; intuition auto). + erewrite proj_value_misaligned with (n := 4%nat); simpl; eauto. + rewrite inj_value_length; simpl; auto. + Qed. + + Lemma proj_bytes_firstn_inj_value_mismatch: + forall n q v, + n <> 0%nat -> + proj_bytes (firstn n (inj_value q v)) = None. + Proof. + intros. + unfold inj_value. + destruct (Val.eq v Vundef), (fits_quantity_dec v q), q, n; simpl; congruence. + Qed. + + Lemma proj_value_firstn_inj_value_mismatch: + forall q q' v, + q <> q' -> + proj_value q (firstn (size_quantity_nat q) (inj_value q' v)) = Vundef. + Proof. + intros. destruct q, q'; try contradiction. + - simpl size_quantity_nat. + unfold inj_value. + destruct (Val.eq v Vundef); auto. + rewrite pred_dec_true by apply any_fits_quantity_Q64. + erewrite proj_value_diff_q; simpl; eauto. + - simpl size_quantity_nat. + unfold inj_value. + destruct (Val.eq v Vundef); auto. + destruct (fits_quantity_dec v Q32). + + erewrite proj_value_diff_q; simpl; eauto. + + simpl; auto. + Qed. + + Lemma gu_overlap: + forall ofs q ofs' q' v fb, + (ofs, q) <> (ofs', q') -> + next_addr ofs q > addr ofs' -> + next_addr ofs' q' > addr ofs -> + get ofs q (set ofs' q' v fb) = Vundef. + Proof. + intros until fb. intros DIFF H H'. + + unfold next_addr, addr, quantity_size in *. + + destruct (Z.eq_dec ofs ofs'). + - subst ofs. + destruct (quantity_eq q q'). + + subst q. contradiction. + + unfold get, set, get_bytes, set_bytes. + rewrite encode_val_q, firstn_inj_value. + destruct q, q'; try contradiction. + * rewrite Mem.getN_setN_prefix. + simpl chunk_of_quantity. unfold decode_val. + rewrite proj_bytes_firstn_inj_value_mismatch by (simpl; auto). + rewrite proj_value_firstn_inj_value_mismatch; auto. + rewrite inj_value_length; simpl; auto. + * simpl size_quantity_nat. + replace 8%nat with (4 + 4)%nat. + rewrite Mem.getN_concat. + unfold inj_value. + destruct (Val.eq v Vundef). + erewrite decode_val_undef; auto. + change 4%nat with (length (list_repeat (size_quantity_nat Q32) Undef)). + rewrite Mem.getN_setN_same; simpl; auto. + destruct (fits_quantity_dec v Q32). + erewrite decode_val_diff_q; eauto. + change 4%nat with (length (inj_value_rec (size_quantity_nat Q32) v Q32)). + rewrite Mem.getN_setN_same; simpl; auto. + erewrite decode_val_undef; auto. + change 4%nat with (length (list_repeat (size_quantity_nat Q32) Undef)). + rewrite Mem.getN_setN_same; simpl; auto. + simpl; auto. + - destruct (zlt (addr ofs) (addr ofs')). + + apply gu_partial_overlap. + unfold next_addr, addr, quantity_size in *. + omega. + + unfold get, set, get_bytes, set_bytes. + rewrite encode_val_q, firstn_inj_value. + assert (q' = Q64). + { destruct q'; auto. simpl in H'. unfold addr in g. omega. } + subst q'. + assert (ofs = ofs' + 1). + { simpl in H'. unfold addr in g. omega. } + subst ofs. unfold addr. + destruct q. + * simpl size_quantity_nat. + replace ((ofs' + 1) * 4) with (ofs' * 4 + Z.of_nat 4) by (simpl; omega). + rewrite Mem.getN_setN_suffix. + unfold inj_value. + destruct (Val.eq v Vundef); auto. + rewrite pred_dec_true by apply any_fits_quantity_Q64. + erewrite decode_val_diff_q; simpl; eauto. congruence. + rewrite inj_value_length; simpl; auto. + * simpl size_quantity_nat. + replace ((ofs' + 1) * 4) with (ofs' * 4 + Z.of_nat 4) by (simpl; omega). + unfold inj_value. + destruct (Val.eq v Vundef). + ** erewrite decode_val_undef; auto. + simpl. left. + replace (ofs' * 4 + 1 + 1 + 1 + 1) with (ofs' * 4 + 4) by omega. + rewrite 3! ZMap.gso by (apply Z.lt_neq; omega). + rewrite ZMap.gss; auto. + ** rewrite pred_dec_true by apply any_fits_quantity_Q64. + simpl chunk_of_quantity; unfold decode_val. + erewrite proj_bytes_fragment. + erewrite proj_value_misaligned with (n := 0%nat). auto. + simpl. + replace (ofs' * 4 + 1 + 1 + 1 + 1) with (ofs' * 4 + 4) by omega. + rewrite 3! ZMap.gso by (apply Z.lt_neq; omega). + rewrite ZMap.gss; auto. + simpl. auto. + simpl. + replace (ofs' * 4 + 1 + 1 + 1 + 1) with (ofs' * 4 + 4) by omega. + rewrite 3! ZMap.gso by (apply Z.lt_neq; omega). + rewrite ZMap.gss; auto. + Qed. + + Lemma gi: + forall ofs q, get ofs q init = Vundef. + Proof. + intros. unfold init, get, get_bytes. + destruct q; simpl; rewrite ZMap.gi; apply decode_val_hd_undef. + Qed. + + Lemma get_has_type: + forall ofs q fb, Val.has_type (get ofs q fb) (typ_of_quantity q). + Proof. + intros. unfold get. + generalize (decode_val_type (chunk_of_quantity q)); intro. + destruct q; auto. + Qed. + + Lemma get_fits_quantity: + forall ofs q fb, fits_quantity (get ofs q fb) q. + Proof. + intros. + generalize (get_has_type ofs q fb); intro. + apply has_type_fits_quantity in H. + destruct q; auto. + Qed. + +End FragBlock. diff --git a/backend/IRC.ml b/backend/IRC.ml index 43955897d9..17f0fe011d 100644 --- a/backend/IRC.ml +++ b/backend/IRC.ml @@ -13,6 +13,7 @@ open Printf open Camlcoq open AST +open Memdata open Registers open Machregs open Locations @@ -105,12 +106,12 @@ let name_of_loc = function | None -> "fixed-reg" | Some s -> s end - | S (Local, ofs, ty) -> - sprintf "L%c%ld" (PrintXTL.short_name_of_type ty) (camlint_of_coqint ofs) - | S (Incoming, ofs, ty) -> - sprintf "I%c%ld" (PrintXTL.short_name_of_type ty) (camlint_of_coqint ofs) - | S (Outgoing, ofs, ty) -> - sprintf "O%c%ld" (PrintXTL.short_name_of_type ty) (camlint_of_coqint ofs) + | S (Local, ofs, q) -> + sprintf "L%c%ld" (PrintXTL.short_name_of_quantity q) (camlint_of_coqint ofs) + | S (Incoming, ofs, q) -> + sprintf "I%c%ld" (PrintXTL.short_name_of_quantity q) (camlint_of_coqint ofs) + | S (Outgoing, ofs, q) -> + sprintf "O%c%ld" (PrintXTL.short_name_of_quantity q) (camlint_of_coqint ofs) let name_of_node n = match n.var with @@ -240,14 +241,15 @@ type graph = { let class_of_type = function | Tint | Tlong -> 0 | Tfloat | Tsingle -> 1 - | Tany32 | Tany64 -> assert false + | Tany32 -> 0 + | Tany64 -> 1 let class_of_reg r = if Conventions1.is_float_reg r then 1 else 0 let class_of_loc = function | R r -> class_of_reg r - | S(_, _, ty) -> class_of_type ty + | S(_, _, q) -> class_of_type (typ_of_quantity q) let no_spill_class = 2 @@ -842,13 +844,13 @@ let find_slot conflicts typ = let al = Z.to_int (Locations.typealign typ) in let rec find curr = function | [] -> - S(Local, Z.of_uint curr, typ) - | S(Local, ofs, typ') :: l -> + S(Local, Z.of_uint curr, quantity_of_typ typ) + | S(Local, ofs, q') :: l -> let ofs = Z.to_int ofs in if curr + sz <= ofs then - S(Local, Z.of_uint curr, typ) + S(Local, Z.of_uint curr, quantity_of_typ typ) else begin - let sz' = Z.to_int (Locations.typesize typ') in + let sz' = Z.to_int (Locations.typesize (typ_of_quantity q')) in let ofs' = align (ofs + sz') al in find (if ofs' <= curr then curr else ofs') l end diff --git a/backend/LTL.v b/backend/LTL.v index 5e7eec8c59..1072f058b2 100644 --- a/backend/LTL.v +++ b/backend/LTL.v @@ -30,8 +30,8 @@ Definition node := positive. Inductive instruction: Type := | Lop (op: operation) (args: list mreg) (res: mreg) | Lload (chunk: memory_chunk) (addr: addressing) (args: list mreg) (dst: mreg) - | Lgetstack (sl: slot) (ofs: Z) (ty: typ) (dst: mreg) - | Lsetstack (src: mreg) (sl: slot) (ofs: Z) (ty: typ) + | Lgetstack (sl: slot) (ofs: Z) (q: quantity) (dst: mreg) + | Lsetstack (src: mreg) (sl: slot) (ofs: Z) (q: quantity) | Lstore (chunk: memory_chunk) (addr: addressing) (args: list mreg) (src: mreg) | Lcall (sg: signature) (ros: mreg + ident) | Ltailcall (sg: signature) (ros: mreg + ident) @@ -66,6 +66,9 @@ Definition funsig (fd: fundef) := Definition genv := Genv.t fundef unit. Definition locset := Locmap.t. +Module Regfile := Locmap.Regfile. + +Open Scope ltl. (** Calling conventions are reflected at the level of location sets (environments mapping locations to values) by the following two @@ -78,9 +81,14 @@ Definition locset := Locmap.t. values as the corresponding outgoing stack slots (used for argument passing) in the caller. - Local and outgoing stack slots are initialized to undefined values. + + Location sets are not really represented as functions, but we would + like them to behave as such. We first give a functional specification + of the behavior of [call_regs], then the actual definition. The + correspondence is proved below. *) -Definition call_regs (caller: locset) : locset := +Definition call_regs_spec (caller: loc -> val) : (loc -> val) := fun (l: loc) => match l with | R r => caller (R r) @@ -89,6 +97,18 @@ Definition call_regs (caller: locset) : locset := | S Outgoing ofs ty => Vundef end. +Definition call_regs (caller: locset) : locset := + match caller with + | (caller_rf, caller_stack) => + (caller_rf, + fun s: slot => + match s with + | Local => ZMap.init Undef + | Incoming => caller_stack Outgoing + | Outgoing => ZMap.init Undef + end) + end. + (** [return_regs caller callee] returns the location set after a call instruction, as a function of the location set [caller] of the caller before the call instruction and of the location @@ -99,9 +119,12 @@ Definition call_regs (caller: locset) : locset := - Local and Incoming stack slots have the same values as in the caller. - Outgoing stack slots are set to Vundef to reflect the fact that they may have been changed by the callee. + + As for [call_regs], we give a functional specification and the real + implementation and prove the correspondence below. *) -Definition return_regs (caller callee: locset) : locset := +Definition return_regs_spec (caller callee: loc -> val) : (loc -> val) := fun (l: loc) => match l with | R r => if is_callee_save r then caller (R r) else callee (R r) @@ -109,11 +132,22 @@ Definition return_regs (caller callee: locset) : locset := | S sl ofs ty => caller (S sl ofs ty) end. +Definition return_regs (caller callee: locset) : locset := + match caller, callee with + | (caller_rf, caller_stack), (callee_rf, _) => + (Regfile.override destroyed_at_call callee_rf caller_rf, + fun s: slot => + match s with + | Outgoing => ZMap.init Undef + | s => caller_stack s + end) + end. + (** [undef_caller_save_regs ls] models the effect of calling an external function: caller-save registers and outgoing locations can change unpredictably, hence we set them to [Vundef]. *) -Definition undef_caller_save_regs (ls: locset) : locset := +Definition undef_caller_save_regs_spec (ls: loc -> val) : (loc -> val) := fun (l: loc) => match l with | R r => if is_callee_save r then ls (R r) else Vundef @@ -121,6 +155,17 @@ Definition undef_caller_save_regs (ls: locset) : locset := | S sl ofs ty => ls (S sl ofs ty) end. +Definition undef_caller_save_regs (ls: locset) : locset := + match ls with + | (rf, stack) => + (Regfile.undef_regs destroyed_at_call rf, + fun s: slot => + match s with + | Outgoing => ZMap.init Undef + | s => stack s + end) + end. + (** LTL execution states. *) Inductive stackframe : Type := @@ -166,7 +211,7 @@ Section RELSEM. Variable ge: genv. Definition reglist (rs: locset) (rl: list mreg) : list val := - List.map (fun r => rs (R r)) rl. + List.map (fun r => rs @ (R r)) rl. Fixpoint undef_regs (rl: list mreg) (rs: locset) : locset := match rl with @@ -182,7 +227,7 @@ Definition destroyed_by_getstack (s: slot): list mreg := Definition find_function (ros: mreg + ident) (rs: locset) : option fundef := match ros with - | inl r => Genv.find_funct ge (rs (R r)) + | inl r => Genv.find_funct ge (rs @ (R r)) | inr symb => match Genv.find_symbol ge symb with | None => None @@ -195,7 +240,7 @@ Definition find_function (ros: mreg + ident) (rs: locset) : option fundef := Definition parent_locset (stack: list stackframe) : locset := match stack with - | nil => Locmap.init Vundef + | nil => Locmap.init | Stackframe f sp ls bb :: stack' => ls end. @@ -215,17 +260,17 @@ Inductive step: state -> trace -> state -> Prop := rs' = Locmap.set (R dst) v (undef_regs (destroyed_by_load chunk addr) rs) -> step (Block s f sp (Lload chunk addr args dst :: bb) rs m) E0 (Block s f sp bb rs' m) - | exec_Lgetstack: forall s f sp sl ofs ty dst bb rs m rs', - rs' = Locmap.set (R dst) (rs (S sl ofs ty)) (undef_regs (destroyed_by_getstack sl) rs) -> - step (Block s f sp (Lgetstack sl ofs ty dst :: bb) rs m) + | exec_Lgetstack: forall s f sp sl ofs q dst bb rs m rs', + rs' = Locmap.set (R dst) (rs @ (S sl ofs q)) (undef_regs (destroyed_by_getstack sl) rs) -> + step (Block s f sp (Lgetstack sl ofs q dst :: bb) rs m) E0 (Block s f sp bb rs' m) - | exec_Lsetstack: forall s f sp src sl ofs ty bb rs m rs', - rs' = Locmap.set (S sl ofs ty) (rs (R src)) (undef_regs (destroyed_by_setstack ty) rs) -> - step (Block s f sp (Lsetstack src sl ofs ty :: bb) rs m) + | exec_Lsetstack: forall s f sp src sl ofs q bb rs m rs', + rs' = Locmap.set (S sl ofs q) (rs @ (R src)) (undef_regs (destroyed_by_setstack q) rs) -> + step (Block s f sp (Lsetstack src sl ofs q :: bb) rs m) E0 (Block s f sp bb rs' m) | exec_Lstore: forall s f sp chunk addr args src bb rs m a rs' m', eval_addressing ge sp addr (reglist rs args) = Some a -> - Mem.storev chunk m a (rs (R src)) = Some m' -> + Mem.storev chunk m a (rs @ (R src)) = Some m' -> rs' = undef_regs (destroyed_by_store chunk addr) rs -> step (Block s f sp (Lstore chunk addr args src :: bb) rs m) E0 (Block s f sp bb rs' m') @@ -242,7 +287,7 @@ Inductive step: state -> trace -> state -> Prop := step (Block s f (Vptr sp Ptrofs.zero) (Ltailcall sig ros :: bb) rs m) E0 (Callstate s fd rs' m') | exec_Lbuiltin: forall s f sp ef args res bb rs m vargs t vres rs' m', - eval_builtin_args ge rs sp m args vargs -> + eval_builtin_args ge (Locmap.read rs) sp m args vargs -> external_call ef ge vargs m t vres m' -> rs' = Locmap.setres res vres (undef_regs (destroyed_by_builtin ef) rs) -> step (Block s f sp (Lbuiltin ef args res :: bb) rs m) @@ -257,7 +302,7 @@ Inductive step: state -> trace -> state -> Prop := step (Block s f sp (Lcond cond args pc1 pc2 :: 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 -> + rs @ (R arg) = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc -> rs' = undef_regs (destroyed_by_jumptable) rs -> step (Block s f sp (Ljumptable arg tbl :: bb) rs m) @@ -295,7 +340,7 @@ Inductive initial_state (p: program): state -> Prop := Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> funsig f = signature_main -> - initial_state p (Callstate nil f (Locmap.init Vundef) m0). + initial_state p (Callstate nil f Locmap.init m0). Inductive final_state: state -> int -> Prop := | final_state_intro: forall rs m retcode, @@ -320,3 +365,67 @@ Fixpoint successors_block (b: bblock) : list node := | Lreturn :: _ => nil | instr :: b' => successors_block b' end. + +(** * General properties of auxiliary definitions for LTL *) + +Lemma call_regs_correct: + forall (l: loc) (caller: locset), + (call_regs caller) @ l = call_regs_spec (Locmap.read caller) l. +Proof. + intros. destruct l, caller as [caller_rf caller_stack]. + - reflexivity. + - simpl. unfold Stack.get. destruct sl; auto using FragBlock.gi. +Qed. + +Local Opaque all_mregs. + +Lemma in_destroyed_at_call: + forall r, In r destroyed_at_call -> is_callee_save r = false. +Proof. + intros. unfold destroyed_at_call in H. + apply filter_In in H. apply negb_true_iff. tauto. +Qed. + +Lemma notin_destroyed_at_call: + forall r, ~ In r destroyed_at_call -> is_callee_save r = true. +Proof. + intros. + apply negb_false_iff. apply not_true_iff_false. + contradict H. + apply filter_In. split; auto using all_mregs_complete. +Qed. + +Lemma return_regs_correct: + forall (l: loc) (caller callee: locset), + (return_regs caller callee) @ l = return_regs_spec (Locmap.read caller) (Locmap.read callee) l. +Proof. + intros. destruct l, caller, callee. + - unfold Locmap.get, return_regs, return_regs_spec. + destruct (In_dec mreg_eq r destroyed_at_call). + + rewrite Regfile.override_in; auto. + rewrite (in_destroyed_at_call r); auto. + + rewrite Regfile.override_notin; auto. + rewrite (notin_destroyed_at_call r); auto. + - simpl. unfold Stack.get. destruct sl; auto using FragBlock.gi. +Qed. + +Lemma undef_caller_save_regs_correct: + forall (l: loc) (ls: locset), + (undef_caller_save_regs ls) @ l = undef_caller_save_regs_spec (Locmap.read ls) l. +Proof. + intros. destruct l, ls. + - unfold Locmap.get, undef_caller_save_regs, undef_caller_save_regs_spec. + destruct (In_dec mreg_eq r destroyed_at_call). + + rewrite Regfile.undef_regs_in; auto. + rewrite (in_destroyed_at_call r); auto. + + rewrite Regfile.undef_regs_notin; auto. + rewrite (notin_destroyed_at_call r); auto. + - simpl. unfold Stack.get. destruct sl; auto using FragBlock.gi. +Qed. + +Lemma LTL_undef_regs_Regfile_undef_regs: + forall rl rf stack, undef_regs rl (rf, stack) = (Regfile.undef_regs rl rf, stack). +Proof. + induction rl; intros. auto. + simpl. rewrite IHrl; auto. +Qed. diff --git a/backend/LTLtyping.v b/backend/LTLtyping.v new file mode 100644 index 0000000000..50995c4d7e --- /dev/null +++ b/backend/LTLtyping.v @@ -0,0 +1,291 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Type-checking LTL code, adapted with tiny changes from [Lineartyping]. *) + +Require Import Coqlib. +Require Import AST. +Require Import Integers. +Require Import Values. +Require Import Globalenvs. +Require Import Memory. +Require Import Events. +Require Import Op. +Require Import Machregs. +Require Import Locations. +Require Import Conventions. +Require Import LTL. + +(** The rules are presented as boolean-valued functions so that we + get an executable type-checker for free. *) + +Section WT_INSTR. + +Variable funct: function. + +Definition slot_valid (sl: slot) (ofs: Z) (q: quantity): bool := + match sl with + | Local => zle 0 ofs + | Outgoing => zle 0 ofs + | Incoming => In_dec Loc.eq (S Incoming ofs q) (regs_of_rpairs (loc_parameters funct.(fn_sig))) + end + && Zdivide_dec (typealign (typ_of_quantity q)) ofs (typealign_pos (typ_of_quantity q)). + +Definition slot_writable (sl: slot) : bool := + match sl with + | Local => true + | Outgoing => true + | Incoming => false + end. + +Definition loc_valid (l: loc) : bool := + match l with + | R r => true + | S Local ofs q => slot_valid Local ofs q + | S _ _ _ => false + end. + +Definition wt_builtin_res (ty: typ) (res: builtin_res mreg) : bool := + match res with + | BR r => subtype ty (mreg_type r) + | BR_none => true + | BR_splitlong hi lo => subtype Tint (mreg_type hi) && subtype Tint (mreg_type lo) + end. + +Definition wt_instr (i: instruction) : bool := + match i with + | Lgetstack sl ofs q r => + subtype (typ_of_quantity q) (mreg_type r) && slot_valid sl ofs q + | Lsetstack r sl ofs q => + slot_valid sl ofs q && slot_writable sl + | Lop op args res => + match is_move_operation op args with + | Some arg => + subtype (mreg_type arg) (mreg_type res) + | None => + let (targs, tres) := type_of_operation op in + subtype tres (mreg_type res) + end + | Lload chunk addr args dst => + subtype (type_of_chunk chunk) (mreg_type dst) + | Ltailcall sg ros => + zeq (size_arguments sg) 0 + | Lbuiltin ef args res => + wt_builtin_res (proj_sig_res (ef_sig ef)) res + && forallb loc_valid (params_of_builtin_args args) + | _ => + true + end. + +End WT_INSTR. + +Definition wt_bblock (f: function) (b: bblock) : bool := + forallb (wt_instr f) b. + +Definition wt_function (f: function) : bool := + let bs := map snd (Maps.PTree.elements f.(fn_code)) in + forallb (wt_bblock f) bs. + +Lemma wt_function_wt_bblock: + forall f pc b, + wt_function f = true -> + Maps.PTree.get pc (fn_code f) = Some b -> + wt_bblock f b = true. +Proof. + intros. apply Maps.PTree.elements_correct in H0. + unfold wt_function, wt_bblock in *. eapply forallb_forall in H; eauto. + change b with (snd (pc, b)). apply in_map; auto. +Qed. + +(** Typing the run-time state. *) + +Definition wt_locset (ls: locset) : Prop := + forall l, Val.has_type (ls @ l) (Loc.type l). + +Lemma well_typed_locset: + forall ls, wt_locset ls. +Proof. + unfold wt_locset. intros. apply Locmap.get_has_type. +Qed. + +(** Soundness of the type system *) + +Definition wt_fundef (fd: fundef) := + match fd with + | Internal f => wt_function f = true + | External ef => True + end. + +Definition wt_program (prog: program): Prop := + forall i fd, In (i, Gfun fd) prog.(prog_defs) -> wt_fundef fd. + +Inductive wt_callstack: list stackframe -> Prop := + | wt_callstack_nil: + wt_callstack nil + | wt_callstack_cons: forall f sp rs b s + (WTSTK: wt_callstack s) + (WTF: wt_function f = true) + (WTB: wt_bblock f b = true), + wt_callstack (Stackframe f sp rs b :: s). + +Inductive wt_state: state -> Prop := + | wt_branch_state: forall s f sp n rs m + (WTSTK: wt_callstack s ) + (WTF: wt_function f = true), + wt_state (State s f sp n rs m) + | wt_regular_state: forall s f sp b rs m + (WTSTK: wt_callstack s ) + (WTF: wt_function f = true) + (WTB: wt_bblock f b = true), + wt_state (Block s f sp b rs m) + | wt_call_state: forall s fd rs m + (WTSTK: wt_callstack s) + (WTFD: wt_fundef fd), + wt_state (Callstate s fd rs m) + | wt_return_state: forall s rs m + (WTSTK: wt_callstack s), + wt_state (Returnstate s rs m). + +(** Preservation of state typing by transitions *) + +Section SOUNDNESS. + +Variable prog: program. +Let ge := Genv.globalenv prog. + +Hypothesis wt_prog: wt_program prog. + +Lemma wt_find_function: + forall ros rs f, find_function ge ros rs = Some f -> wt_fundef f. +Proof. + intros. + assert (X: exists i, In (i, Gfun f) prog.(prog_defs)). + { + destruct ros as [r | s]; simpl in H. + eapply Genv.find_funct_inversion; eauto. + destruct (Genv.find_symbol ge s) as [b|]; try discriminate. + eapply Genv.find_funct_ptr_inversion; eauto. + } + destruct X as [i IN]. eapply wt_prog; eauto. +Qed. + +Theorem step_type_preservation: + forall S1 t S2, step ge S1 t S2 -> wt_state S1 -> wt_state S2. +Proof. +Local Opaque mreg_type. + induction 1; intros WTS; inv WTS. +- (* startblock *) + econstructor; eauto. + apply Maps.PTree.elements_correct in H. + unfold wt_function in WTF. eapply forallb_forall in WTF; eauto. + change bb with (snd (pc, bb)). apply in_map; auto. +- (* op *) + simpl in *. destruct (is_move_operation op args) as [src | ] eqn:ISMOVE. + + (* move *) + InvBooleans. exploit is_move_operation_correct; eauto. intros [EQ1 EQ2]; subst. + simpl in H. inv H. + econstructor; eauto. + + (* other ops *) + destruct (type_of_operation op) as [ty_args ty_res] eqn:TYOP. InvBooleans. + econstructor; eauto. +- (* load *) + simpl in *; InvBooleans. + econstructor; eauto. +- (* getstack *) + simpl in *; InvBooleans. + econstructor; eauto. +- (* setstack *) + simpl in *; InvBooleans. + econstructor; eauto. +- (* store *) + simpl in *; InvBooleans. + econstructor. eauto. eauto. eauto. +- (* call *) + simpl in *; InvBooleans. + econstructor; eauto. econstructor; eauto. + eapply wt_find_function; eauto. +- (* tailcall *) + simpl in *; InvBooleans. + econstructor; eauto. + eapply wt_find_function; eauto. +- (* builtin *) + simpl in *; InvBooleans. + econstructor; eauto. +- (* branch *) + simpl in *. econstructor; eauto. +- (* cond branch *) + simpl in *. econstructor; auto. +- (* jumptable *) + simpl in *. econstructor; auto. +- (* return *) + simpl in *. InvBooleans. + econstructor; eauto. +- (* internal function *) + simpl in WTFD. + econstructor. eauto. eauto. +- (* external function *) + econstructor. auto. +- (* return *) + inv WTSTK. econstructor; eauto. +Qed. + +Theorem wt_initial_state: + forall S, initial_state prog S -> wt_state S. +Proof. + induction 1. econstructor. constructor. + unfold ge0 in H1. exploit Genv.find_funct_ptr_inversion; eauto. + intros [id IN]. eapply wt_prog; eauto. +Qed. + +End SOUNDNESS. + +(** Properties of well-typed states that are used in [Allocproof]. *) + +Lemma wt_state_getstack: + forall s f sp sl ofs q rd c rs m, + wt_state (Block s f sp (Lgetstack sl ofs q rd :: c) rs m) -> + slot_valid f sl ofs q = true. +Proof. + intros. inv H. simpl in WTB; InvBooleans. auto. +Qed. + +Lemma wt_state_setstack: + forall s f sp sl ofs q r c rs m, + wt_state (Block s f sp (Lsetstack r sl ofs q :: c) rs m) -> + slot_valid f sl ofs q = true /\ slot_writable sl = true. +Proof. + intros. inv H. simpl in WTB; InvBooleans. intuition. +Qed. + +Lemma wt_state_tailcall: + forall s f sp sg ros c rs m, + wt_state (Block s f sp (Ltailcall sg ros :: c) rs m) -> + size_arguments sg = 0. +Proof. + intros. inv H. simpl in WTB; InvBooleans. auto. +Qed. + +Lemma wt_state_builtin: + forall s f sp ef args res c rs m, + wt_state (Block s f sp (Lbuiltin ef args res :: c) rs m) -> + forallb (loc_valid f) (params_of_builtin_args args) = true. +Proof. + intros. inv H. simpl in WTB; InvBooleans. auto. +Qed. + +Lemma wt_callstate_wt_regs: + forall s f rs m, + wt_state (Callstate s f rs m) -> + forall r, Val.has_type (rs @ (R r)) (mreg_type r). +Proof. + intros. apply well_typed_locset. +Qed. diff --git a/backend/Linear.v b/backend/Linear.v index 447c6ba656..789e1ab972 100644 --- a/backend/Linear.v +++ b/backend/Linear.v @@ -25,8 +25,8 @@ Require Import Op Locations LTL Conventions. Definition label := positive. Inductive instruction: Type := - | Lgetstack: slot -> Z -> typ -> mreg -> instruction - | Lsetstack: mreg -> slot -> Z -> typ -> instruction + | Lgetstack: slot -> Z -> quantity -> mreg -> instruction + | Lsetstack: mreg -> slot -> Z -> quantity -> instruction | Lop: operation -> list mreg -> mreg -> instruction | Lload: memory_chunk -> addressing -> list mreg -> mreg -> instruction | Lstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction @@ -95,7 +95,7 @@ Variable ge: genv. Definition find_function (ros: mreg + ident) (rs: locset) : option fundef := match ros with - | inl r => Genv.find_funct ge (rs (R r)) + | inl r => Genv.find_funct ge (rs @ (R r)) | inr symb => match Genv.find_symbol ge symb with | None => None @@ -138,20 +138,20 @@ Inductive state: Type := of the caller function. *) Definition parent_locset (stack: list stackframe) : locset := match stack with - | nil => Locmap.init Vundef + | nil => Locmap.init | Stackframe f sp ls c :: stack' => ls end. Inductive step: state -> trace -> state -> Prop := | exec_Lgetstack: - forall s f sp sl ofs ty dst b rs m rs', - rs' = Locmap.set (R dst) (rs (S sl ofs ty)) (undef_regs (destroyed_by_getstack sl) rs) -> - step (State s f sp (Lgetstack sl ofs ty dst :: b) rs m) + forall s f sp sl ofs q dst b rs m rs', + rs' = Locmap.set (R dst) (rs @ (S sl ofs q)) (undef_regs (destroyed_by_getstack sl) rs) -> + step (State s f sp (Lgetstack sl ofs q dst :: b) rs m) E0 (State s f sp b rs' m) | exec_Lsetstack: - forall s f sp src sl ofs ty b rs m rs', - rs' = Locmap.set (S sl ofs ty) (rs (R src)) (undef_regs (destroyed_by_setstack ty) rs) -> - step (State s f sp (Lsetstack src sl ofs ty :: b) rs m) + forall s f sp src sl ofs q b rs m rs', + rs' = Locmap.set (S sl ofs q) (rs @ (R src)) (undef_regs (destroyed_by_setstack q) rs) -> + step (State s f sp (Lsetstack src sl ofs q :: b) rs m) E0 (State s f sp b rs' m) | exec_Lop: forall s f sp op args res b rs m v rs', @@ -169,7 +169,7 @@ Inductive step: state -> trace -> state -> Prop := | exec_Lstore: forall s f sp chunk addr args src b rs m m' a rs', eval_addressing ge sp addr (reglist rs args) = Some a -> - Mem.storev chunk m a (rs (R src)) = Some m' -> + Mem.storev chunk m a (rs @ (R src)) = Some m' -> rs' = undef_regs (destroyed_by_store chunk addr) rs -> step (State s f sp (Lstore chunk addr args src :: b) rs m) E0 (State s f sp b rs' m') @@ -189,7 +189,7 @@ Inductive step: state -> trace -> state -> Prop := E0 (Callstate s f' rs' m') | exec_Lbuiltin: forall s f sp rs m ef args res b vargs t vres rs' m', - eval_builtin_args ge rs sp m args vargs -> + eval_builtin_args ge (Locmap.read rs) sp m args vargs -> external_call ef ge vargs m t vres m' -> rs' = Locmap.setres res vres (undef_regs (destroyed_by_builtin ef) rs) -> step (State s f sp (Lbuiltin ef args res :: b) rs m) @@ -218,7 +218,7 @@ Inductive step: state -> trace -> state -> Prop := E0 (State s f sp b rs' m) | exec_Ljumptable: forall s f sp arg tbl b rs m n lbl b' rs', - rs (R arg) = Vint n -> + rs @ (R arg) = Vint n -> list_nth_z tbl (Int.unsigned n) = Some lbl -> find_label lbl f.(fn_code) = Some b' -> rs' = undef_regs (destroyed_by_jumptable) rs -> @@ -256,7 +256,7 @@ Inductive initial_state (p: program): state -> Prop := Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> funsig f = signature_main -> - initial_state p (Callstate nil f (Locmap.init Vundef) m0). + initial_state p (Callstate nil f Locmap.init m0). Inductive final_state: state -> int -> Prop := | final_state_intro: forall rs m retcode, diff --git a/backend/Linearizeproof.v b/backend/Linearizeproof.v index 10a3d8b2a4..5f3e13490d 100644 --- a/backend/Linearizeproof.v +++ b/backend/Linearizeproof.v @@ -516,7 +516,7 @@ Inductive match_states: LTL.state -> Linear.state -> Prop := (STACKS: list_forall2 match_stackframes s ts) (TRF: transf_function f = OK tf) (REACH: (reachable f)!!pc = true) - (ARG: ls (R arg) = Vint n) + (ARG: ls @ (R arg) = Vint n) (JUMP: list_nth_z tbl (Int.unsigned n) = Some pc), match_states (LTL.State s f sp pc (undef_regs destroyed_by_jumptable ls) m) (Linear.State ts tf sp (Ljumptable arg tbl :: c) ls m) @@ -710,7 +710,7 @@ Lemma transf_initial_states: Proof. intros. inversion H. exploit function_ptr_translated; eauto. intros [tf [A B]]. - exists (Callstate nil tf (Locmap.init Vundef) m0); split. + exists (Callstate nil tf Locmap.init m0); split. econstructor; eauto. eapply (Genv.init_mem_transf_partial TRANSF); eauto. rewrite (match_program_main TRANSF). rewrite symbols_preserved. eauto. diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v index fc16371956..6c388a4cba 100644 --- a/backend/Lineartyping.v +++ b/backend/Lineartyping.v @@ -33,13 +33,13 @@ Section WT_INSTR. Variable funct: function. -Definition slot_valid (sl: slot) (ofs: Z) (ty: typ): bool := +Definition slot_valid (sl: slot) (ofs: Z) (q: quantity): bool := match sl with | Local => zle 0 ofs | Outgoing => zle 0 ofs - | Incoming => In_dec Loc.eq (S Incoming ofs ty) (regs_of_rpairs (loc_parameters funct.(fn_sig))) + | Incoming => In_dec Loc.eq (S Incoming ofs q) (regs_of_rpairs (loc_parameters funct.(fn_sig))) end - && Zdivide_dec (typealign ty) ofs (typealign_pos ty). + && Zdivide_dec (typealign (typ_of_quantity q)) ofs (typealign_pos (typ_of_quantity q)). Definition slot_writable (sl: slot) : bool := match sl with @@ -51,23 +51,23 @@ Definition slot_writable (sl: slot) : bool := Definition loc_valid (l: loc) : bool := match l with | R r => true - | S Local ofs ty => slot_valid Local ofs ty + | S Local ofs q => slot_valid Local ofs q | S _ _ _ => false end. -Fixpoint wt_builtin_res (ty: typ) (res: builtin_res mreg) : bool := +Definition wt_builtin_res (ty: typ) (res: builtin_res mreg) : bool := match res with | BR r => subtype ty (mreg_type r) | BR_none => true - | BR_splitlong hi lo => wt_builtin_res Tint hi && wt_builtin_res Tint lo + | BR_splitlong hi lo => subtype Tint (mreg_type hi) && subtype Tint (mreg_type lo) end. Definition wt_instr (i: instruction) : bool := match i with - | Lgetstack sl ofs ty r => - subtype ty (mreg_type r) && slot_valid sl ofs ty - | Lsetstack r sl ofs ty => - slot_valid sl ofs ty && slot_writable sl + | Lgetstack sl ofs q r => + subtype (typ_of_quantity q) (mreg_type r) && slot_valid sl ofs q + | Lsetstack r sl ofs q => + slot_valid sl ofs q && slot_writable sl | Lop op args res => match is_move_operation op args with | Some arg => @@ -98,101 +98,12 @@ Definition wt_function (f: function) : bool := (** Typing the run-time state. *) Definition wt_locset (ls: locset) : Prop := - forall l, Val.has_type (ls l) (Loc.type l). + forall l, Val.has_type (ls @ l) (Loc.type l). -Lemma wt_setreg: - forall ls r v, - Val.has_type v (mreg_type r) -> wt_locset ls -> wt_locset (Locmap.set (R r) v ls). +Lemma well_typed_locset: + forall ls, wt_locset ls. Proof. - intros; red; intros. - unfold Locmap.set. - destruct (Loc.eq (R r) l). - subst l; auto. - destruct (Loc.diff_dec (R r) l). auto. red. auto. -Qed. - -Lemma wt_setstack: - forall ls sl ofs ty v, - wt_locset ls -> wt_locset (Locmap.set (S sl ofs ty) v ls). -Proof. - intros; red; intros. - unfold Locmap.set. - destruct (Loc.eq (S sl ofs ty) l). - subst l. simpl. - generalize (Val.load_result_type (chunk_of_type ty) v). - replace (type_of_chunk (chunk_of_type ty)) with ty. auto. - destruct ty; reflexivity. - destruct (Loc.diff_dec (S sl ofs ty) l). auto. red. auto. -Qed. - -Lemma wt_undef_regs: - forall rs ls, wt_locset ls -> wt_locset (undef_regs rs ls). -Proof. - induction rs; simpl; intros. auto. apply wt_setreg; auto. red; auto. -Qed. - -Lemma wt_call_regs: - forall ls, wt_locset ls -> wt_locset (call_regs ls). -Proof. - intros; red; intros. unfold call_regs. destruct l. auto. - destruct sl. - red; auto. - change (Loc.type (S Incoming pos ty)) with (Loc.type (S Outgoing pos ty)). auto. - red; auto. -Qed. - -Lemma wt_return_regs: - forall caller callee, - wt_locset caller -> wt_locset callee -> wt_locset (return_regs caller callee). -Proof. - intros; red; intros. - unfold return_regs. destruct l. -- destruct (is_callee_save r); auto. -- destruct sl; auto; red; auto. -Qed. - -Lemma wt_undef_caller_save_regs: - forall ls, wt_locset ls -> wt_locset (undef_caller_save_regs ls). -Proof. - intros; red; intros. unfold undef_caller_save_regs. - destruct l. - destruct (is_callee_save r); auto; simpl; auto. - destruct sl; auto; red; auto. -Qed. - -Lemma wt_init: - wt_locset (Locmap.init Vundef). -Proof. - red; intros. unfold Locmap.init. red; auto. -Qed. - -Lemma wt_setpair: - forall sg v rs, - Val.has_type v (proj_sig_res sg) -> - wt_locset rs -> - wt_locset (Locmap.setpair (loc_result sg) v rs). -Proof. - intros. generalize (loc_result_pair sg) (loc_result_type sg). - destruct (loc_result sg); simpl Locmap.setpair. -- intros. apply wt_setreg; auto. eapply Val.has_subtype; eauto. -- intros A B. decompose [and] A. - apply wt_setreg. eapply Val.has_subtype; eauto. destruct v; exact I. - apply wt_setreg. eapply Val.has_subtype; eauto. destruct v; exact I. - auto. -Qed. - -Lemma wt_setres: - forall res ty v rs, - wt_builtin_res ty res = true -> - Val.has_type v ty -> - wt_locset rs -> - wt_locset (Locmap.setres res v rs). -Proof. - induction res; simpl; intros. -- apply wt_setreg; auto. eapply Val.has_subtype; eauto. -- auto. -- InvBooleans. eapply IHres2; eauto. destruct v; exact I. - eapply IHres1; eauto. destruct v; exact I. + unfold wt_locset. intros. apply Locmap.get_has_type. Qed. Lemma wt_find_label: @@ -218,13 +129,13 @@ Qed. Definition agree_outgoing_arguments (sg: signature) (ls pls: locset) : Prop := forall ty ofs, In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments sg)) -> - ls (S Outgoing ofs ty) = pls (S Outgoing ofs ty). + ls @ (S Outgoing ofs ty) = pls @ (S Outgoing ofs ty). (** For return points, we have that all [Outgoing] stack locations have been set to [Vundef]. *) Definition outgoing_undef (ls: locset) : Prop := - forall ty ofs, ls (S Outgoing ofs ty) = Vundef. + forall ty ofs, ls @ (S Outgoing ofs ty) = Vundef. (** Soundness of the type system *) @@ -240,35 +151,23 @@ Inductive wt_callstack: list stackframe -> Prop := | wt_callstack_cons: forall f sp rs c s (WTSTK: wt_callstack s) (WTF: wt_function f = true) - (WTC: wt_code f c = true) - (WTRS: wt_locset rs), + (WTC: wt_code f c = true), wt_callstack (Stackframe f sp rs c :: s). -Lemma wt_parent_locset: - forall s, wt_callstack s -> wt_locset (parent_locset s). -Proof. - induction 1; simpl. -- apply wt_init. -- auto. -Qed. - Inductive wt_state: state -> Prop := | wt_regular_state: forall s f sp c rs m (WTSTK: wt_callstack s ) (WTF: wt_function f = true) - (WTC: wt_code f c = true) - (WTRS: wt_locset rs), + (WTC: wt_code f c = true), wt_state (State s f sp c rs m) | wt_call_state: forall s fd rs m (WTSTK: wt_callstack s) (WTFD: wt_fundef fd) - (WTRS: wt_locset rs) (AGCS: agree_callee_save rs (parent_locset s)) (AGARGS: agree_outgoing_arguments (funsig fd) rs (parent_locset s)), wt_state (Callstate s fd rs m) | wt_return_state: forall s rs m (WTSTK: wt_callstack s) - (WTRS: wt_locset rs) (AGCS: agree_callee_save rs (parent_locset s)) (UOUT: outgoing_undef rs), wt_state (Returnstate s rs m). @@ -305,37 +204,24 @@ Local Opaque mreg_type. - (* getstack *) simpl in *; InvBooleans. econstructor; eauto. - eapply wt_setreg; eauto. eapply Val.has_subtype; eauto. apply WTRS. - apply wt_undef_regs; auto. - (* setstack *) simpl in *; InvBooleans. econstructor; eauto. - apply wt_setstack. apply wt_undef_regs; auto. - (* op *) simpl in *. destruct (is_move_operation op args) as [src | ] eqn:ISMOVE. + (* move *) InvBooleans. exploit is_move_operation_correct; eauto. intros [EQ1 EQ2]; subst. simpl in H. inv H. - econstructor; eauto. apply wt_setreg. eapply Val.has_subtype; eauto. apply WTRS. - apply wt_undef_regs; auto. + econstructor; eauto. + (* other ops *) destruct (type_of_operation op) as [ty_args ty_res] eqn:TYOP. InvBooleans. econstructor; eauto. - apply wt_setreg; auto. eapply Val.has_subtype; eauto. - change ty_res with (snd (ty_args, ty_res)). rewrite <- TYOP. eapply type_of_operation_sound; eauto. - red; intros; subst op. simpl in ISMOVE. - destruct args; try discriminate. destruct args; discriminate. - apply wt_undef_regs; auto. - (* load *) simpl in *; InvBooleans. econstructor; eauto. - apply wt_setreg. eapply Val.has_subtype; eauto. - destruct a; simpl in H0; try discriminate. eapply Mem.load_type; eauto. - apply wt_undef_regs; auto. - (* store *) simpl in *; InvBooleans. econstructor. eauto. eauto. eauto. - apply wt_undef_regs; auto. - (* call *) simpl in *; InvBooleans. econstructor; eauto. econstructor; eauto. @@ -346,45 +232,47 @@ Local Opaque mreg_type. simpl in *; InvBooleans. econstructor; eauto. eapply wt_find_function; eauto. - apply wt_return_regs; auto. apply wt_parent_locset; auto. - red; simpl; intros. destruct l; simpl in *. rewrite H3; auto. destruct sl; auto; congruence. + red; simpl; intros. + rewrite return_regs_correct. unfold Locmap.get. + destruct l; simpl in *. rewrite H3; auto. destruct sl; auto; congruence. red; simpl; intros. apply zero_size_arguments_tailcall_possible in H. apply H in H3. contradiction. - (* builtin *) simpl in *; InvBooleans. econstructor; eauto. - eapply wt_setres; eauto. eapply external_call_well_typed; eauto. - apply wt_undef_regs; auto. - (* label *) simpl in *. econstructor; eauto. - (* goto *) simpl in *. econstructor; eauto. eapply wt_find_label; eauto. - (* cond branch, taken *) simpl in *. econstructor. auto. auto. eapply wt_find_label; eauto. - apply wt_undef_regs; auto. - (* cond branch, not taken *) simpl in *. econstructor. auto. auto. auto. - apply wt_undef_regs; auto. - (* jumptable *) simpl in *. econstructor. auto. auto. eapply wt_find_label; eauto. - apply wt_undef_regs; auto. - (* return *) simpl in *. InvBooleans. econstructor; eauto. - apply wt_return_regs; auto. apply wt_parent_locset; auto. - red; simpl; intros. destruct l; simpl in *. rewrite H0; auto. destruct sl; auto; congruence. - red; simpl; intros. auto. + red; simpl; intros. + rewrite return_regs_correct. unfold Locmap.get. + destruct l; simpl in *. rewrite H0; auto. destruct sl; auto; congruence. + red; intros. + rewrite return_regs_correct. unfold Locmap.get, return_regs_spec. auto. - (* internal function *) simpl in WTFD. econstructor. eauto. eauto. eauto. - apply wt_undef_regs. apply wt_call_regs. auto. - (* external function *) - econstructor. auto. apply wt_setpair. - eapply external_call_well_typed; eauto. - apply wt_undef_caller_save_regs; auto. - red; simpl; intros. destruct l; simpl in *. - rewrite locmap_get_set_loc_result by auto. simpl. rewrite H; auto. - rewrite locmap_get_set_loc_result by auto. simpl. destruct sl; auto; congruence. - red; simpl; intros. rewrite locmap_get_set_loc_result by auto. auto. + econstructor. auto. + red; simpl; intros. destruct l. + rewrite locmap_get_set_loc_result by auto. + rewrite undef_caller_save_regs_correct. unfold undef_caller_save_regs_spec. rewrite H; auto. + apply AGCS; auto. + rewrite locmap_get_set_loc_result by auto. + rewrite undef_caller_save_regs_correct. unfold undef_caller_save_regs_spec. + destruct sl; try apply AGCS; auto. + simpl in H; contradiction. + red; intros. rewrite locmap_get_set_loc_result by auto. + rewrite undef_caller_save_regs_correct. + unfold Locmap.get, undef_caller_save_regs_spec. auto. - (* return *) inv WTSTK. econstructor; eauto. Qed. @@ -395,7 +283,6 @@ Proof. induction 1. econstructor. constructor. unfold ge0 in H1. exploit Genv.find_funct_ptr_inversion; eauto. intros [id IN]. eapply wt_prog; eauto. - apply wt_init. red; auto. red; auto. Qed. @@ -405,17 +292,17 @@ End SOUNDNESS. (** Properties of well-typed states that are used in [Stackingproof]. *) Lemma wt_state_getstack: - forall s f sp sl ofs ty rd c rs m, - wt_state (State s f sp (Lgetstack sl ofs ty rd :: c) rs m) -> - slot_valid f sl ofs ty = true. + forall s f sp sl ofs q rd c rs m, + wt_state (State s f sp (Lgetstack sl ofs q rd :: c) rs m) -> + slot_valid f sl ofs q = true. Proof. intros. inv H. simpl in WTC; InvBooleans. auto. Qed. Lemma wt_state_setstack: - forall s f sp sl ofs ty r c rs m, - wt_state (State s f sp (Lsetstack r sl ofs ty :: c) rs m) -> - slot_valid f sl ofs ty = true /\ slot_writable sl = true. + forall s f sp sl ofs q r c rs m, + wt_state (State s f sp (Lsetstack r sl ofs q :: c) rs m) -> + slot_valid f sl ofs q = true /\ slot_writable sl = true. Proof. intros. inv H. simpl in WTC; InvBooleans. intuition. Qed. @@ -439,9 +326,9 @@ Qed. Lemma wt_callstate_wt_regs: forall s f rs m, wt_state (Callstate s f rs m) -> - forall r, Val.has_type (rs (R r)) (mreg_type r). + forall r, Val.has_type (rs @ (R r)) (mreg_type r). Proof. - intros. inv H. apply WTRS. + intros. apply well_typed_locset. Qed. Lemma wt_callstate_agree: diff --git a/backend/Locations.v b/backend/Locations.v index c437df5dd6..fca5149f51 100644 --- a/backend/Locations.v +++ b/backend/Locations.v @@ -19,7 +19,11 @@ Require Import Maps. Require Import Ordered. Require Import AST. Require Import Values. +Require Import Memory. +Require Export Memdata. +Require Export FragmentBlock. Require Export Machregs. +Require Export Registerfile. (** * Representation of locations *) @@ -54,9 +58,9 @@ slots of its caller function. The type of a slot indicates how it will be accessed later once mapped to actual memory locations inside a memory-allocated activation record: -as 32-bit integers/pointers (type [Tint]) or as 64-bit floats (type [Tfloat]). +as 32-bit or 64-bit quantities ([Q32] or [Q64]). -The offset of a slot, combined with its type and its kind, identifies +The offset of a slot, combined with its quantity and its kind, identifies uniquely the slot and will determine later where it resides within the memory-allocated activation record. Offsets are always positive. *) @@ -106,6 +110,20 @@ Proof. intros. exists (typesize ty / typealign ty); destruct ty; reflexivity. Qed. +Definition quantity_align (q: quantity) : Z := 1. + +Lemma quantity_align_pos: + forall (q: quantity), quantity_align q > 0. +Proof. + destruct q; compute; auto. +Qed. + +Lemma quantity_align_typesize: + forall (q: quantity), (quantity_align q | typesize (typ_of_quantity q)). +Proof. + intros. destruct q; simpl. exists 1; eauto. exists 2; eauto. +Qed. + (** ** Locations *) (** Locations are just the disjoint union of machine registers and @@ -113,21 +131,27 @@ Qed. Inductive loc : Type := | R (r: mreg) - | S (sl: slot) (pos: Z) (ty: typ). + | S (sl: slot) (pos: Z) (q: quantity). Module Loc. Definition type (l: loc) : typ := match l with | R r => mreg_type r - | S sl pos ty => ty + | S sl pos q => typ_of_quantity q + end. + + Definition quantity (l: loc) : quantity := + match l with + | R r => quantity_of_mreg r + | S sl pos q => q end. Lemma eq: forall (p q: loc), {p = q} + {p <> q}. Proof. decide equality. apply mreg_eq. - apply typ_eq. + apply quantity_eq. apply zeq. apply slot_eq. Defined. @@ -147,8 +171,8 @@ Module Loc. match l1, l2 with | R r1, R r2 => r1 <> r2 - | S s1 d1 t1, S s2 d2 t2 => - s1 <> s2 \/ d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1 + | S s1 d1 q1, S s2 d2 q2 => + s1 <> s2 \/ d1 + typesize (typ_of_quantity q1) <= d2 \/ d2 + typesize (typ_of_quantity q2) <= d1 | _, _ => True end. @@ -157,7 +181,7 @@ Module Loc. forall l, ~(diff l l). Proof. destruct l; unfold diff; auto. - red; intros. destruct H; auto. generalize (typesize_pos ty); omega. + red; intros. destruct H; auto. generalize (typesize_pos (typ_of_quantity q)); omega. Qed. Lemma diff_not_eq: @@ -180,9 +204,9 @@ Module Loc. - left; auto. - left; auto. - destruct (slot_eq sl sl0). - destruct (zle (pos + typesize ty) pos0). + destruct (zle (pos + typesize (typ_of_quantity q)) pos0). left; auto. - destruct (zle (pos0 + typesize ty0) pos). + destruct (zle (pos0 + typesize (typ_of_quantity q0)) pos). left; auto. right; red; intros [P | [P | P]]. congruence. omega. omega. left; auto. @@ -298,6 +322,143 @@ Module Loc. End Loc. +(** * Representation of the stack frame *) + +(** The [Stack] module defines mappings from stack slots to values, + represented as a mapping from addresses to bytes. The stack frame is + represented as three distinct such mappings, one for each kind of slot; a + stack slot [S sl ofs q] is accessed in the mapping for its kind [sl] at an + address computed from its offset [ofs], using [q] to inform how many bytes to + access and how to encode/decode them. *) + +Module Stack. + + Definition t := slot -> FragBlock.t. + + Definition init : t := fun _ => FragBlock.init. + + (* A slot's address: The index of its first byte. *) + Definition addr (ofs: Z) : Z := FragBlock.addr ofs. + + (* The address one byte past the end of a slot with [ofs] and [q]. The next + nonoverlapping slot may start here. *) + Definition next_addr (ofs: Z) (q: quantity) : Z := FragBlock.next_addr ofs q. + + Definition get_bytes sl ofs q (stack: t) : list memval := + FragBlock.get_bytes ofs q (stack sl). + + Definition get sl ofs q (stack: t) : val := + FragBlock.get ofs q (stack sl). + + Definition set_bytes sl ofs q (bytes: list memval) (stack: t) : t := + fun slot => + if slot_eq slot sl then + FragBlock.set_bytes ofs q bytes (stack sl) + else + stack slot. + + Definition set sl ofs q (v: val) (stack: t) : t := + fun slot => + if slot_eq slot sl then + FragBlock.set ofs q v (stack sl) + else + stack slot. + + (* The [Loc.diff] predicate on stack slots is compatible with [FragBlock]'s + notion of non-overlapping accesses. *) + Lemma diff_compat: + forall sl ofs q ofs' q', + Loc.diff (S sl ofs q) (S sl ofs' q') -> + FragBlock.next_addr ofs q <= FragBlock.addr ofs' \/ + FragBlock.next_addr ofs' q' <= FragBlock.addr ofs. + Proof. + intros. + unfold Loc.diff in H. destruct H; try contradiction. + unfold FragBlock.next_addr, FragBlock.addr, FragBlock.quantity_size. + destruct q, q'; simpl in *; omega. + Qed. + + Lemma gss: + forall sl ofs q v stack, + get sl ofs q (set sl ofs q v stack) = Val.load_result (chunk_of_quantity q) v. + Proof. + intros. unfold get, set. rewrite dec_eq_true, FragBlock.gss; auto. + Qed. + + Lemma gso: + forall sl ofs q sl' ofs' q' v stack, + Loc.diff (S sl ofs q) (S sl' ofs' q') -> + get sl ofs q (set sl' ofs' q' v stack) = get sl ofs q stack. + Proof. + intros. unfold get, set. + destruct (slot_eq sl sl'); subst; auto. + apply FragBlock.gso. eauto using diff_compat. + Qed. + + Lemma gss_bytes: + forall sl ofs q bs rf, + length bs = size_quantity_nat q -> + get_bytes sl ofs q (set_bytes sl ofs q bs rf) = bs. + Proof. + intros. unfold get_bytes, set_bytes. + rewrite dec_eq_true, FragBlock.gss_bytes; auto. + Qed. + + Lemma gso_bytes: + forall sl ofs q sl' ofs' q' bs stack, + length bs = size_quantity_nat q' -> + Loc.diff (S sl ofs q) (S sl' ofs' q') -> + get_bytes sl ofs q (set_bytes sl' ofs' q' bs stack) = get_bytes sl ofs q stack. + Proof. + intros. unfold get_bytes, set_bytes. + destruct (slot_eq sl sl'); subst; auto. + apply FragBlock.gso_bytes; auto. + eauto using diff_compat. + Qed. + + Lemma gu_overlap: + forall sl ofs q sl' ofs' q' v stack, + S sl ofs q <> S sl' ofs' q' -> + ~ Loc.diff (S sl ofs q) (S sl' ofs' q') -> + get sl ofs q (set sl' ofs' q' v stack) = Vundef. + Proof. + intros. unfold get, set. + simpl in H0. apply Decidable.not_or in H0. destruct H0. + assert (sl = sl'). { destruct (slot_eq sl sl'). auto. contradiction. } subst sl'. clear H0. + apply Decidable.not_or in H1. destruct H1. + apply Z.nle_gt in H0. apply Z.nle_gt in H1. + rewrite pred_dec_true by auto. + apply FragBlock.gu_overlap. + congruence. + unfold FragBlock.next_addr, FragBlock.addr, FragBlock.quantity_size. + destruct q, q'; simpl in *; omega. + unfold FragBlock.next_addr, FragBlock.addr, FragBlock.quantity_size. + destruct q, q'; simpl in *; omega. + Qed. + + Lemma get_has_type: + forall sl ofs q stack, + Val.has_type (get sl ofs q stack) (Loc.type (S sl ofs q)). + Proof. + intros. unfold get. apply FragBlock.get_has_type. + Qed. + + Lemma get_bytes_compat: + forall sl ofs q stack, + get sl ofs q stack = decode_val (chunk_of_quantity q) (get_bytes sl ofs q stack). + Proof. + reflexivity. + Qed. + + Lemma set_bytes_compat: + forall sl ofs q v stack, + set sl ofs q v stack = set_bytes sl ofs q (encode_val (chunk_of_quantity q) v) stack. + Proof. + reflexivity. + Qed. + +End Stack. + (** * Mappings from locations to values *) (** The [Locmap] module defines mappings from locations to values, @@ -308,11 +469,30 @@ Set Implicit Arguments. Module Locmap. - Definition t := loc -> val. + Module Regfile := Regfile(Mreg). - Definition init (x: val) : t := fun (_: loc) => x. + Definition t := (Regfile.t * Stack.t)%type. - Definition get (l: loc) (m: t) : val := m l. + Definition chunk_of_loc (l: loc) : memory_chunk := + chunk_of_type (Loc.type l). + + Lemma chunk_of_loc_charact: + forall l, + chunk_of_loc l = chunk_of_type (Loc.type l). + Proof. + destruct l; auto. + Qed. + + Definition init : t := (Regfile.init, Stack.init). + + Definition get (l: loc) (m: t) : val := + match l, m with + | R r, (rf, stack) => Regfile.get r rf + | S sl ofs q, (rf, stack) => Stack.get sl ofs q stack + end. + + (* Auxiliary for some places where a function of type [loc -> val] is expected. *) + Definition read (m: t) : loc -> val := fun (l: loc) => get l m. (** The [set] operation over location mappings reflects the overlapping properties of locations: changing the value of a location [l] @@ -329,36 +509,39 @@ Module Locmap. are normalized according to the type of the slot. *) Definition set (l: loc) (v: val) (m: t) : t := - fun (p: loc) => - if Loc.eq l p then - match l with R r => v | S sl ofs ty => Val.load_result (chunk_of_type ty) v end - else if Loc.diff_dec l p then - m p - else Vundef. + match l, m with + | R r, (rf, stack) => (Regfile.set r v rf, stack) + | S sl ofs q, (rf, stack) => (rf, Stack.set sl ofs q v stack) + end. Lemma gss: forall l v m, - (set l v m) l = - match l with R r => v | S sl ofs ty => Val.load_result (chunk_of_type ty) v end. + get l (set l v m) = Val.load_result (chunk_of_type (Loc.type l)) v. Proof. - intros. unfold set. apply dec_eq_true. + intros. + destruct l, m. apply Regfile.gss. + unfold get, set. simpl. destruct q; apply Stack.gss. Qed. - Lemma gss_reg: forall r v m, (set (R r) v m) (R r) = v. + Lemma gss_reg: forall r v m, Val.has_type v (mreg_type r) -> get (R r) (set (R r) v m) = v. Proof. - intros. unfold set. rewrite dec_eq_true. auto. + intros. rewrite gss. apply Val.load_result_same; auto. Qed. - Lemma gss_typed: forall l v m, Val.has_type v (Loc.type l) -> (set l v m) l = v. + Lemma gss_typed: forall l v m, Val.has_type v (Loc.type l) -> get l (set l v m) = v. Proof. - intros. rewrite gss. destruct l. auto. apply Val.load_result_same; auto. + intros. rewrite gss. destruct l; simpl. apply Val.load_result_same; auto. + replace (chunk_of_quantity q) with (chunk_of_type (Loc.type (S sl pos q))) by (destruct q; auto). + apply Val.load_result_same; auto. Qed. - Lemma gso: forall l v m p, Loc.diff l p -> (set l v m) p = m p. + Lemma gso: forall l v m p, Loc.diff l p -> get p (set l v m) = get p m. Proof. - intros. unfold set. destruct (Loc.eq l p). - subst p. elim (Loc.same_not_diff _ H). - destruct (Loc.diff_dec l p). - auto. + intros. + destruct l, m. + destruct p; simpl in H; auto. apply Regfile.gso; auto. + unfold get, set. + destruct (Loc.diff_dec (S sl pos q) p). + destruct p; auto using Stack.gso, Loc.diff_sym. contradiction. Qed. @@ -368,28 +551,46 @@ Module Locmap. | l1 :: ll' => undef ll' (set l1 Vundef m) end. - Lemma guo: forall ll l m, Loc.notin l ll -> (undef ll m) l = m l. + Lemma guo: forall ll l m, Loc.notin l ll -> get l (undef ll m) = get l m. Proof. induction ll; simpl; intros. auto. destruct H. rewrite IHll; auto. apply gso. apply Loc.diff_sym; auto. Qed. - Lemma gus: forall ll l m, In l ll -> (undef ll m) l = Vundef. + Lemma gus: forall ll l m, In l ll -> get l (undef ll m) = Vundef. Proof. - assert (P: forall ll l m, m l = Vundef -> (undef ll m) l = Vundef). - induction ll; simpl; intros. auto. apply IHll. - unfold set. destruct (Loc.eq a l). - destruct a. auto. destruct ty; reflexivity. - destruct (Loc.diff_dec a l); auto. + assert (P: forall ll l m, get l m = Vundef -> get l (undef ll m) = Vundef). + { induction ll; simpl; intros. auto. apply IHll. + destruct (Loc.eq a l). subst; apply gss_typed. simpl; auto. + destruct a, l, m; simpl in n. + simpl. rewrite Regfile.gso; auto; congruence. + rewrite gso; simpl; auto. + rewrite gso; simpl; auto. + unfold get, set. simpl in H. (*rewrite dec_eq_false; auto.*) + destruct (Loc.diff_dec (S sl0 pos0 q0) (S sl pos q)). + rewrite Stack.gso; auto. + apply Stack.gu_overlap; auto. } induction ll; simpl; intros. contradiction. destruct H. apply P. subst a. apply gss_typed. exact I. auto. Qed. + Lemma gu_overlap: + forall l l' v m, + l <> l' -> + ~ Loc.diff l l' -> + get l (set l' v m) = Vundef. + Proof. + intros. + destruct l, l'; try (simpl in H0; contradict H0; auto; congruence). + destruct m as [rf stack]; simpl. + auto using Stack.gu_overlap. + Qed. + Definition getpair (p: rpair loc) (m: t) : val := match p with - | One l => m l - | Twolong l1 l2 => Val.longofwords (m l1) (m l2) + | One l => get l m + | Twolong l1 l2 => Val.longofwords (get l1 m) (get l2 m) end. Definition setpair (p: rpair mreg) (v: val) (m: t) : t := @@ -400,7 +601,7 @@ Module Locmap. Lemma getpair_exten: forall p ls1 ls2, - (forall l, In l (regs_of_rpair p) -> ls2 l = ls1 l) -> + (forall l, In l (regs_of_rpair p) -> get l ls2 = get l ls1) -> getpair p ls2 = getpair p ls1. Proof. intros. destruct p; simpl. @@ -410,23 +611,34 @@ Module Locmap. Lemma gpo: forall p v m l, - forall_rpair (fun r => Loc.diff l (R r)) p -> setpair p v m l = m l. + forall_rpair (fun r => Loc.diff l (R r)) p -> get l (setpair p v m) = get l m. Proof. - intros; destruct p; simpl in *. + intros; destruct p; unfold setpair. - apply gso. apply Loc.diff_sym; auto. - destruct H. rewrite ! gso by (apply Loc.diff_sym; auto). auto. Qed. - Fixpoint setres (res: builtin_res mreg) (v: val) (m: t) : t := + Definition setres (res: builtin_res mreg) (v: val) (m: t) : t := match res with | BR r => set (R r) v m | BR_none => m | BR_splitlong hi lo => - setres lo (Val.loword v) (setres hi (Val.hiword v) m) + set (R lo) (Val.loword v) (set (R hi) (Val.hiword v) m) end. + Lemma get_has_type: + forall l m, Val.has_type (get l m) (Loc.type l). + Proof. + intros. + destruct m as [rf stack], l. + - apply Regfile.get_has_type. + - apply Stack.get_has_type. + Qed. + End Locmap. +Notation "a @ b" := (Locmap.get b a) (at level 1) : ltl. + (** * Total ordering over locations *) Module IndexedTyp <: INDEXED_TYPE. @@ -447,6 +659,20 @@ End IndexedTyp. Module OrderedTyp := OrderedIndexed(IndexedTyp). +Module IndexedQuantity <: INDEXED_TYPE. + Definition t := quantity. + Definition index (q: t) := + match q with + | Q32 => 1%positive + | Q64 => 2%positive + end. + Lemma index_inj: forall x y, index x = index y -> x = y. + Proof. destruct x, y; simpl; congruence. Qed. + Definition eq := quantity_eq. +End IndexedQuantity. + +Module OrderedQuantity := OrderedIndexed(IndexedQuantity). + Module IndexedSlot <: INDEXED_TYPE. Definition t := slot. Definition index (x: t) := @@ -466,9 +692,9 @@ Module OrderedLoc <: OrderedType. | R r1, R r2 => Plt (IndexedMreg.index r1) (IndexedMreg.index r2) | R _, S _ _ _ => True | S _ _ _, R _ => False - | S sl1 ofs1 ty1, S sl2 ofs2 ty2 => + | S sl1 ofs1 q1, S sl2 ofs2 q2 => OrderedSlot.lt sl1 sl2 \/ (sl1 = sl2 /\ - (ofs1 < ofs2 \/ (ofs1 = ofs2 /\ OrderedTyp.lt ty1 ty2))) + (ofs1 < ofs2 \/ (ofs1 = ofs2 /\ OrderedQuantity.lt q1 q2))) end. Lemma eq_refl : forall x : t, eq x x. Proof (@eq_refl t). @@ -489,7 +715,7 @@ Module OrderedLoc <: OrderedType. destruct H. right. split. auto. intuition. - right; split. congruence. eapply OrderedTyp.lt_trans; eauto. + right; split. congruence. eapply OrderedQuantity.lt_trans; eauto. Qed. Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y. Proof. @@ -498,7 +724,7 @@ Module OrderedLoc <: OrderedType. eelim Plt_strict; eauto. destruct H. eelim OrderedSlot.lt_not_eq; eauto. red; auto. destruct H. destruct H0. omega. - destruct H0. eelim OrderedTyp.lt_not_eq; eauto. red; auto. + destruct H0. eelim OrderedQuantity.lt_not_eq; eauto. red; auto. Qed. Definition compare : forall x y : t, Compare lt eq x y. Proof. @@ -513,9 +739,9 @@ Module OrderedLoc <: OrderedType. + apply LT. red; auto. + destruct (OrderedZ.compare pos pos0). * apply LT. red. auto. - * destruct (OrderedTyp.compare ty ty0). + * destruct (OrderedQuantity.compare q q0). apply LT. red; auto. - apply EQ. red; red in e; red in e0; red in e1. congruence. + apply EQ. congruence. apply GT. red; auto. * apply GT. red. auto. + apply GT. red; auto. @@ -527,44 +753,44 @@ Module OrderedLoc <: OrderedType. Definition diff_low_bound (l: loc) : loc := match l with | R mr => l - | S sl ofs ty => S sl (ofs - 1) Tany64 + | S sl ofs q => S sl (ofs - 1) Q64 end. Definition diff_high_bound (l: loc) : loc := match l with | R mr => l - | S sl ofs ty => S sl (ofs + typesize ty - 1) Tlong + | S sl ofs q => S sl (ofs + typesize (typ_of_quantity q) - 1) Q64 end. Lemma outside_interval_diff: forall l l', lt l' (diff_low_bound l) \/ lt (diff_high_bound l) l' -> Loc.diff l l'. Proof. intros. - destruct l as [mr | sl ofs ty]; destruct l' as [mr' | sl' ofs' ty']; simpl in *; auto. + destruct l as [mr | sl ofs q]; destruct l' as [mr' | sl' ofs' q']; simpl in *; auto. - assert (IndexedMreg.index mr <> IndexedMreg.index mr'). { destruct H. apply not_eq_sym. apply Plt_ne; auto. apply Plt_ne; auto. } congruence. - - assert (RANGE: forall ty, 1 <= typesize ty <= 2). - { intros; unfold typesize. destruct ty0; omega. } + - assert (RANGE: forall q, 1 <= typesize q <= 2). + { intros; unfold typesize. destruct q0; omega. } destruct H. + destruct H. left. apply not_eq_sym. apply OrderedSlot.lt_not_eq; auto. destruct H. right. - destruct H0. right. generalize (RANGE ty'); omega. + destruct H0. right. generalize (RANGE (typ_of_quantity q')); omega. destruct H0. - assert (ty' = Tint \/ ty' = Tsingle \/ ty' = Tany32). - { unfold OrderedTyp.lt in H1. destruct ty'; auto; compute in H1; congruence. } - right. destruct H2 as [E|[E|E]]; subst ty'; simpl typesize; omega. + assert (typ_of_quantity q' = Tany32). + { unfold OrderedTyp.lt in H1. destruct q'; auto; compute in H1; congruence. } + right. rewrite H2; simpl typesize; omega. + destruct H. left. apply OrderedSlot.lt_not_eq; auto. destruct H. right. destruct H0. left; omega. - destruct H0. exfalso. destruct ty'; compute in H1; congruence. + destruct H0. exfalso. destruct q'; compute in H1; congruence. Qed. Lemma diff_outside_interval: forall l l', Loc.diff l l' -> lt l' (diff_low_bound l) \/ lt (diff_high_bound l) l'. Proof. intros. - destruct l as [mr | sl ofs ty]; destruct l' as [mr' | sl' ofs' ty']; simpl in *; auto. + destruct l as [mr | sl ofs q]; destruct l' as [mr' | sl' ofs' q']; simpl in *; auto. - unfold Plt, Pos.lt. destruct (Pos.compare (IndexedMreg.index mr) (IndexedMreg.index mr')) eqn:C. elim H. apply IndexedMreg.index_inj. apply Pos.compare_eq_iff. auto. auto. @@ -574,11 +800,8 @@ Module OrderedLoc <: OrderedType. destruct H. right; right; split; auto. left; omega. left; right; split; auto. - assert (EITHER: typesize ty' = 1 /\ OrderedTyp.lt ty' Tany64 \/ typesize ty' = 2). - { destruct ty'; compute; auto. } - destruct (zlt ofs' (ofs - 1)). left; auto. - destruct EITHER as [[P Q] | P]. - right; split; auto. omega. + destruct q'; simpl in H. destruct (zlt ofs' (ofs - 1)). + left; auto. right; split. omega. compute; auto. left; omega. Qed. diff --git a/backend/Mach.v b/backend/Mach.v index 9fdee9ebff..fe2312885e 100644 --- a/backend/Mach.v +++ b/backend/Mach.v @@ -52,9 +52,9 @@ Require Stacklayout. Definition label := positive. Inductive instruction: Type := - | Mgetstack: ptrofs -> typ -> mreg -> instruction - | Msetstack: mreg -> ptrofs -> typ -> instruction - | Mgetparam: ptrofs -> typ -> mreg -> instruction + | Mgetstack: ptrofs -> quantity -> mreg -> instruction + | Msetstack: mreg -> ptrofs -> quantity -> instruction + | Mgetparam: ptrofs -> quantity -> mreg -> instruction | Mop: operation -> list mreg -> mreg -> instruction | Mload: memory_chunk -> addressing -> list mreg -> mreg -> instruction | Mstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction @@ -94,11 +94,11 @@ Definition genv := Genv.t fundef unit. on the interpretation of stack slot accesses. In Mach, these accesses are interpreted as memory accesses relative to the stack pointer. More precisely: -- [Mgetstack ofs ty r] is a memory load at offset [ofs * 4] relative +- [Mgetstack ofs q r] is a memory load at offset [ofs * 4] relative to the stack pointer. -- [Msetstack r ofs ty] is a memory store at offset [ofs * 4] relative +- [Msetstack r ofs q] is a memory store at offset [ofs * 4] relative to the stack pointer. -- [Mgetparam ofs ty r] is a memory load at offset [ofs * 4] +- [Mgetparam ofs q r] is a memory load at offset [ofs * 4] relative to the pointer found at offset 0 from the stack pointer. The semantics maintain a linked structure of activation records, with the current record containing a pointer to the record of the @@ -118,23 +118,15 @@ value of the return address that the Asm code generated later will store in the reserved location. *) -Definition load_stack (m: mem) (sp: val) (ty: typ) (ofs: ptrofs) := - Mem.loadv (chunk_of_type ty) m (Val.offset_ptr sp ofs). +Definition load_stack (m: mem) (sp: val) (q: quantity) (ofs: ptrofs) := + Mem.loadv (chunk_of_quantity q) m (Val.offset_ptr sp ofs). -Definition store_stack (m: mem) (sp: val) (ty: typ) (ofs: ptrofs) (v: val) := - Mem.storev (chunk_of_type ty) m (Val.offset_ptr sp ofs) v. +Definition store_stack (m: mem) (sp: val) (q: quantity) (ofs: ptrofs) (v: val) := + Mem.storev (chunk_of_quantity q) m (Val.offset_ptr sp ofs) v. -Module RegEq. - Definition t := mreg. - Definition eq := mreg_eq. -End RegEq. +Module Regmap := Regfile(Mreg). -Module Regmap := EMap(RegEq). - -Definition regset := Regmap.t val. - -Notation "a ## b" := (List.map a b) (at level 1). -Notation "a # b <- c" := (Regmap.set b c a) (at level 1, b at next level). +Definition regset := Regmap.t. Fixpoint undef_regs (rl: list mreg) (rs: regset) {struct rl} : regset := match rl with @@ -142,34 +134,65 @@ Fixpoint undef_regs (rl: list mreg) (rs: regset) {struct rl} : regset := | r1 :: rl' => Regmap.set r1 Vundef (undef_regs rl' rs) end. +Definition regmap_get (r: mreg) (rs: regset) : val := + Regmap.get r rs. + +Definition regmap_read (rs: regset) : mreg -> val := + fun r => regmap_get r rs. + +Definition regmap_set (r: mreg) (v: val) (rs: regset) : regset := + Regmap.set r v rs. + +Notation "a # b" := (regmap_get b a) (at level 1, b at next level). +Notation "a ## b" := (List.map (regmap_read a) b) (at level 1). +Notation "a # b <- c" := (regmap_set b c a) (at level 1, b at next level). + Lemma undef_regs_other: - forall r rl rs, ~In r rl -> undef_regs rl rs r = rs r. + forall r rl rs, ~In r rl -> (undef_regs rl rs) # r = rs # r. Proof. induction rl; simpl; intros. auto. rewrite Regmap.gso. apply IHrl. intuition. intuition. Qed. Lemma undef_regs_same: - forall r rl rs, In r rl -> undef_regs rl rs r = Vundef. + forall r rl rs, In r rl -> (undef_regs rl rs) # r = Vundef. Proof. induction rl; simpl; intros. tauto. - destruct H. subst a. apply Regmap.gss. - unfold Regmap.set. destruct (RegEq.eq r a); auto. + destruct H. + - subst a. rewrite Regmap.gss. destruct (Mreg.chunk_of r); auto. + - destruct (Mreg.eq_dec r a). + + subst a. rewrite Regmap.gss. destruct (Mreg.chunk_of r); auto. + + rewrite Regmap.gso; auto. Qed. -Definition undef_caller_save_regs (rs: regset) : regset := +Definition undef_caller_save_regs_spec (rs: mreg -> val) : mreg -> val := fun r => if is_callee_save r then rs r else Vundef. +Definition undef_caller_save_regs (rs: regset) : regset := + undef_regs destroyed_at_call rs. + +Lemma undef_caller_save_regs_correct: + forall rs r, + (undef_caller_save_regs rs) # r = undef_caller_save_regs_spec (fun r' => rs # r') r. +Proof. + intros. unfold undef_caller_save_regs, undef_caller_save_regs_spec. + destruct (In_dec mreg_eq r destroyed_at_call) as [IN | NOTIN]. + - rewrite undef_regs_same; auto. + apply LTL.in_destroyed_at_call in IN. rewrite IN; auto. + - rewrite undef_regs_other; auto. + apply LTL.notin_destroyed_at_call in NOTIN. rewrite NOTIN; auto. +Qed. + Definition set_pair (p: rpair mreg) (v: val) (rs: regset) : regset := match p with | One r => rs#r <- v | Twolong rhi rlo => rs#rhi <- (Val.hiword v) #rlo <- (Val.loword v) end. -Fixpoint set_res (res: builtin_res mreg) (v: val) (rs: regset) : regset := +Definition set_res (res: builtin_res mreg) (v: val) (rs: regset) : regset := match res with | BR r => Regmap.set r v rs | BR_none => rs - | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) + | BR_splitlong hi lo => Regmap.set lo (Val.loword v) (Regmap.set hi (Val.hiword v) rs) end. Definition is_label (lbl: label) (instr: instruction) : bool := @@ -215,7 +238,7 @@ Definition find_function_ptr (ge: genv) (ros: mreg + ident) (rs: regset) : option block := match ros with | inl r => - match rs r with + match rs # r with | Vptr b ofs => if Ptrofs.eq ofs Ptrofs.zero then Some b else None | _ => None end @@ -227,10 +250,10 @@ Definition find_function_ptr Inductive extcall_arg (rs: regset) (m: mem) (sp: val): loc -> val -> Prop := | extcall_arg_reg: forall r, - extcall_arg rs m sp (R r) (rs r) - | extcall_arg_stack: forall ofs ty v, - load_stack m sp ty (Ptrofs.repr (Stacklayout.fe_ofs_arg + 4 * ofs)) = Some v -> - extcall_arg rs m sp (S Outgoing ofs ty) v. + extcall_arg rs m sp (R r) (rs # r) + | extcall_arg_stack: forall ofs q v, + load_stack m sp q (Ptrofs.repr (Stacklayout.fe_ofs_arg + 4 * ofs)) = Some v -> + extcall_arg rs m sp (S Outgoing ofs q) v. Inductive extcall_arg_pair (rs: regset) (m: mem) (sp: val): rpair loc -> val -> Prop := | extcall_arg_one: forall l v, @@ -245,6 +268,15 @@ Definition extcall_arguments (rs: regset) (m: mem) (sp: val) (sg: signature) (args: list val) : Prop := list_forall2 (extcall_arg_pair rs m sp) (loc_arguments sg) args. +(** Typing of builtin results. *) + +Definition wt_builtin_res (ty: typ) (res: builtin_res mreg) : bool := + match res with + | BR r => subtype ty (mreg_type r) + | BR_none => true + | BR_splitlong hi lo => subtype Tint (mreg_type hi) && subtype Tint (mreg_type lo) + end. + (** Mach execution states. *) (** Mach execution states. *) @@ -296,23 +328,23 @@ Inductive step: state -> trace -> state -> Prop := step (State s f sp (Mlabel lbl :: c) rs m) E0 (State s f sp c rs m) | exec_Mgetstack: - forall s f sp ofs ty dst c rs m v, - load_stack m sp ty ofs = Some v -> - step (State s f sp (Mgetstack ofs ty dst :: c) rs m) + forall s f sp ofs q dst c rs m v, + load_stack m sp q ofs = Some v -> + step (State s f sp (Mgetstack ofs q dst :: c) rs m) E0 (State s f sp c (rs#dst <- v) m) | exec_Msetstack: - forall s f sp src ofs ty c rs m m' rs', - store_stack m sp ty ofs (rs src) = Some m' -> - rs' = undef_regs (destroyed_by_setstack ty) rs -> - step (State s f sp (Msetstack src ofs ty :: c) rs m) + forall s f sp src ofs q c rs m m' rs', + store_stack m sp q ofs (rs # src) = Some m' -> + rs' = undef_regs (destroyed_by_setstack q) rs -> + step (State s f sp (Msetstack src ofs q :: c) rs m) E0 (State s f sp c rs' m') | exec_Mgetparam: - forall s fb f sp ofs ty dst c rs m v rs', + forall s fb f sp ofs q dst c rs m v rs', Genv.find_funct_ptr ge fb = Some (Internal f) -> - load_stack m sp Tptr f.(fn_link_ofs) = Some (parent_sp s) -> - load_stack m (parent_sp s) ty ofs = Some v -> + load_stack m sp (quantity_of_typ Tptr) f.(fn_link_ofs) = Some (parent_sp s) -> + load_stack m (parent_sp s) q ofs = Some v -> rs' = (rs # temp_for_parent_frame <- Vundef # dst <- v) -> - step (State s fb sp (Mgetparam ofs ty dst :: c) rs m) + step (State s fb sp (Mgetparam ofs q dst :: c) rs m) E0 (State s fb sp c rs' m) | exec_Mop: forall s f sp op args res c rs m v rs', @@ -330,7 +362,7 @@ Inductive step: state -> trace -> state -> Prop := | exec_Mstore: forall s f sp chunk addr args src c rs m m' a rs', eval_addressing ge sp addr rs##args = Some a -> - Mem.storev chunk m a (rs src) = Some m' -> + Mem.storev chunk m a (rs # src) = Some m' -> rs' = undef_regs (destroyed_by_store chunk addr) rs -> step (State s f sp (Mstore chunk addr args src :: c) rs m) E0 (State s f sp c rs' m') @@ -346,15 +378,16 @@ Inductive step: state -> trace -> state -> Prop := forall s fb stk soff sig ros c rs m f f' m', find_function_ptr ge ros rs = Some f' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> - load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp s) -> - load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra s) -> + load_stack m (Vptr stk soff) (quantity_of_typ Tptr) f.(fn_link_ofs) = Some (parent_sp s) -> + load_stack m (Vptr stk soff) (quantity_of_typ Tptr) f.(fn_retaddr_ofs) = Some (parent_ra s) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> step (State s fb (Vptr stk soff) (Mtailcall sig ros :: c) rs m) E0 (Callstate s f' rs m') | exec_Mbuiltin: forall s f sp rs m ef args res b vargs t vres rs' m', - eval_builtin_args ge rs sp m args vargs -> + eval_builtin_args ge (regmap_read rs) sp m args vargs -> external_call ef ge vargs m t vres m' -> + wt_builtin_res (proj_sig_res (ef_sig ef)) res = true -> rs' = set_res res vres (undef_regs (destroyed_by_builtin ef) rs) -> step (State s f sp (Mbuiltin ef args res :: b) rs m) t (State s f sp b rs' m') @@ -380,7 +413,7 @@ Inductive step: state -> trace -> state -> Prop := E0 (State s f sp c rs' m) | exec_Mjumptable: forall s fb f sp arg tbl c rs m n lbl c' rs', - rs arg = Vint n -> + rs # arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some lbl -> Genv.find_funct_ptr ge fb = Some (Internal f) -> find_label lbl f.(fn_code) = Some c' -> @@ -390,8 +423,8 @@ Inductive step: state -> trace -> state -> Prop := | exec_Mreturn: forall s fb stk soff c rs m f m', Genv.find_funct_ptr ge fb = Some (Internal f) -> - load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp s) -> - load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra s) -> + load_stack m (Vptr stk soff) (quantity_of_typ Tptr) f.(fn_link_ofs) = Some (parent_sp s) -> + load_stack m (Vptr stk soff) (quantity_of_typ Tptr) f.(fn_retaddr_ofs) = Some (parent_ra s) -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> step (State s fb (Vptr stk soff) (Mreturn :: c) rs m) E0 (Returnstate s rs m') @@ -400,8 +433,8 @@ Inductive step: state -> trace -> state -> Prop := Genv.find_funct_ptr ge fb = Some (Internal f) -> Mem.alloc m 0 f.(fn_stacksize) = (m1, stk) -> let sp := Vptr stk Ptrofs.zero in - store_stack m1 sp Tptr f.(fn_link_ofs) (parent_sp s) = Some m2 -> - store_stack m2 sp Tptr f.(fn_retaddr_ofs) (parent_ra s) = Some m3 -> + store_stack m1 sp (quantity_of_typ Tptr) f.(fn_link_ofs) (parent_sp s) = Some m2 -> + store_stack m2 sp (quantity_of_typ Tptr) f.(fn_retaddr_ofs) (parent_ra s) = Some m3 -> rs' = undef_regs destroyed_at_function_entry rs -> step (Callstate s fb rs m) E0 (State s fb sp f.(fn_code) rs' m3) @@ -425,12 +458,12 @@ Inductive initial_state (p: program): state -> Prop := let ge := Genv.globalenv p in Genv.init_mem p = Some m0 -> Genv.find_symbol ge p.(prog_main) = Some fb -> - initial_state p (Callstate nil fb (Regmap.init Vundef) m0). + initial_state p (Callstate nil fb Regmap.init m0). Inductive final_state: state -> int -> Prop := | final_state_intro: forall rs m r retcode, loc_result signature_main = One r -> - rs r = Vint retcode -> + rs # r = Vint retcode -> final_state (Returnstate nil rs m) retcode. Definition semantics (rao: function -> code -> ptrofs -> Prop) (p: program) := diff --git a/backend/PrintAsmaux.ml b/backend/PrintAsmaux.ml index f9ed569f19..243822ef17 100644 --- a/backend/PrintAsmaux.ml +++ b/backend/PrintAsmaux.ml @@ -258,7 +258,7 @@ let print_asm_argument print_preg oc modifier = function let builtin_arg_of_res = function | BR r -> BA r - | BR_splitlong(BR hi, BR lo) -> BA_splitlong(BA hi, BA lo) + | BR_splitlong(hi, lo) -> BA_splitlong(BA hi, BA lo) | _ -> assert false let re_asm_param_1 = Str.regexp "%%\\|%[QR]?[0-9]+" diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml index d055707355..61eb989bc6 100644 --- a/backend/PrintLTL.ml +++ b/backend/PrintLTL.ml @@ -31,19 +31,19 @@ let rec mregs pp = function | [r] -> mreg pp r | r1::rl -> fprintf pp "%a, %a" mreg r1 mregs rl -let slot pp (sl, ofs, ty) = +let slot pp (sl, ofs, q) = match sl with | Locations.Local -> - fprintf pp "local(%ld,%s)" (camlint_of_coqint ofs) (name_of_type ty) + fprintf pp "local(%ld,%s)" (camlint_of_coqint ofs) (name_of_quantity q) | Locations.Incoming -> - fprintf pp "incoming(%ld,%s)" (camlint_of_coqint ofs) (name_of_type ty) + fprintf pp "incoming(%ld,%s)" (camlint_of_coqint ofs) (name_of_quantity q) | Locations.Outgoing -> - fprintf pp "outgoing(%ld,%s)" (camlint_of_coqint ofs) (name_of_type ty) + fprintf pp "outgoing(%ld,%s)" (camlint_of_coqint ofs) (name_of_quantity q) let loc pp l = match l with | Locations.R r -> mreg pp r - | Locations.S(sl, ofs, ty) -> slot pp (sl, ofs, ty) + | Locations.S(sl, ofs, q) -> slot pp (sl, ofs, q) let rec locs pp = function | [] -> () diff --git a/backend/PrintMach.ml b/backend/PrintMach.ml index 517f303711..239ab190be 100644 --- a/backend/PrintMach.ml +++ b/backend/PrintMach.ml @@ -36,15 +36,15 @@ let ros pp = function let print_instruction pp i = match i with - | Mgetstack(ofs, ty, res) -> + | Mgetstack(ofs, q, res) -> fprintf pp "\t%a = stack(%ld, %s)\n" - reg res (camlint_of_coqint ofs) (name_of_type ty) - | Msetstack(arg, ofs, ty) -> + reg res (camlint_of_coqint ofs) (name_of_quantity q) + | Msetstack(arg, ofs, q) -> fprintf pp "\tstack(%ld, %s) = %a\n" - (camlint_of_coqint ofs) (name_of_type ty) reg arg - | Mgetparam(ofs, ty, res) -> + (camlint_of_coqint ofs) (name_of_quantity q) reg arg + | Mgetparam(ofs, q, res) -> fprintf pp "\t%a = param(%ld, %s)\n" - reg res (camlint_of_coqint ofs) (name_of_type ty) + reg res (camlint_of_coqint ofs) (name_of_quantity q) | Mop(op, args, res) -> fprintf pp "\t%a = %a\n" reg res (PrintOp.print_operation reg) (op, args) diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml index cc1f7d49cb..dcfef43d0b 100644 --- a/backend/PrintXTL.ml +++ b/backend/PrintXTL.ml @@ -17,6 +17,7 @@ open Camlcoq open Datatypes open Maps open AST +open Memdata open PrintAST open PrintOp open XTL @@ -34,14 +35,18 @@ let short_name_of_type = function | Tany32 -> 'w' | Tany64 -> 'd' +let short_name_of_quantity = function + | Q32 -> 'w' (* "word" *) + | Q64 -> 'd' (* "doubleword" *) + let loc pp = function | Locations.R r -> mreg pp r - | Locations.S(Locations.Local, ofs, ty) -> - fprintf pp "L%c%ld" (short_name_of_type ty) (camlint_of_coqint ofs) - | Locations.S(Locations.Incoming, ofs, ty) -> - fprintf pp "I%c%ld" (short_name_of_type ty) (camlint_of_coqint ofs) - | Locations.S(Locations.Outgoing, ofs, ty) -> - fprintf pp "O%c%ld" (short_name_of_type ty) (camlint_of_coqint ofs) + | Locations.S(Locations.Local, ofs, q) -> + fprintf pp "L%c%ld" (short_name_of_quantity q) (camlint_of_coqint ofs) + | Locations.S(Locations.Incoming, ofs, q) -> + fprintf pp "I%c%ld" (short_name_of_quantity q) (camlint_of_coqint ofs) + | Locations.S(Locations.Outgoing, ofs, q) -> + fprintf pp "O%c%ld" (short_name_of_quantity q) (camlint_of_coqint ofs) let current_alloc = ref (None: (var -> Locations.loc) option) let current_liveness = ref (None: VSet.t PMap.t option) diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml index 19aba4f683..eb21ff9f6f 100644 --- a/backend/Regalloc.ml +++ b/backend/Regalloc.ml @@ -33,6 +33,7 @@ open Datatypes open Coqlib open Maps open AST +open Memdata open Kildall open Op open Machregs @@ -174,7 +175,7 @@ let convert_builtin_res tyenv = function | BR r -> let ty = tyenv r in if Archi.splitlong && ty = Tlong - then BR_splitlong(BR(V(r, Tint)), BR(V(twin_reg r, Tint))) + then BR_splitlong(V(r, Tint), V(twin_reg r, Tint)) else BR(V(r, ty)) | BR_none -> BR_none | BR_splitlong _ -> assert false @@ -201,15 +202,22 @@ let rec constrain_builtin_args al cl = let (al', cl2) = constrain_builtin_args al cl1 in (a' :: al', cl2) -let rec constrain_builtin_res a cl = - match a, cl with - | BR x, None :: cl' -> (a, cl') - | BR x, Some mr :: cl' -> (BR (L(R mr)), cl') - | BR_splitlong(hi, lo), _ -> - let (hi', cl1) = constrain_builtin_res hi cl in - let (lo', cl2) = constrain_builtin_res lo cl1 in +let constrain_builtin_res_var x cl = + match cl with + | None :: cl' -> (x, cl') + | Some mr :: cl' -> (L(R mr), cl') + | _ -> (x, cl) + +let constrain_builtin_res a cl = + match a with + | BR x -> + let (v, cl') = constrain_builtin_res_var x cl in + (BR v, cl') + | BR_splitlong(hi, lo) -> + let (hi', cl1) = constrain_builtin_res_var hi cl in + let (lo', cl2) = constrain_builtin_res_var lo cl1 in (BR_splitlong(hi', lo'), cl2) - | _, _ -> (a, cl) + | _ -> (a, cl) (* Return the XTL basic block corresponding to the given RTL instruction. Move and parallel move instructions are introduced to honor calling @@ -346,11 +354,11 @@ let rec vset_addarg a after = let vset_addargs al after = List.fold_right vset_addarg al after -let rec vset_removeres r after = +let vset_removeres r after = match r with | BR v -> VSet.remove v after | BR_none -> after - | BR_splitlong(hi, lo) -> vset_removeres hi (vset_removeres lo after) + | BR_splitlong(hi, lo) -> VSet.remove hi (VSet.remove lo after) let live_before instr after = match instr with @@ -883,15 +891,15 @@ let save_var tospill eqs v = (t, [Xspill(t, v)], add v t (kill v eqs)) end -let rec save_res tospill eqs = function +let save_res tospill eqs = function | BR v -> let (t, c1, eqs1) = save_var tospill eqs v in (BR t, c1, eqs1) | BR_none -> (BR_none, [], eqs) | BR_splitlong(hi, lo) -> - let (hi', c1, eqs1) = save_res tospill eqs hi in - let (lo', c2, eqs2) = save_res tospill eqs1 lo in + let (hi', c1, eqs1) = save_var tospill eqs hi in + let (lo', c2, eqs2) = save_var tospill eqs1 lo in (BR_splitlong(hi', lo'), c1 @ c2, eqs2) (* Trimming equations when we have too many or when they are too old. @@ -1055,22 +1063,24 @@ let make_parmove srcs dsts itmp ftmp k = let n = Array.length src in assert (Array.length dst = n); let status = Array.make n To_move in - let temp_for cls = + let temp_for_cls cls = match cls with 0 -> itmp | 1 -> ftmp | _ -> assert false in + let temp_for_q q = + match q with Q32 -> itmp | Q64 -> ftmp in let code = ref [] in let add_move s d = match s, d with | R rs, R rd -> code := LTL.Lop(Omove, [rs], rd) :: !code - | R rs, Locations.S(sl, ofs, ty) -> - code := LTL.Lsetstack(rs, sl, ofs, ty) :: !code - | Locations.S(sl, ofs, ty), R rd -> - code := LTL.Lgetstack(sl, ofs, ty, rd) :: !code - | Locations.S(sls, ofss, tys), Locations.S(sld, ofsd, tyd) -> - let tmp = temp_for (class_of_type tys) in + | R rs, Locations.S(sl, ofs, q) -> + code := LTL.Lsetstack(rs, sl, ofs, q) :: !code + | Locations.S(sl, ofs, q), R rd -> + code := LTL.Lgetstack(sl, ofs, q, rd) :: !code + | Locations.S(sls, ofss, qs), Locations.S(sld, ofsd, qd) -> + let tmp = temp_for_q qs in (* code will be reversed at the end *) - code := LTL.Lsetstack(tmp, sld, ofsd, tyd) :: - LTL.Lgetstack(sls, ofss, tys, tmp) :: !code + code := LTL.Lsetstack(tmp, sld, ofsd, qd) :: + LTL.Lgetstack(sls, ofss, qs, tmp) :: !code in let rec move_one i = if src.(i) <> dst.(i) then begin @@ -1081,7 +1091,7 @@ let make_parmove srcs dsts itmp ftmp k = | To_move -> move_one j | Being_moved -> - let tmp = R (temp_for (class_of_loc src.(j))) in + let tmp = R (temp_for_cls (class_of_loc src.(j))) in add_move src.(j) tmp; src.(j) <- tmp | Moved -> diff --git a/backend/Registerfile.v b/backend/Registerfile.v new file mode 100644 index 0000000000..2f36fd7357 --- /dev/null +++ b/backend/Registerfile.v @@ -0,0 +1,342 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Gergö Barany, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +Require Import OrderedType. +Require Import Coqlib. +Require Import Maps. +Require Import Ordered. +Require Import AST. +Require Import Values. +Require Import Memdata. +Require Import Memory. +Require Import FragmentBlock. +Require Export Machregs. + +Open Scope Z_scope. + +(** * Auxiliary definitions *) + +Definition quantity_of_mreg (r: mreg) : quantity := + quantity_of_typ (mreg_type r). + +Definition chunk_of_mreg (r: mreg) : memory_chunk := + chunk_of_type (mreg_type r). + +Definition typesize ty := + match ty with + | Tint | Tsingle | Tany32 => 1 + | Tlong | Tfloat | Tany64 => 2 + end. + +Lemma typesize_pos: + forall ty, typesize ty > 0. +Proof. + destruct ty; simpl; omega. +Qed. + +(** * Definition of a general register model *) + +(** We will be using slightly different models of the target CPU's register + architecture: [mreg] models data registers only, with no separation of + different kinds of registers, while [preg] models integer and floating-point + registers differently and includes condition code registers and the program + counter as well. + + We will want to use the same basic representation of the register file for + both of these register models, so we capture their commonalities in this + module type. *) + +Module Type REGISTER_MODEL. + + Parameter reg: Set. + Axiom eq_dec: forall r s: reg, {r = s} + {r <> s}. + + Parameter type: reg -> typ. + Parameter quantity_of: reg -> quantity. + Parameter chunk_of: reg -> memory_chunk. + + Axiom type_cases: + forall r, type r = Tany32 \/ type r = Tany64. + + Parameter ofs: reg -> Z. + Parameter addr: reg -> Z. + Parameter next_addr: reg -> Z. + + Axiom addr_pos: forall r, addr r > 0. + + Axiom addr_compat: forall r, FragBlock.addr (ofs r) = addr r. + Axiom next_addr_compat: forall r, FragBlock.next_addr (ofs r) (quantity_of r) = next_addr r. + + Axiom size_compat: + forall r, + AST.typesize (type r) = FragBlock.quantity_size (quantity_of r). + + Axiom quantity_of_compat: + forall r, + quantity_of r = quantity_of_typ (type r). + + Axiom chunk_of_reg_compat: + forall r, + chunk_of r = chunk_of_quantity (quantity_of r). + + Axiom chunk_of_reg_type: + forall r, + chunk_of r = chunk_of_type (type r). + + Axiom diff_outside_interval: + forall r s, r <> s -> next_addr r <= addr s \/ next_addr s <= addr r. + +End REGISTER_MODEL. + +Module Mreg <: REGISTER_MODEL. + + Definition reg := mreg. + Definition eq_dec := mreg_eq. + + Definition type := mreg_type. + Definition quantity_of := quantity_of_mreg. + Definition chunk_of := chunk_of_mreg. + + Definition type_cases := mreg_type_cases. + + (* A register's offset, in words, from an arbitrary starting point. *) + Definition ofs (r: mreg) : Z := + Zpos (IndexedMreg.index r). + + (* A register's address: The index of its first byte. *) + Definition addr (r: mreg) : Z := Zpos (IndexedMreg.index r) * 4. + + Remark addr_pos: forall r, addr r > 0. + Proof. + intros. unfold addr. xomega. + Qed. + + (* The address one byte past the end of register [r]. The next nonoverlapping + register may start here. *) + Definition next_addr (r: mreg) : Z := addr r + AST.typesize (mreg_type r). + + (* Our notions of offset and address are compatible with FragBlock's addresses. *) + Lemma addr_compat: forall r, FragBlock.addr (ofs r) = addr r. + Proof. + unfold addr, ofs, FragBlock.addr; intros. auto. + Qed. + + Lemma size_compat: + forall r, + AST.typesize (mreg_type r) = FragBlock.quantity_size (quantity_of_mreg r). + Proof. + intros. unfold quantity_of_mreg. destruct (mreg_type r); simpl; auto. + Qed. + + Lemma quantity_of_compat: + forall r, + quantity_of r = quantity_of_typ (type r). + Proof. + reflexivity. + Qed. + + Lemma chunk_of_reg_type: + forall r, + chunk_of r = chunk_of_type (type r). + Proof. + reflexivity. + Qed. + + Lemma chunk_of_reg_compat: + forall r, + chunk_of r = chunk_of_quantity (quantity_of_mreg r). + Proof. + intros. + rewrite quantity_of_compat, chunk_of_quantity_of_typ, chunk_of_reg_type; auto. + apply type_cases. + Qed. + + Lemma next_addr_compat: forall r, FragBlock.next_addr (ofs r) (quantity_of_mreg r) = next_addr r. + Proof. + unfold next_addr, addr, ofs, FragBlock.next_addr, FragBlock.addr; intros. + rewrite size_compat. auto. + Qed. + + Lemma address_lt: + forall r s, + r <> s -> + addr r < addr s -> + next_addr r <= addr s. + Proof. + intros. unfold addr in H0. + apply IndexedMreg.scaled_index_with_size; omega. + Qed. + + Lemma outside_interval_diff: + forall r s, next_addr r <= addr s \/ next_addr s <= addr r -> r <> s. + Proof. + intros. destruct H; unfold next_addr in H. + generalize (AST.typesize_pos (mreg_type r)); intros. contradict H. subst. omega. + generalize (AST.typesize_pos (mreg_type s)); intros. contradict H. subst. omega. + Qed. + + Lemma diff_outside_interval: + forall r s, r <> s -> next_addr r <= addr s \/ next_addr s <= addr r. + Proof. + intros. + assert (Neq: addr r <> addr s). + { unfold addr. contradict H. apply IndexedMreg.index_inj. xomega. } + assert (Cases: addr r < addr s \/ addr s < addr r) by omega. + destruct Cases. + - left. apply address_lt; auto. + - right. apply address_lt; auto. + Qed. + +End Mreg. + +(** * Representation of the register file *) + +(** The [Regfile] module defines mappings from registers to values. The register + file is represented as a kind of memory block containing bytes addressed using + register numbers. + + The register offset numbers are interpreted as words and scaled internally to + bytes. The indices of adjacent 64-bit registers must therefore differ by at + least 2. *) + +Module Regfile(M: REGISTER_MODEL). + + Definition t := FragBlock.t. + + Definition init := FragBlock.init. + + (* The offset just past the end of register [r]. The next nonoverlapping + register may start here. *) + Definition next_ofs (r: M.reg) : Z := M.ofs r + typesize (M.type r). + + Lemma chunk_length: + forall r v, + Z.to_nat (AST.typesize (mreg_type r)) = length (encode_val (chunk_of_mreg r) v). + Proof. + intros. rewrite encode_val_length. + unfold chunk_of_mreg. destruct (mreg_type r); auto. + Qed. + + Definition get_bytes (r: M.reg) (rf: t) : list memval := + FragBlock.get_bytes (M.ofs r) (M.quantity_of r) rf. + + Definition get (r: M.reg) (rf: t) : val := + FragBlock.get (M.ofs r) (M.quantity_of r) rf. + + Definition set_bytes (r: M.reg) (bytes: list memval) (rf: t) : t := + FragBlock.set_bytes (M.ofs r) (M.quantity_of r) bytes rf. + + Definition set (r: M.reg) (v: val) (rf: t) : t := + FragBlock.set (M.ofs r) (M.quantity_of r) v rf. + + (* Update the [old] register file by choosing the values for the registers in + [rs] from [new]. *) + Fixpoint override (rs: list M.reg) (new old: t) : t := + match rs with + | nil => old + | r :: rs' => set r (get r new) (override rs' new old) + end. + + Fixpoint undef_regs (rs: list M.reg) (rf: t) : t := + match rs with + | nil => rf + | r :: rs => set r Vundef (undef_regs rs rf) + end. + + Lemma gss: + forall r v rf, + get r (set r v rf) = Val.load_result (M.chunk_of r) v. + Proof. + intros. unfold get, set. + rewrite FragBlock.gss, M.chunk_of_reg_compat; auto. + Qed. + + Lemma gso: + forall r s v rf, + r <> s -> + get r (set s v rf) = get r rf. + Proof. + intros. unfold get, set. + rewrite FragBlock.gso; auto. + rewrite !M.next_addr_compat, !M.addr_compat. + apply M.diff_outside_interval; auto. + Qed. + + Lemma get_has_type: + forall r rf, Val.has_type (get r rf) (M.type r). + Proof. + intros. unfold get. + unfold quantity_of_mreg. + rewrite M.quantity_of_compat. + destruct (M.type_cases r) as [T | T]; rewrite T; apply FragBlock.get_has_type. + Qed. + + Lemma get_load_result: + forall r rf, Val.load_result (M.chunk_of r) (get r rf) = get r rf. + Proof. + intros. rewrite M.chunk_of_reg_type, Val.load_result_same; auto using get_has_type. + Qed. + + Lemma get_bytes_compat: + forall r rf, get r rf = decode_val (M.chunk_of r) (get_bytes r rf). + Proof. + intros. unfold get, FragBlock.get. rewrite M.chunk_of_reg_compat. reflexivity. + Qed. + + Lemma set_bytes_compat: + forall r v rf, set r v rf = set_bytes r (encode_val (M.chunk_of r) v) rf. + Proof. + intros. unfold set, FragBlock.set. rewrite M.chunk_of_reg_compat. reflexivity. + Qed. + + Lemma override_in: + forall rs r new old, + In r rs -> get r (override rs new old) = get r new. + Proof. + intros. induction rs; try contradiction. + destruct (M.eq_dec r a). + - subst; simpl; rewrite gss. + rewrite M.chunk_of_reg_type, Val.load_result_same; auto. + apply get_has_type. + - inversion H; try congruence. + simpl; rewrite gso; auto. + Qed. + + Lemma override_notin: + forall r rs new old, + ~ In r rs -> get r (override rs new old) = get r old. + Proof. + intros. induction rs; auto. + apply not_in_cons in H. simpl. rewrite gso; intuition auto. + Qed. + + Lemma undef_regs_in: + forall r rs rf, + In r rs -> get r (undef_regs rs rf) = Vundef. + Proof. + induction rs; simpl; intros. contradiction. + destruct (M.eq_dec r a). + - subst; simpl; rewrite gss. + destruct (M.chunk_of a); auto. + - inversion H; try congruence. + simpl; rewrite gso; auto. + Qed. + + Lemma undef_regs_notin: + forall r rs rf, + ~ In r rs -> get r (undef_regs rs rf) = get r rf. + Proof. + induction rs; auto; intros. + apply not_in_cons in H. simpl. rewrite gso; intuition auto. + Qed. + +End Regfile. diff --git a/backend/Stacking.v b/backend/Stacking.v index 7b382d0594..34662c5a0d 100644 --- a/backend/Stacking.v +++ b/backend/Stacking.v @@ -39,7 +39,7 @@ Fixpoint save_callee_save_rec (rl: list mreg) (ofs: Z) (k: Mach.code) := let ty := mreg_type r in let sz := AST.typesize ty in let ofs1 := align ofs sz in - Msetstack r (Ptrofs.repr ofs1) ty :: save_callee_save_rec rl (ofs1 + sz) k + Msetstack r (Ptrofs.repr ofs1) (quantity_of_typ ty) :: save_callee_save_rec rl (ofs1 + sz) k end. Definition save_callee_save (fe: frame_env) (k: Mach.code) := @@ -56,7 +56,7 @@ Fixpoint restore_callee_save_rec (rl: list mreg) (ofs: Z) (k: Mach.code) := let ty := mreg_type r in let sz := AST.typesize ty in let ofs1 := align ofs sz in - Mgetstack (Ptrofs.repr ofs1) ty r :: restore_callee_save_rec rl (ofs1 + sz) k + Mgetstack (Ptrofs.repr ofs1) (quantity_of_typ ty) r :: restore_callee_save_rec rl (ofs1 + sz) k end. Definition restore_callee_save (fe: frame_env) (k: Mach.code) := @@ -82,8 +82,8 @@ Definition transl_addr (fe: frame_env) (addr: addressing) := Fixpoint transl_builtin_arg (fe: frame_env) (a: builtin_arg loc) : builtin_arg mreg := match a with | BA (R r) => BA r - | BA (S Local ofs ty) => - BA_loadstack (chunk_of_type ty) (Ptrofs.repr (offset_local fe ofs)) + | BA (S Local ofs q) => + BA_loadstack (chunk_of_quantity q) (Ptrofs.repr (offset_local fe ofs)) | BA (S _ _ _) => BA_int Int.zero (**r never happens *) | BA_int n => BA_int n | BA_long n => BA_long n @@ -113,23 +113,23 @@ Fixpoint transl_builtin_arg (fe: frame_env) (a: builtin_arg loc) : builtin_arg m Definition transl_instr (fe: frame_env) (i: Linear.instruction) (k: Mach.code) : Mach.code := match i with - | Lgetstack sl ofs ty r => + | Lgetstack sl ofs q r => match sl with | Local => - Mgetstack (Ptrofs.repr (offset_local fe ofs)) ty r :: k + Mgetstack (Ptrofs.repr (offset_local fe ofs)) q r :: k | Incoming => - Mgetparam (Ptrofs.repr (offset_arg ofs)) ty r :: k + Mgetparam (Ptrofs.repr (offset_arg ofs)) q r :: k | Outgoing => - Mgetstack (Ptrofs.repr (offset_arg ofs)) ty r :: k + Mgetstack (Ptrofs.repr (offset_arg ofs)) q r :: k end - | Lsetstack r sl ofs ty => + | Lsetstack r sl ofs q => match sl with | Local => - Msetstack r (Ptrofs.repr (offset_local fe ofs)) ty :: k + Msetstack r (Ptrofs.repr (offset_local fe ofs)) q :: k | Incoming => k (* should not happen *) | Outgoing => - Msetstack r (Ptrofs.repr (offset_arg ofs)) ty :: k + Msetstack r (Ptrofs.repr (offset_arg ofs)) q :: k end | Lop op args res => Mop (transl_op fe op) args res :: k diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v index ffd9b2278d..5b334f4e3f 100644 --- a/backend/Stackingproof.v +++ b/backend/Stackingproof.v @@ -18,7 +18,8 @@ Require Import Coqlib Errors. Require Import Integers AST Linking. Require Import Values Memory Separation Events Globalenvs Smallstep. Require Import LTL Op Locations Linear Mach. -Require Import Bounds Conventions Stacklayout Lineartyping. +Require Import Bounds Conventions Conventions1 Stacklayout Lineartyping. +Require Import FragmentBlock. Require Import Stacking. Local Open Scope sep_scope. @@ -52,6 +53,12 @@ Proof. destruct ty; reflexivity. Qed. +Remark size_quantity_chunk: + forall q, size_chunk (chunk_of_quantity q) = size_quantity q. +Proof. + destruct q; reflexivity. +Qed. + Lemma slot_outgoing_argument_valid: forall f ofs ty sg, In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments sg)) -> slot_valid f Outgoing ofs ty = true. @@ -59,8 +66,7 @@ Proof. intros. exploit loc_arguments_acceptable_2; eauto. intros [A B]. unfold slot_valid. unfold proj_sumbool. rewrite zle_true by omega. - rewrite pred_dec_true by auto. - auto. + destruct ty; simpl; rewrite pred_dec_true; auto. Qed. Lemma load_result_inject: @@ -143,9 +149,9 @@ Local Opaque Z.add Z.mul Z.divide. (** Accessing the stack frame using [load_stack] and [store_stack]. *) Lemma contains_get_stack: - forall spec m ty sp ofs, - m |= contains (chunk_of_type ty) sp ofs spec -> - exists v, load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr ofs) = Some v /\ spec v. + forall spec m q sp ofs, + m |= contains (chunk_of_quantity q) sp ofs spec -> + exists v, load_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr ofs) = Some v /\ spec v. Proof. intros. unfold load_stack. replace (Val.offset_ptr (Vptr sp Ptrofs.zero) (Ptrofs.repr ofs)) with (Vptr sp (Ptrofs.repr ofs)). @@ -154,20 +160,20 @@ Proof. Qed. Lemma hasvalue_get_stack: - forall ty m sp ofs v, - m |= hasvalue (chunk_of_type ty) sp ofs v -> - load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr ofs) = Some v. + forall q m sp ofs v, + m |= hasvalue (chunk_of_quantity q) sp ofs v -> + load_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr ofs) = Some v. Proof. intros. exploit contains_get_stack; eauto. intros (v' & A & B). congruence. Qed. Lemma contains_set_stack: - forall (spec: val -> Prop) v spec1 m ty sp ofs P, - m |= contains (chunk_of_type ty) sp ofs spec1 ** P -> - spec (Val.load_result (chunk_of_type ty) v) -> + forall (spec: val -> Prop) v spec1 m q sp ofs P, + m |= contains (chunk_of_quantity q) sp ofs spec1 ** P -> + spec (Val.load_result (chunk_of_quantity q) v) -> exists m', - store_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr ofs) v = Some m' - /\ m' |= contains (chunk_of_type ty) sp ofs spec ** P. + store_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr ofs) v = Some m' + /\ m' |= contains (chunk_of_quantity q) sp ofs spec ** P. Proof. intros. unfold store_stack. replace (Val.offset_ptr (Vptr sp Ptrofs.zero) (Ptrofs.repr ofs)) with (Vptr sp (Ptrofs.repr ofs)). @@ -188,9 +194,9 @@ Program Definition contains_locations (j: meminj) (sp: block) (pos bound: Z) (sl m_pred := fun m => (8 | pos) /\ 0 <= pos /\ pos + 4 * bound <= Ptrofs.modulus /\ Mem.range_perm m sp pos (pos + 4 * bound) Cur Freeable /\ - forall ofs ty, 0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) -> - exists v, Mem.load (chunk_of_type ty) m sp (pos + 4 * ofs) = Some v - /\ Val.inject j (ls (S sl ofs ty)) v; + forall ofs q, 0 <= ofs -> ofs + typesize (typ_of_quantity q) <= bound -> (typealign (typ_of_quantity q) | ofs) -> + exists v, Mem.load (chunk_of_quantity q) m sp (pos + 4 * ofs) = Some v + /\ Val.inject j (ls @ (S sl ofs q)) v; m_footprint := fun b ofs => b = sp /\ pos <= ofs < pos + 4 * bound |}. @@ -199,51 +205,51 @@ Next Obligation. - red; intros. eapply Mem.perm_unchanged_on; eauto. simpl; auto. - exploit H4; eauto. intros (v & A & B). exists v; split; auto. eapply Mem.load_unchanged_on; eauto. - simpl; intros. rewrite size_type_chunk, typesize_typesize in H8. - split; auto. omega. + simpl; intros. + split; auto. destruct q; simpl in *; omega. Qed. Next Obligation. eauto with mem. Qed. Remark valid_access_location: - forall m sp pos bound ofs ty p, + forall m sp pos bound ofs q p, (8 | pos) -> Mem.range_perm m sp pos (pos + 4 * bound) Cur Freeable -> - 0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) -> - Mem.valid_access m (chunk_of_type ty) sp (pos + 4 * ofs) p. + 0 <= ofs -> ofs + typesize (typ_of_quantity q) <= bound -> (typealign (typ_of_quantity q) | ofs) -> + Mem.valid_access m (chunk_of_quantity q) sp (pos + 4 * ofs) p. Proof. intros; split. - red; intros. apply Mem.perm_implies with Freeable; auto with mem. - apply H0. rewrite size_type_chunk, typesize_typesize in H4. omega. -- rewrite align_type_chunk. apply Z.divide_add_r. - apply Z.divide_trans with 8; auto. - exists (8 / (4 * typealign ty)); destruct ty; reflexivity. - apply Z.mul_divide_mono_l. auto. + apply H0. destruct q; simpl in *; omega. +- replace (align_chunk _) with 4 by (destruct q; auto). + apply Z.divide_add_r. + apply Z.divide_trans with 8; auto. exists 2; auto. + apply Z.divide_factor_l. Qed. Lemma get_location: - forall m j sp pos bound sl ls ofs ty, + forall m j sp pos bound sl ls ofs q, m |= contains_locations j sp pos bound sl ls -> - 0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) -> + 0 <= ofs -> ofs + typesize (typ_of_quantity q) <= bound -> (typealign (typ_of_quantity q) | ofs) -> exists v, - load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (pos + 4 * ofs)) = Some v - /\ Val.inject j (ls (S sl ofs ty)) v. + load_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr (pos + 4 * ofs)) = Some v + /\ Val.inject j (ls @ (S sl ofs q)) v. Proof. intros. destruct H as (D & E & F & G & H). exploit H; eauto. intros (v & U & V). exists v; split; auto. unfold load_stack; simpl. rewrite Ptrofs.add_zero_l, Ptrofs.unsigned_repr; auto. - unfold Ptrofs.max_unsigned. generalize (typesize_pos ty). omega. + unfold Ptrofs.max_unsigned. generalize (typesize_pos (typ_of_quantity q)). omega. Qed. Lemma set_location: - forall m j sp pos bound sl ls P ofs ty v v', + forall m j sp pos bound sl ls P ofs q v v', m |= contains_locations j sp pos bound sl ls ** P -> - 0 <= ofs -> ofs + typesize ty <= bound -> (typealign ty | ofs) -> + 0 <= ofs -> ofs + typesize (typ_of_quantity q) <= bound -> (typealign (typ_of_quantity q) | ofs) -> Val.inject j v v' -> exists m', - store_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (pos + 4 * ofs)) v' = Some m' - /\ m' |= contains_locations j sp pos bound sl (Locmap.set (S sl ofs ty) v ls) ** P. + store_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr (pos + 4 * ofs)) v' = Some m' + /\ m' |= contains_locations j sp pos bound sl (Locmap.set (S sl ofs q) v ls) ** P. Proof. intros. destruct H as (A & B & C). destruct A as (D & E & F & G & H). edestruct Mem.valid_access_store as [m' STORE]. @@ -252,55 +258,56 @@ Proof. { red; intros; eauto with mem. } exists m'; split. - unfold store_stack; simpl. rewrite Ptrofs.add_zero_l, Ptrofs.unsigned_repr; eauto. - unfold Ptrofs.max_unsigned. generalize (typesize_pos ty). omega. -- simpl. intuition auto. -+ unfold Locmap.set. - destruct (Loc.eq (S sl ofs ty) (S sl ofs0 ty0)); [|destruct (Loc.diff_dec (S sl ofs ty) (S sl ofs0 ty0))]. + unfold Ptrofs.max_unsigned. generalize (typesize_pos (typ_of_quantity q)). omega. +- set (LS := Locmap.set (S sl ofs q) v ls). simpl. intuition auto. fold (LS @ (S sl ofs0 q0)). ++ subst LS. unfold Locmap.get, Locmap.set. destruct ls. + destruct (Loc.eq (S sl ofs q) (S sl ofs0 q0)); [|destruct (Loc.diff_dec (S sl ofs q) (S sl ofs0 q0))]. * (* same location *) - inv e. rename ofs0 into ofs. rename ty0 into ty. - exists (Val.load_result (chunk_of_type ty) v'); split. - eapply Mem.load_store_similar_2; eauto. omega. - apply Val.load_result_inject; auto. + inv e. rename ofs0 into ofs. rename q0 into q. + exists (Val.load_result (chunk_of_quantity q) v'). split. + exploit Mem.load_store_similar_2; eauto. omega. + rewrite Stack.gss. auto using Val.load_result_inject. * (* different locations *) exploit H; eauto. intros (v0 & X & Y). exists v0; split; auto. rewrite <- X; eapply Mem.load_store_other; eauto. - destruct d. congruence. right. rewrite ! size_type_chunk, ! typesize_typesize. omega. + destruct d. congruence. right. destruct q, q0; simpl in *; omega. + rewrite Stack.gso; auto using Loc.diff_sym. * (* overlapping locations *) - destruct (Mem.valid_access_load m' (chunk_of_type ty0) sp (pos + 4 * ofs0)) as [v'' LOAD]. + destruct (Mem.valid_access_load m' (chunk_of_quantity q0) sp (pos + 4 * ofs0)) as [v'' LOAD]. apply Mem.valid_access_implies with Writable; auto with mem. eapply valid_access_location; eauto. - exists v''; auto. + exists v''; rewrite Stack.gu_overlap; auto using Loc.diff_sym. + apply (m_invar P) with m; auto. eapply Mem.store_unchanged_on; eauto. - intros i; rewrite size_type_chunk, typesize_typesize. intros; red; intros. - eelim C; eauto. simpl. split; auto. omega. + intros i. intros; red; intros. + eelim C; eauto. simpl. split; auto. destruct q; simpl in *; omega. Qed. Lemma initial_locations: forall j sp pos bound P sl ls m, m |= range sp pos (pos + 4 * bound) ** P -> (8 | pos) -> - (forall ofs ty, ls (S sl ofs ty) = Vundef) -> + (forall ofs q, ls @ (S sl ofs q) = Vundef) -> m |= contains_locations j sp pos bound sl ls ** P. Proof. intros. destruct H as (A & B & C). destruct A as (D & E & F). split. - simpl; intuition auto. red; intros; eauto with mem. - destruct (Mem.valid_access_load m (chunk_of_type ty) sp (pos + 4 * ofs)) as [v LOAD]. + destruct (Mem.valid_access_load m (chunk_of_quantity q) sp (pos + 4 * ofs)) as [v LOAD]. eapply valid_access_location; eauto. red; intros; eauto with mem. - exists v; split; auto. rewrite H1; auto. + exists v; split; auto. fold (ls @ (S sl ofs q)); rewrite H1; auto. - split; assumption. Qed. Lemma contains_locations_exten: forall ls ls' j sp pos bound sl, - (forall ofs ty, Val.lessdef (ls' (S sl ofs ty)) (ls (S sl ofs ty))) -> + (forall ofs q, Val.lessdef (ls' @ (S sl ofs q)) (ls @ (S sl ofs q))) -> massert_imp (contains_locations j sp pos bound sl ls) (contains_locations j sp pos bound sl ls'). Proof. intros; split; simpl; intros; auto. intuition auto. exploit H5; eauto. intros (v & A & B). exists v; split; auto. - specialize (H ofs ty). inv H. congruence. auto. + specialize (H ofs q). inv H. congruence. auto. Qed. Lemma contains_locations_incr: @@ -326,7 +333,7 @@ Fixpoint contains_callee_saves (j: meminj) (sp: block) (pos: Z) (rl: list mreg) let ty := mreg_type r in let sz := AST.typesize ty in let pos1 := align pos sz in - contains (chunk_of_type ty) sp pos1 (fun v => Val.inject j (ls (R r)) v) + contains (chunk_of_type ty) sp pos1 (fun v => Val.inject j (ls @ (R r)) v) ** contains_callee_saves j sp (pos1 + sz) rl ls end. @@ -344,13 +351,13 @@ Qed. Lemma contains_callee_saves_exten: forall j sp ls ls' rl pos, - (forall r, In r rl -> ls' (R r) = ls (R r)) -> + (forall r, In r rl -> ls' @ (R r) = ls @ (R r)) -> massert_eqv (contains_callee_saves j sp pos rl ls) (contains_callee_saves j sp pos rl ls'). Proof. induction rl as [ | r1 rl]; simpl; intros. - reflexivity. -- apply sepconj_morph_2; auto. rewrite H by auto. reflexivity. +- unfold Locmap.get in IHrl. apply sepconj_morph_2; auto. rewrite H by auto. reflexivity. Qed. (** Separation logic assertions describing the stack frame at [sp]. @@ -369,8 +376,8 @@ represents the Linear stack data. *) Definition frame_contents_1 (j: meminj) (sp: block) (ls ls0: locset) (parent retaddr: val) := contains_locations j sp fe.(fe_ofs_local) b.(bound_local) Local ls ** contains_locations j sp fe_ofs_arg b.(bound_outgoing) Outgoing ls - ** hasvalue Mptr sp fe.(fe_ofs_link) parent - ** hasvalue Mptr sp fe.(fe_ofs_retaddr) retaddr + ** hasvalue Mptr_any sp fe.(fe_ofs_link) parent + ** hasvalue Mptr_any sp fe.(fe_ofs_retaddr) retaddr ** contains_callee_saves j sp fe.(fe_ofs_callee_save) b.(used_callee_save) ls0. Definition frame_contents (j: meminj) (sp: block) (ls ls0: locset) (parent retaddr: val) := @@ -381,12 +388,12 @@ Definition frame_contents (j: meminj) (sp: block) (ls ls0: locset) (parent retad (** Accessing components of the frame. *) Lemma frame_get_local: - forall ofs ty j sp ls ls0 parent retaddr m P, + forall ofs q j sp ls ls0 parent retaddr m P, m |= frame_contents j sp ls ls0 parent retaddr ** P -> - slot_within_bounds b Local ofs ty -> slot_valid f Local ofs ty = true -> + slot_within_bounds b Local ofs q -> slot_valid f Local ofs q = true -> exists v, - load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (offset_local fe ofs)) = Some v - /\ Val.inject j (ls (S Local ofs ty)) v. + load_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr (offset_local fe ofs)) = Some v + /\ Val.inject j (ls @ (S Local ofs q)) v. Proof. unfold frame_contents, frame_contents_1; intros. unfold slot_valid in H1; InvBooleans. apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_proj1 in H. @@ -394,12 +401,12 @@ Proof. Qed. Lemma frame_get_outgoing: - forall ofs ty j sp ls ls0 parent retaddr m P, + forall ofs q j sp ls ls0 parent retaddr m P, m |= frame_contents j sp ls ls0 parent retaddr ** P -> - slot_within_bounds b Outgoing ofs ty -> slot_valid f Outgoing ofs ty = true -> + slot_within_bounds b Outgoing ofs q -> slot_valid f Outgoing ofs q = true -> exists v, - load_stack m (Vptr sp Ptrofs.zero) ty (Ptrofs.repr (offset_arg ofs)) = Some v - /\ Val.inject j (ls (S Outgoing ofs ty)) v. + load_stack m (Vptr sp Ptrofs.zero) q (Ptrofs.repr (offset_arg ofs)) = Some v + /\ Val.inject j (ls @ (S Outgoing ofs q)) v. Proof. unfold frame_contents, frame_contents_1; intros. unfold slot_valid in H1; InvBooleans. apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick2 in H. @@ -409,21 +416,21 @@ Qed. Lemma frame_get_parent: forall j sp ls ls0 parent retaddr m P, m |= frame_contents j sp ls ls0 parent retaddr ** P -> - load_stack m (Vptr sp Ptrofs.zero) Tptr (Ptrofs.repr fe.(fe_ofs_link)) = Some parent. + load_stack m (Vptr sp Ptrofs.zero) (quantity_of_typ Tptr) (Ptrofs.repr fe.(fe_ofs_link)) = Some parent. Proof. unfold frame_contents, frame_contents_1; intros. - apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick3 in H. rewrite <- chunk_of_Tptr in H. - eapply hasvalue_get_stack; eauto. + apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick3 in H. (*rewrite <- chunk_of_Tptr in H.*) + eapply hasvalue_get_stack; eauto; rewrite chunk_of_quantity_of_Tptr; auto. Qed. Lemma frame_get_retaddr: forall j sp ls ls0 parent retaddr m P, m |= frame_contents j sp ls ls0 parent retaddr ** P -> - load_stack m (Vptr sp Ptrofs.zero) Tptr (Ptrofs.repr fe.(fe_ofs_retaddr)) = Some retaddr. + load_stack m (Vptr sp Ptrofs.zero) (quantity_of_typ Tptr) (Ptrofs.repr fe.(fe_ofs_retaddr)) = Some retaddr. Proof. unfold frame_contents, frame_contents_1; intros. - apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick4 in H. rewrite <- chunk_of_Tptr in H. - eapply hasvalue_get_stack; eauto. + apply mconj_proj1 in H. apply sep_proj1 in H. apply sep_pick4 in H. (*rewrite <- chunk_of_Tptr in H.*) + eapply hasvalue_get_stack; eauto; rewrite chunk_of_quantity_of_Tptr; auto. Qed. (** Assigning a [Local] or [Outgoing] stack slot. *) @@ -446,7 +453,7 @@ Proof. assert (forall i k p, Mem.perm m sp i k p -> Mem.perm m' sp i k p). { intros. unfold store_stack in A; simpl in A. eapply Mem.perm_store_1; eauto. } eapply frame_mconj. eauto. - unfold frame_contents_1; rewrite ! sep_assoc; exact B. + unfold frame_contents_1; rewrite ! sep_assoc; destruct ls; exact B. eapply sep_preserved. eapply sep_proj1. eapply mconj_proj2. eassumption. intros; eapply range_preserved; eauto. @@ -471,7 +478,7 @@ Proof. assert (forall i k p, Mem.perm m sp i k p -> Mem.perm m' sp i k p). { intros. unfold store_stack in A; simpl in A. eapply Mem.perm_store_1; eauto. } eapply frame_mconj. eauto. - unfold frame_contents_1; rewrite ! sep_assoc, sep_swap; eauto. + unfold frame_contents_1; rewrite ! sep_assoc, sep_swap; destruct ls; eauto. eapply sep_preserved. eapply sep_proj1. eapply mconj_proj2. eassumption. intros; eapply range_preserved; eauto. @@ -482,9 +489,9 @@ Qed. Lemma frame_contents_exten: forall ls ls0 ls' ls0' j sp parent retaddr P m, - (forall ofs ty, Val.lessdef (ls' (S Local ofs ty)) (ls (S Local ofs ty))) -> - (forall ofs ty, Val.lessdef (ls' (S Outgoing ofs ty)) (ls (S Outgoing ofs ty))) -> - (forall r, In r b.(used_callee_save) -> ls0' (R r) = ls0 (R r)) -> + (forall ofs ty, Val.lessdef (ls' @ (S Local ofs ty)) (ls @ (S Local ofs ty))) -> + (forall ofs ty, Val.lessdef (ls' @ (S Outgoing ofs ty)) (ls @ (S Outgoing ofs ty))) -> + (forall r, In r b.(used_callee_save) -> ls0' @ (R r) = ls0 @ (R r)) -> m |= frame_contents j sp ls ls0 parent retaddr ** P -> m |= frame_contents j sp ls' ls0' parent retaddr ** P. Proof. @@ -502,6 +509,8 @@ Corollary frame_set_reg: m |= frame_contents j sp (Locmap.set (R r) v ls) ls0 parent retaddr ** P. Proof. intros. apply frame_contents_exten with ls ls0; auto. + intros; rewrite Locmap.gso; simpl; auto. + intros; rewrite Locmap.gso; simpl; auto. Qed. Corollary frame_undef_regs: @@ -533,7 +542,7 @@ Proof. induction res; simpl; intros. - apply frame_set_reg; auto. - auto. -- eauto. +- destruct ls; eauto. Qed. (** Invariance by change of memory injection. *) @@ -556,7 +565,7 @@ Qed. (** Agreement with Mach register states *) Definition agree_regs (j: meminj) (ls: locset) (rs: regset) : Prop := - forall r, Val.inject j (ls (R r)) (rs r). + forall r, Val.inject j (ls @ (R r)) (rs # r). (** Agreement over locations *) @@ -565,14 +574,14 @@ Record agree_locs (ls ls0: locset) : Prop := (** Unused registers have the same value as in the caller *) agree_unused_reg: - forall r, ~(mreg_within_bounds b r) -> ls (R r) = ls0 (R r); + forall r, ~(mreg_within_bounds b r) -> ls @ (R r) = ls0 @ (R r); (** Incoming stack slots have the same value as the corresponding Outgoing stack slots in the caller *) agree_incoming: forall ofs ty, In (S Incoming ofs ty) (regs_of_rpairs (loc_parameters f.(Linear.fn_sig))) -> - ls (S Incoming ofs ty) = ls0 (S Outgoing ofs ty) + ls @ (S Incoming ofs ty) = ls0 @ (S Outgoing ofs ty) }. (** ** Properties of [agree_regs]. *) @@ -581,7 +590,7 @@ Record agree_locs (ls ls0: locset) : Prop := Lemma agree_reg: forall j ls rs r, - agree_regs j ls rs -> Val.inject j (ls (R r)) (rs r). + agree_regs j ls rs -> Val.inject j (ls @ (R r)) (rs # r). Proof. intros. auto. Qed. @@ -591,7 +600,9 @@ Lemma agree_reglist: agree_regs j ls rs -> Val.inject_list j (reglist ls rl) (rs##rl). Proof. induction rl; simpl; intros. - auto. constructor; auto using agree_reg. + auto. + fold (ls @ (R a)). unfold regmap_read. + constructor; auto using agree_reg. Qed. Hint Resolve agree_reg agree_reglist: stacking. @@ -602,23 +613,26 @@ Lemma agree_regs_set_reg: forall j ls rs r v v', agree_regs j ls rs -> Val.inject j v v' -> + Val.has_type v' (mreg_type r) -> agree_regs j (Locmap.set (R r) v ls) (Regmap.set r v' rs). Proof. intros; red; intros. - unfold Regmap.set. destruct (RegEq.eq r0 r). subst r0. - rewrite Locmap.gss; auto. - rewrite Locmap.gso; auto. red. auto. + destruct (Mreg.eq_dec r0 r). subst r0. + rewrite Locmap.gss, Regmap.gss. + rewrite Mreg.chunk_of_reg_type. auto using Val.load_result_inject. + rewrite Locmap.gso, Regmap.gso; auto. red. auto. Qed. Lemma agree_regs_set_pair: forall j p v v' ls rs, agree_regs j ls rs -> Val.inject j v v' -> + Val.has_type_rpair v' p Val.loword Val.hiword mreg_type -> agree_regs j (Locmap.setpair p v ls) (set_pair p v' rs). Proof. - intros. destruct p; simpl. + intros. destruct p; try destruct H1; simpl. - apply agree_regs_set_reg; auto. -- apply agree_regs_set_reg. apply agree_regs_set_reg; auto. +- apply agree_regs_set_reg; auto. apply agree_regs_set_reg; auto. apply Val.hiword_inject; auto. apply Val.loword_inject; auto. Qed. @@ -626,12 +640,14 @@ Lemma agree_regs_set_res: forall j res v v' ls rs, agree_regs j ls rs -> Val.inject j v v' -> + Val.has_type_builtin_res v' res Val.loword Val.hiword mreg_type -> agree_regs j (Locmap.setres res v ls) (set_res res v' rs). Proof. - induction res; simpl; intros. + destruct res eqn:RES; simpl; intros. - apply agree_regs_set_reg; auto. - auto. -- apply IHres2. apply IHres1. auto. +- destruct H1. + repeat apply agree_regs_set_reg; try tauto. apply Val.hiword_inject; auto. apply Val.loword_inject; auto. Qed. @@ -639,7 +655,7 @@ Qed. Lemma agree_regs_exten: forall j ls rs ls' rs', agree_regs j ls rs -> - (forall r, ls' (R r) = Vundef \/ ls' (R r) = ls (R r) /\ rs' r = rs r) -> + (forall r, ls' @ (R r) = Vundef \/ ls' @ (R r) = ls @ (R r) /\ rs' # r = rs # r) -> agree_regs j ls' rs'. Proof. intros; red; intros. @@ -656,6 +672,7 @@ Proof. induction rl; simpl; intros. auto. apply agree_regs_set_reg; auto. + simpl; auto. Qed. Lemma agree_regs_undef_caller_save_regs: @@ -663,9 +680,12 @@ Lemma agree_regs_undef_caller_save_regs: agree_regs j ls rs -> agree_regs j (LTL.undef_caller_save_regs ls) (Mach.undef_caller_save_regs rs). Proof. - intros; red; intros. - unfold LTL.undef_caller_save_regs, Mach.undef_caller_save_regs. - destruct (is_callee_save r); auto. + intros; red; intros. + rewrite LTL.undef_caller_save_regs_correct, Mach.undef_caller_save_regs_correct. + unfold LTL.undef_caller_save_regs_spec, Mach.undef_caller_save_regs_spec. + destruct (is_callee_save r). + apply H. + auto. Qed. (** Preservation under assignment of stack slot *) @@ -695,7 +715,9 @@ Lemma agree_regs_call_regs: agree_regs j (call_regs ls) rs. Proof. intros. - unfold call_regs; intros; red; intros; auto. + unfold agree_regs, Locmap.get; intros. + fold ((call_regs ls) @ (R r)). rewrite call_regs_correct. + simpl; fold (ls @ (R r)); auto. Qed. (** ** Properties of [agree_locs] *) @@ -708,8 +730,8 @@ Lemma agree_locs_set_reg: mreg_within_bounds b r -> agree_locs (Locmap.set (R r) v ls) ls0. Proof. - intros. inv H; constructor; auto; intros. - rewrite Locmap.gso. auto. red. intuition congruence. + intros. inv H. + constructor; intros; rewrite Locmap.gso; auto; red; intuition congruence. Qed. Lemma caller_save_reg_within_bounds: @@ -742,7 +764,7 @@ Proof. induction res; simpl; intros. - eapply agree_locs_set_reg; eauto. - auto. -- apply IHres2; auto using in_or_app. +- repeat eapply agree_locs_set_reg; eauto. Qed. Lemma agree_locs_undef_regs: @@ -899,9 +921,7 @@ Variable sp: block. Variable ls: locset. Hypothesis ls_temp_undef: - forall ty r, In r (destroyed_by_setstack ty) -> ls (R r) = Vundef. - -Hypothesis wt_ls: forall r, Val.has_type (ls (R r)) (mreg_type r). + forall ty r, In r (destroyed_by_setstack ty) -> ls @ (R r) = Vundef. Lemma save_callee_save_rec_correct: forall k l pos rs m P, @@ -923,7 +943,7 @@ Local Opaque mreg_type. split. rewrite sep_pure; split; auto. eapply sep_drop; eauto. split. auto. auto. -- set (ty := mreg_type r) in *. +- set (ty := typ_of_quantity (quantity_of_typ (mreg_type r))) in *. set (sz := AST.typesize ty) in *. set (pos1 := align pos sz) in *. assert (SZPOS: sz > 0) by (apply AST.typesize_pos). @@ -933,68 +953,92 @@ Local Opaque mreg_type. { unfold pos1. apply Z.divide_trans with sz. unfold sz; rewrite <- size_type_chunk. apply align_size_chunk_divides. apply align_divides; auto. } - apply range_drop_left with (mid := pos1) in SEP; [ | omega ]. - apply range_split with (mid := pos1 + sz) in SEP; [ | omega ]. + assert (R_SZ: AST.typesize (mreg_type r) = sz). + { unfold sz, ty. rewrite (typ_of_quantity_of_typ (mreg_type r) (mreg_type_cases r)); auto. } + apply range_drop_left with (mid := pos1) in SEP; [ | rewrite R_SZ; fold pos1; omega ]. + apply range_split with (mid := pos1 + sz) in SEP; [ | rewrite R_SZ; fold pos1; omega ]. unfold sz at 1 in SEP. rewrite <- size_type_chunk in SEP. apply range_contains in SEP; auto. - exploit (contains_set_stack (fun v' => Val.inject j (ls (R r)) v') (rs r)). + exploit (contains_set_stack (fun v' => Val.inject j (ls @ (R r)) v') (rs # r)). + unfold ty in SEP. rewrite chunk_of_typ_of_quantity in SEP. eexact SEP. - apply load_result_inject; auto. apply wt_ls. + rewrite (chunk_of_quantity_of_typ (mreg_type r) (mreg_type_cases r)). + apply load_result_inject; auto. + apply well_typed_locset. clear SEP; intros (m1 & STORE & SEP). - set (rs1 := undef_regs (destroyed_by_setstack ty) rs). + set (rs1 := undef_regs (destroyed_by_setstack (quantity_of_typ (mreg_type r))) rs). assert (AG1: agree_regs j ls rs1). - { red; intros. unfold rs1. destruct (In_dec mreg_eq r0 (destroyed_by_setstack ty)). - erewrite ls_temp_undef by eauto. auto. - rewrite undef_regs_other by auto. apply AG. } + { red; intros. unfold rs1. + destruct (In_dec mreg_eq r0 (destroyed_by_setstack (quantity_of_typ (mreg_type r)))). + - erewrite ls_temp_undef; eauto. + - rewrite undef_regs_other by auto. apply AG. } rewrite sep_swap in SEP. exploit (IHl (pos1 + sz) rs1 m1); eauto. + rewrite R_SZ in SEP; eauto. intros (rs2 & m2 & A & B & C & D). exists rs2, m2. - split. eapply star_left; eauto. constructor. exact STORE. auto. traceEq. - split. rewrite sep_assoc, sep_swap. exact B. + split. eapply star_left; eauto. rewrite R_SZ; constructor. exact STORE. auto. traceEq. + split. rewrite sep_assoc, sep_swap, R_SZ. + rewrite (chunk_of_quantity_of_typ (mreg_type r) (mreg_type_cases r)) in B. exact B. split. intros. apply C. unfold store_stack in STORE; simpl in STORE. eapply Mem.perm_store_1; eauto. auto. Qed. End SAVE_CALLEE_SAVE. -Remark LTL_undef_regs_same: - forall r rl ls, In r rl -> LTL.undef_regs rl ls (R r) = Vundef. +Lemma undef_regs_same: + forall r rl rf, In r rl -> Regfile.get r (Regfile.undef_regs rl rf) = Vundef. Proof. induction rl; simpl; intros. contradiction. - unfold Locmap.set. destruct (Loc.eq (R a) (R r)). auto. - destruct (Loc.diff_dec (R a) (R r)); auto. - apply IHrl. intuition congruence. + destruct (mreg_eq a r). + - subst. rewrite Regfile.gss. destruct (Mreg.chunk_of r); auto. + - rewrite Regfile.gso; auto. apply IHrl; tauto. +Qed. + +Remark LTL_undef_regs_same: + forall r rl ls, In r rl -> (LTL.undef_regs rl ls) @ (R r) = Vundef. +Proof. + intros. simpl. + destruct ls. rewrite LTL_undef_regs_Regfile_undef_regs. apply undef_regs_same; auto. Qed. Remark LTL_undef_regs_others: - forall r rl ls, ~In r rl -> LTL.undef_regs rl ls (R r) = ls (R r). + forall r rl ls, ~In r rl -> (LTL.undef_regs rl ls) @ (R r) = ls @ (R r). Proof. - induction rl; simpl; intros. auto. - rewrite Locmap.gso. apply IHrl. intuition. red. intuition. + induction rl; intros. auto. + destruct ls as [rf stack]. + generalize (IHrl (rf, stack)); intro IH. + rewrite LTL_undef_regs_Regfile_undef_regs in *; simpl in *. + rewrite Regfile.gso. auto. intuition. Qed. Remark LTL_undef_regs_slot: - forall sl ofs ty rl ls, LTL.undef_regs rl ls (S sl ofs ty) = ls (S sl ofs ty). + forall sl ofs q rl ls, (LTL.undef_regs rl ls) @ (S sl ofs q) = ls @ (S sl ofs q). Proof. induction rl; simpl; intros. auto. - rewrite Locmap.gso. apply IHrl. red; auto. + destruct ls as [rf stack]. rewrite LTL_undef_regs_Regfile_undef_regs; auto. Qed. Remark undef_regs_type: forall ty l rl ls, - Val.has_type (ls l) ty -> Val.has_type (LTL.undef_regs rl ls l) ty. + Val.has_type (ls @ l) ty -> Val.has_type ((LTL.undef_regs rl ls) @ l) ty. Proof. induction rl; simpl; intros. - auto. -- unfold Locmap.set. destruct (Loc.eq (R a) l). red; auto. - destruct (Loc.diff_dec (R a) l); auto. red; auto. +- destruct ls as [rf stack]. rewrite LTL_undef_regs_Regfile_undef_regs. + destruct (Loc.eq (R a) l). + + subst l. simpl. rewrite Regfile.gss. + destruct (Mreg.chunk_of a); simpl; auto. + + generalize (IHrl (rf, stack) H); intros. + rewrite LTL_undef_regs_Regfile_undef_regs in H0. + unfold Locmap.get. destruct l; auto. + rewrite Regfile.gso. auto. congruence. Qed. Lemma save_callee_save_correct: forall j ls ls0 rs sp cs fb k m P, m |= range sp fe.(fe_ofs_callee_save) (size_callee_save_area b fe.(fe_ofs_callee_save)) ** P -> - (forall r, Val.has_type (ls (R r)) (mreg_type r)) -> + (forall r, Val.has_type (ls @ (R r)) (mreg_type r)) -> agree_callee_save ls ls0 -> agree_regs j ls rs -> let ls1 := LTL.undef_regs destroyed_at_function_entry (LTL.call_regs ls) in @@ -1010,7 +1054,6 @@ Proof. intros until P; intros SEP TY AGCS AG; intros ls1 rs1. exploit (save_callee_save_rec_correct j cs fb sp ls1). - intros. unfold ls1. apply LTL_undef_regs_same. eapply destroyed_by_setstack_function_entry; eauto. -- intros. unfold ls1. apply undef_regs_type. apply TY. - exact b.(used_callee_save_prop). - eexact SEP. - instantiate (1 := rs1). apply agree_regs_undef_regs. apply agree_regs_call_regs. auto. @@ -1019,7 +1062,7 @@ Proof. split. eexact EXEC. split. rewrite (contains_callee_saves_exten j sp ls0 ls1). exact SEP. intros. apply b.(used_callee_save_prop) in H. - unfold ls1. rewrite LTL_undef_regs_others. unfold call_regs. + unfold ls1. rewrite LTL_undef_regs_others, call_regs_correct. apply AGCS; auto. red; intros. assert (existsb is_callee_save destroyed_at_function_entry = false) @@ -1040,7 +1083,7 @@ Lemma function_prologue_correct: agree_regs j ls rs -> agree_callee_save ls ls0 -> agree_outgoing_arguments (Linear.fn_sig f) ls ls0 -> - (forall r, Val.has_type (ls (R r)) (mreg_type r)) -> + (forall r, Val.has_type (ls @ (R r)) (mreg_type r)) -> ls1 = LTL.undef_regs destroyed_at_function_entry (LTL.call_regs ls) -> rs1 = undef_regs destroyed_at_function_entry rs -> Mem.alloc m1 0 f.(Linear.fn_stacksize) = (m2, sp) -> @@ -1048,8 +1091,8 @@ Lemma function_prologue_correct: m1' |= minjection j m1 ** globalenv_inject ge j ** P -> exists j', exists rs', exists m2', exists sp', exists m3', exists m4', exists m5', Mem.alloc m1' 0 tf.(fn_stacksize) = (m2', sp') - /\ store_stack m2' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_link_ofs) parent = Some m3' - /\ store_stack m3' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_retaddr_ofs) ra = Some m4' + /\ store_stack m2' (Vptr sp' Ptrofs.zero) (quantity_of_typ Tptr) tf.(fn_link_ofs) parent = Some m3' + /\ store_stack m3' (Vptr sp' Ptrofs.zero) (quantity_of_typ Tptr) tf.(fn_retaddr_ofs) ra = Some m4' /\ star step tge (State cs fb (Vptr sp' Ptrofs.zero) (save_callee_save fe k) rs1 m4') E0 (State cs fb (Vptr sp' Ptrofs.zero) k rs' m5') @@ -1087,16 +1130,26 @@ Local Opaque b fe. apply (frame_env_separated b) in SEP. replace (make_env b) with fe in SEP by auto. (* Store of parent *) rewrite sep_swap3 in SEP. - apply (range_contains Mptr) in SEP; [|tauto]. - exploit (contains_set_stack (fun v' => v' = parent) parent (fun _ => True) m2' Tptr). - rewrite chunk_of_Tptr; eexact SEP. apply Val.load_result_same; auto. + apply (range_contains Mptr_any) in SEP; [|tauto]. + exploit (contains_set_stack (fun v' => v' = parent) parent (fun _ => True) m2' (quantity_of_typ Tptr)). + rewrite chunk_of_quantity_of_Tptr. + eexact SEP. + unfold Tptr in *. destruct Archi.ptr64 eqn:PTR64; auto. + change (chunk_of_quantity _) with (chunk_of_type Tany32). + apply Val.load_result_same. + apply Val.has_subtype with (ty1 := Tint); auto. clear SEP; intros (m3' & STORE_PARENT & SEP). rewrite sep_swap3 in SEP. (* Store of return address *) rewrite sep_swap4 in SEP. - apply (range_contains Mptr) in SEP; [|tauto]. - exploit (contains_set_stack (fun v' => v' = ra) ra (fun _ => True) m3' Tptr). - rewrite chunk_of_Tptr; eexact SEP. apply Val.load_result_same; auto. + apply (range_contains Mptr_any) in SEP; [|try apply align_chunk_Mptr_any; tauto]. + exploit (contains_set_stack (fun v' => v' = ra) ra (fun _ => True) m3' (quantity_of_typ Tptr)). + rewrite chunk_of_quantity_of_Tptr. + eexact SEP. + unfold Tptr in *. destruct Archi.ptr64 eqn:PTR64; auto. + change (chunk_of_quantity _) with (chunk_of_type Tany32). + apply Val.load_result_same. + apply Val.has_subtype with (ty1 := Tint); auto. clear SEP; intros (m4' & STORE_RETADDR & SEP). rewrite sep_swap4 in SEP. (* Saving callee-save registers *) @@ -1109,20 +1162,20 @@ Local Opaque b fe. rewrite sep_swap5 in SEP. (* Materializing the Local and Outgoing locations *) exploit (initial_locations j'). eexact SEP. tauto. - instantiate (1 := Local). instantiate (1 := ls1). - intros; rewrite LS1. rewrite LTL_undef_regs_slot. reflexivity. + instantiate (1 := ls1). instantiate (1 := Local). + intros; rewrite LS1. rewrite LTL_undef_regs_slot, call_regs_correct. reflexivity. clear SEP; intros SEP. rewrite sep_swap in SEP. exploit (initial_locations j'). eexact SEP. tauto. - instantiate (1 := Outgoing). instantiate (1 := ls1). - intros; rewrite LS1. rewrite LTL_undef_regs_slot. reflexivity. + instantiate (1 := ls1). instantiate (1 := Outgoing). + intros; rewrite LS1. rewrite LTL_undef_regs_slot, call_regs_correct. reflexivity. clear SEP; intros SEP. rewrite sep_swap in SEP. (* Now we frame this *) assert (SEPFINAL: m5' |= frame_contents j' sp' ls1 ls0 parent ra ** minjection j' m2 ** globalenv_inject ge j' ** P). { eapply frame_mconj. eexact SEPCONJ. - rewrite chunk_of_Tptr in SEP. - unfold frame_contents_1; rewrite ! sep_assoc. exact SEP. + unfold frame_contents_1; rewrite ! sep_assoc. + rewrite !chunk_of_quantity_of_Tptr in SEP. exact SEP. assert (forall ofs k p, Mem.perm m2' sp' ofs k p -> Mem.perm m5' sp' ofs k p). { intros. apply PERMS. unfold store_stack in STORE_PARENT, STORE_RETADDR. @@ -1141,9 +1194,10 @@ Local Opaque b fe. split. eexact SAVE_CS. split. exact AGREGS'. split. rewrite LS1. apply agree_locs_undef_locs; [|reflexivity]. - constructor; intros. unfold call_regs. apply AGCS. + constructor; intros. rewrite call_regs_correct. apply AGCS. unfold mreg_within_bounds in H; tauto. - unfold call_regs. apply AGARGS. apply incoming_slot_in_parameters; auto. + rewrite call_regs_correct. unfold call_regs_spec, Locmap.read. apply AGARGS. + apply incoming_slot_in_parameters; auto. split. exact SEPFINAL. split. exact SAME. exact INCR. Qed. @@ -1163,7 +1217,7 @@ Variable ls0: locset. Variable m: mem. Definition agree_unused (ls0: locset) (rs: regset) : Prop := - forall r, ~(mreg_within_bounds b r) -> Val.inject j (ls0 (R r)) (rs r). + forall r, ~(mreg_within_bounds b r) -> Val.inject j (ls0 @ (R r)) (rs # r). Lemma restore_callee_save_rec_correct: forall l ofs rs k, @@ -1174,8 +1228,8 @@ Lemma restore_callee_save_rec_correct: star step tge (State cs fb (Vptr sp Ptrofs.zero) (restore_callee_save_rec l ofs k) rs m) E0 (State cs fb (Vptr sp Ptrofs.zero) k rs' m) - /\ (forall r, In r l -> Val.inject j (ls0 (R r)) (rs' r)) - /\ (forall r, ~(In r l) -> rs' r = rs r) + /\ (forall r, In r l -> Val.inject j (ls0 @ (R r)) (rs' # r)) + /\ (forall r, ~(In r l) -> rs' # r = rs # r) /\ agree_unused ls0 rs'. Proof. Local Opaque mreg_type. @@ -1183,27 +1237,36 @@ Local Opaque mreg_type. - (* base case *) exists rs. intuition auto. apply star_refl. - (* inductive case *) - set (ty := mreg_type r) in *. + set (ty := typ_of_quantity (quantity_of_typ (mreg_type r))) in *. set (sz := AST.typesize ty) in *. set (ofs1 := align ofs sz). assert (SZPOS: sz > 0) by (apply AST.typesize_pos). assert (OFSLE: ofs <= ofs1) by (apply align_le; auto). assert (BOUND: mreg_within_bounds b r) by eauto. + assert (R_SZ: AST.typesize (mreg_type r) = sz). + { unfold sz, ty. rewrite (typ_of_quantity_of_typ (mreg_type r) (mreg_type_cases r)); auto. } exploit contains_get_stack. + rewrite <- (chunk_of_quantity_of_typ (mreg_type r) (mreg_type_cases r)) in H. eapply sep_proj1; eassumption. intros (v & LOAD & SPEC). exploit (IHl (ofs1 + sz) (rs#r <- v)). + rewrite <- (typ_of_quantity_of_typ (mreg_type r) (mreg_type_cases r)) in H. + fold ty in H. eapply sep_proj2; eassumption. red; intros. rewrite Regmap.gso. auto. intuition congruence. eauto. intros (rs' & A & B & C & D). exists rs'. split. eapply star_step; eauto. - econstructor. exact LOAD. traceEq. + rewrite R_SZ in *. econstructor. exact LOAD. traceEq. split. intros. - destruct (In_dec mreg_eq r0 l). auto. + fold (ls0 @ (R r0)). destruct (In_dec mreg_eq r0 l). auto. assert (r = r0) by tauto. subst r0. - rewrite C by auto. rewrite Regmap.gss. exact SPEC. + rewrite C by auto. rewrite Regmap.gss, Mreg.chunk_of_reg_type. unfold Mreg.type. + rewrite Val.load_result_same. + exact SPEC. + apply Mem.loadv_type in LOAD. + destruct (mreg_type_cases r) as [T | T]; rewrite T in *; auto. split. intros. rewrite C by tauto. apply Regmap.gso. intuition auto. exact D. @@ -1220,9 +1283,9 @@ Lemma restore_callee_save_correct: (State cs fb (Vptr sp Ptrofs.zero) (restore_callee_save fe k) rs m) E0 (State cs fb (Vptr sp Ptrofs.zero) k rs' m) /\ (forall r, - is_callee_save r = true -> Val.inject j (ls0 (R r)) (rs' r)) + is_callee_save r = true -> Val.inject j (ls0 @ (R r)) (rs' # r)) /\ (forall r, - is_callee_save r = false -> rs' r = rs r). + is_callee_save r = false -> rs' # r = rs # r). Proof. intros. unfold frame_contents, frame_contents_1 in H. @@ -1252,8 +1315,8 @@ Lemma function_epilogue_correct: j sp = Some(sp', fe.(fe_stack_data)) -> Mem.free m sp 0 f.(Linear.fn_stacksize) = Some m1 -> exists rs1, exists m1', - load_stack m' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_link_ofs) = Some pa - /\ load_stack m' (Vptr sp' Ptrofs.zero) Tptr tf.(fn_retaddr_ofs) = Some ra + load_stack m' (Vptr sp' Ptrofs.zero) (quantity_of_typ Tptr) tf.(fn_link_ofs) = Some pa + /\ load_stack m' (Vptr sp' Ptrofs.zero) (quantity_of_typ Tptr) tf.(fn_retaddr_ofs) = Some ra /\ Mem.free m' sp' 0 tf.(fn_stacksize) = Some m1' /\ star step tge (State cs fb (Vptr sp' Ptrofs.zero) (restore_callee_save fe k) rs m') @@ -1279,8 +1342,10 @@ Proof. (* Reloading the back link and return address *) unfold frame_contents in SEP; apply mconj_proj1 in SEP. unfold frame_contents_1 in SEP; rewrite ! sep_assoc in SEP. - exploit (hasvalue_get_stack Tptr). rewrite chunk_of_Tptr. eapply sep_pick3; eexact SEP. intros LOAD_LINK. - exploit (hasvalue_get_stack Tptr). rewrite chunk_of_Tptr. eapply sep_pick4; eexact SEP. intros LOAD_RETADDR. + exploit (hasvalue_get_stack (quantity_of_typ Tptr)). rewrite chunk_of_quantity_of_Tptr. + eapply sep_pick3; eexact SEP. intros LOAD_LINK. + exploit (hasvalue_get_stack (quantity_of_typ Tptr)). rewrite chunk_of_quantity_of_Tptr. + eapply sep_pick4; eexact SEP. intros LOAD_RETADDR. clear SEP. (* Conclusions *) rewrite unfold_transf_function; simpl. @@ -1289,12 +1354,12 @@ Proof. split. assumption. split. assumption. split. eassumption. - split. red; unfold return_regs; intros. + split. red; intros. rewrite return_regs_correct. unfold return_regs_spec. destruct (is_callee_save r) eqn:C. apply CS; auto. rewrite NCS by auto. apply AGR. - split. red; unfold return_regs; intros. - destruct l. rewrite H; auto. destruct sl; auto; contradiction. + split. red; intros. rewrite return_regs_correct. unfold return_regs_spec. + destruct l; auto. rewrite H; auto. destruct sl; auto; contradiction. assumption. Qed. @@ -1330,9 +1395,9 @@ Inductive match_stacks (j: meminj): (INJ: j sp = Some(sp', (fe_stack_data (make_env (function_bounds f))))) (TY_RA: Val.has_type ra Tptr) (AGL: agree_locs f ls (parent_locset cs)) - (ARGS: forall ofs ty, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments sg)) -> - slot_within_bounds (function_bounds f) Outgoing ofs ty) + (ARGS: forall ofs q, + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments sg)) -> + slot_within_bounds (function_bounds f) Outgoing ofs q) (STK: match_stacks j cs cs' (Linear.fn_sig f)), match_stacks j (Linear.Stackframe f (Vptr sp Ptrofs.zero) ls c :: cs) @@ -1561,7 +1626,7 @@ Lemma find_function_translated: /\ transf_fundef f = OK tf. Proof. intros until f; intros AG [bound [_ [?????]]] FF. - destruct ros; simpl in FF. + destruct ros; unfold find_function in FF. - exploit Genv.find_funct_inv; eauto. intros [b EQ]. rewrite EQ in FF. rewrite Genv.find_funct_find_funct_ptr in FF. exploit function_ptr_translated; eauto. intros [tf [A B]]. @@ -1595,20 +1660,20 @@ Hypothesis SEP: m' |= stack_contents j cs cs'. Lemma transl_external_argument: forall l, In l (regs_of_rpairs (loc_arguments sg)) -> - exists v, extcall_arg rs m' (parent_sp cs') l v /\ Val.inject j (ls l) v. + exists v, extcall_arg rs m' (parent_sp cs') l v /\ Val.inject j (ls @ l) v. Proof. intros. assert (loc_argument_acceptable l) by (apply loc_arguments_acceptable_2 with sg; auto). destruct l; red in H0. -- exists (rs r); split. constructor. auto. +- exists (rs # r); split. constructor. auto. - destruct sl; try contradiction. inv MS. + elim (H1 _ H). + simpl in SEP. unfold parent_sp. - assert (slot_valid f Outgoing pos ty = true). + assert (slot_valid f Outgoing pos q = true). { destruct H0. unfold slot_valid, proj_sumbool. - rewrite zle_true by omega. rewrite pred_dec_true by auto. reflexivity. } - assert (slot_within_bounds (function_bounds f) Outgoing pos ty) by eauto. + rewrite zle_true by omega. destruct q; rewrite pred_dec_true by auto; reflexivity. } + assert (slot_within_bounds (function_bounds f) Outgoing pos q) by eauto. exploit frame_get_outgoing; eauto. intros (v & A & B). exists v; split. constructor. exact A. rewrite AGARGS by auto. exact B. @@ -1676,11 +1741,11 @@ Hypothesis SEP: m' |= frame_contents f j sp' ls ls0 parent retaddr ** minjection Lemma transl_builtin_arg_correct: forall a v, - eval_builtin_arg ge ls (Vptr sp Ptrofs.zero) m a v -> + eval_builtin_arg ge (Locmap.read ls) (Vptr sp Ptrofs.zero) m a v -> (forall l, In l (params_of_builtin_arg a) -> loc_valid f l = true) -> - (forall sl ofs ty, In (S sl ofs ty) (params_of_builtin_arg a) -> slot_within_bounds b sl ofs ty) -> + (forall sl ofs q, In (S sl ofs q) (params_of_builtin_arg a) -> slot_within_bounds b sl ofs q) -> exists v', - eval_builtin_arg ge rs (Vptr sp' Ptrofs.zero) m' (transl_builtin_arg fe a) v' + eval_builtin_arg ge (regmap_read rs) (Vptr sp' Ptrofs.zero) m' (transl_builtin_arg fe a) v' /\ Val.inject j v v'. Proof. assert (SYMB: forall id ofs, Val.inject j (Senv.symbol_address ge id ofs) (Senv.symbol_address ge id ofs)). @@ -1692,10 +1757,10 @@ Proof. Local Opaque fe. induction 1; simpl; intros VALID BOUNDS. - assert (loc_valid f x = true) by auto. - destruct x as [r | [] ofs ty]; try discriminate. - + exists (rs r); auto with barg. + destruct x as [r | [] ofs q]; try discriminate. + + exists (rs # r); unfold Locmap.read. split. constructor. auto. + exploit frame_get_local; eauto. intros (v & A & B). - exists v; split; auto. constructor; auto. + exists v; split; auto. constructor; destruct q; auto. - econstructor; eauto with barg. - econstructor; eauto with barg. - econstructor; eauto with barg. @@ -1722,11 +1787,11 @@ Qed. Lemma transl_builtin_args_correct: forall al vl, - eval_builtin_args ge ls (Vptr sp Ptrofs.zero) m al vl -> + eval_builtin_args ge (Locmap.read ls) (Vptr sp Ptrofs.zero) m al vl -> (forall l, In l (params_of_builtin_args al) -> loc_valid f l = true) -> - (forall sl ofs ty, In (S sl ofs ty) (params_of_builtin_args al) -> slot_within_bounds b sl ofs ty) -> + (forall sl ofs q, In (S sl ofs q) (params_of_builtin_args al) -> slot_within_bounds b sl ofs q) -> exists vl', - eval_builtin_args ge rs (Vptr sp' Ptrofs.zero) m' (List.map (transl_builtin_arg fe) al) vl' + eval_builtin_args ge (regmap_read rs) (Vptr sp' Ptrofs.zero) m' (List.map (transl_builtin_arg fe) al) vl' /\ Val.inject_list j vl vl'. Proof. induction 1; simpl; intros VALID BOUNDS. @@ -1821,8 +1886,11 @@ Proof. exploit frame_get_local; eauto. intros (v & A & B). econstructor; split. apply plus_one. apply exec_Mgetstack. exact A. - econstructor; eauto with coqlib. + econstructor; destruct rs; eauto with coqlib. apply agree_regs_set_reg; auto. + inversion WTS; subst. simpl in WTC; InvBooleans. + apply Val.has_subtype with (ty1 := typ_of_quantity q); auto. + apply Mem.load_type in A. destruct q; auto. apply agree_locs_set_reg; auto. + (* Lgetstack, incoming *) unfold slot_valid in SV. InvBooleans. @@ -1839,16 +1907,22 @@ Proof. apply plus_one. eapply exec_Mgetparam; eauto. rewrite (unfold_transf_function _ _ TRANSL). unfold fn_link_ofs. eapply frame_get_parent. eexact SEP. - econstructor; eauto with coqlib. econstructor; eauto. - apply agree_regs_set_reg. apply agree_regs_set_reg. auto. auto. + econstructor; destruct rs; eauto with coqlib. econstructor; eauto. + apply agree_regs_set_reg. apply agree_regs_set_reg. auto. auto. simpl; auto. erewrite agree_incoming by eauto. exact B. + inversion WTS; subst. simpl in WTC; InvBooleans. + apply Val.has_subtype with (ty1 := typ_of_quantity q); auto. + apply Mem.load_type in A. destruct q; auto. apply agree_locs_set_reg; auto. apply agree_locs_undef_locs; auto. + (* Lgetstack, outgoing *) exploit frame_get_outgoing; eauto. intros (v & A & B). econstructor; split. apply plus_one. apply exec_Mgetstack. exact A. - econstructor; eauto with coqlib. + econstructor; destruct rs; eauto with coqlib. apply agree_regs_set_reg; auto. + inversion WTS; subst. simpl in WTC; InvBooleans. + apply Val.has_subtype with (ty1 := typ_of_quantity q); auto. + apply Mem.load_type in A. destruct q; auto. apply agree_locs_set_reg; auto. - (* Lsetstack *) @@ -1858,11 +1932,11 @@ Proof. | Incoming => 0 (* dummy *) | Outgoing => offset_arg ofs end). - eapply frame_undef_regs with (rl := destroyed_by_setstack ty) in SEP. + eapply frame_undef_regs with (rl := destroyed_by_setstack q) in SEP. assert (A: exists m'', - store_stack m' (Vptr sp' Ptrofs.zero) ty (Ptrofs.repr ofs') (rs0 src) = Some m'' - /\ m'' |= frame_contents f j sp' (Locmap.set (S sl ofs ty) (rs (R src)) - (LTL.undef_regs (destroyed_by_setstack ty) rs)) + store_stack m' (Vptr sp' Ptrofs.zero) q (Ptrofs.repr ofs') (rs0 # src) = Some m'' + /\ m'' |= frame_contents f j sp' (Locmap.set (S sl ofs q) (rs @ (R src)) + (LTL.undef_regs (destroyed_by_setstack q) rs)) (parent_locset s) (parent_sp cs') (parent_ra cs') ** stack_contents j s cs' ** minjection j m ** globalenv_inject ge j). { unfold ofs'; destruct sl; try discriminate. @@ -1894,6 +1968,24 @@ Proof. econstructor; eauto with coqlib. apply agree_regs_set_reg; auto. rewrite transl_destroyed_by_op. apply agree_regs_undef_regs; auto. + inversion WTS; subst. simpl in WTC; InvBooleans. + destruct (is_move_operation op args) eqn:MOVE. + (* is move *) + apply is_move_operation_correct in MOVE; destruct MOVE. subst. + simpl in A; inversion A. + apply Val.has_subtype with (ty1 := mreg_type m0); auto. + unfold regmap_read, regmap_get. apply Regmap.get_has_type. + (* is not move *) + generalize (is_not_move_operation ge _ _ args m H MOVE); intros. + assert (subtype (snd (type_of_operation op)) (mreg_type res) = true). + { rewrite <- H0. destruct (type_of_operation op); reflexivity. } + assert (Val.has_type v (snd (type_of_operation op))) by eauto using type_of_operation_sound. + apply Val.has_subtype with (ty1 := snd (type_of_operation op)); auto. + apply type_of_operation_sound in A. + assert (TOP: forall env, type_of_operation (transl_op env op) = type_of_operation op). + { intros. destruct op; simpl; auto; try (destruct a; simpl; auto). } + rewrite TOP in A. auto. + destruct op; simpl; auto. congruence. try congruence. apply agree_locs_set_reg; auto. apply agree_locs_undef_locs. auto. apply destroyed_by_op_caller_save. apply frame_set_reg. apply frame_undef_regs. exact SEP. @@ -1913,8 +2005,11 @@ Proof. apply plus_one. econstructor. instantiate (1 := a'). rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved. eexact C. eauto. - econstructor; eauto with coqlib. + econstructor; destruct rs; eauto with coqlib. apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto. + inversion WTS; subst. simpl in WTC; InvBooleans. + destruct a'; try inversion C. apply Mem.load_type in H4. + apply Val.has_subtype with (ty1 := type_of_chunk chunk); auto. apply agree_locs_set_reg. apply agree_locs_undef_locs. auto. apply destroyed_by_load_caller_save. auto. - (* Lstore *) @@ -1985,9 +2080,19 @@ Proof. apply plus_one. econstructor; eauto. eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved. eapply external_call_symbols_preserved; eauto. apply senv_preserved. + inversion WTS; subst. simpl in WTC; InvBooleans. destruct res; auto. eapply match_states_intro with (j := j'); eauto with coqlib. eapply match_stacks_change_meminj; eauto. - apply agree_regs_set_res; auto. apply agree_regs_undef_regs; auto. eapply agree_regs_inject_incr; eauto. + eapply agree_regs_set_res; eauto. apply agree_regs_undef_regs; auto. eapply agree_regs_inject_incr; eauto. + inversion WTS; subst. simpl in WTC; InvBooleans. + unfold wt_builtin_res in H3. + unfold Val.has_type_builtin_res. + apply external_call_well_typed in EC. + destruct res; auto. + eapply Val.has_subtype; eauto. + InvBooleans; split. + eapply Val.has_subtype; eauto. destruct res'; auto; constructor. + eapply Val.has_subtype; eauto. destruct res'; auto; constructor. apply agree_locs_set_res; auto. apply agree_locs_undef_regs; auto. apply frame_set_res. apply frame_undef_regs. apply frame_contents_incr with j; auto. rewrite sep_swap2. apply stack_contents_change_meminj with j; auto. rewrite sep_swap2. @@ -2028,7 +2133,7 @@ Proof. apply frame_undef_regs; auto. - (* Ljumptable *) - assert (rs0 arg = Vint n). + assert (rs0 # arg = Vint n). { generalize (AGREGS arg). rewrite H. intro IJ; inv IJ; auto. } econstructor; split. apply plus_one; eapply exec_Mjumptable; eauto. @@ -2083,8 +2188,9 @@ Proof. eapply match_states_return with (j := j'). eapply match_stacks_change_meminj; eauto. apply agree_regs_set_pair. apply agree_regs_undef_caller_save_regs. - apply agree_regs_inject_incr with j; auto. - auto. + apply agree_regs_inject_incr with j; auto. auto. + apply external_call_well_typed in A. + apply loc_result_has_type; auto. apply stack_contents_change_meminj with j; auto. rewrite sep_comm, sep_assoc; auto. @@ -2114,7 +2220,7 @@ Proof. set (j := Mem.flat_inj (Mem.nextblock m0)). eapply match_states_call with (j := j); eauto. constructor. red; intros. rewrite H3, loc_arguments_main in H. contradiction. - red; simpl; auto. + red; intros. unfold Locmap.get, Regfile.get; simpl. rewrite FragBlock.gi; auto. simpl. rewrite sep_pure. split; auto. split;[|split]. eapply Genv.initmem_inject; eauto. simpl. exists (Mem.nextblock m0); split. apply Ple_refl. @@ -2138,7 +2244,7 @@ Proof. - generalize (loc_result_type signature_main). rewrite LR. discriminate. } destruct R as [rres EQ]. rewrite EQ in H1. simpl in H1. - generalize (AGREGS rres). rewrite H1. intros A; inv A. + generalize (AGREGS rres). fold (rs @ (R rres)) in H1. rewrite H1. intros A; inv A. econstructor; eauto. Qed. diff --git a/backend/Tunnelingproof.v b/backend/Tunnelingproof.v index 4f95ac9b82..3e73084980 100644 --- a/backend/Tunnelingproof.v +++ b/backend/Tunnelingproof.v @@ -201,7 +201,7 @@ Definition tunneled_code (f: function) := PTree.map1 (tunneled_block f) (fn_code f). Definition locmap_lessdef (ls1 ls2: locset) : Prop := - forall l, Val.lessdef (ls1 l) (ls2 l). + forall l, Val.lessdef (ls1 @ l) (ls2 @ l). Inductive match_stackframes: stackframe -> stackframe -> Prop := | match_stackframes_intro: @@ -254,39 +254,88 @@ Lemma reglist_lessdef: forall rl ls1 ls2, locmap_lessdef ls1 ls2 -> Val.lessdef_list (reglist ls1 rl) (reglist ls2 rl). Proof. - induction rl; simpl; intros; auto. + induction rl; simpl; intros. auto. + destruct ls1, ls2. + apply Val.lessdef_list_cons. exact (H (R a)). auto. +Qed. + +Lemma regfile_set_lessdef: + forall rf1 stack1 rf2 stack2 v1 v2 r, + locmap_lessdef (rf1, stack1) (rf2, stack2) -> + Val.lessdef v1 v2 -> + locmap_lessdef (Regfile.set r v1 rf1, stack1) (Regfile.set r v2 rf2, stack2). +Proof. + intros; red; destruct l; unfold Locmap.get. +- destruct (mreg_eq r r0). + + subst. rewrite !Regfile.gss. auto using Val.load_result_lessdef. + + rewrite !Regfile.gso; auto. exact (H (R r0)). +- exact (H (S sl pos q)). +Qed. + +Lemma locmap_set_reg_lessdef: + forall ls1 ls2 v1 v2 r, + locmap_lessdef ls1 ls2 -> Val.lessdef v1 v2 -> locmap_lessdef (Locmap.set (R r) v1 ls1) (Locmap.set (R r) v2 ls2). +Proof. + intros; red; intros l. unfold Locmap.set. destruct l, ls1, ls2. +- apply regfile_set_lessdef; auto. +- exact (H (S sl pos q)). Qed. Lemma locmap_set_lessdef: forall ls1 ls2 v1 v2 l, locmap_lessdef ls1 ls2 -> Val.lessdef v1 v2 -> locmap_lessdef (Locmap.set l v1 ls1) (Locmap.set l v2 ls2). Proof. - intros; red; intros l'. unfold Locmap.set. destruct (Loc.eq l l'). -- destruct l; auto using Val.load_result_lessdef. -- destruct (Loc.diff_dec l l'); auto. + intros; red; intros l'. + destruct (Loc.eq l l'). + rewrite e, !Locmap.gss. auto using Val.load_result_lessdef. + destruct (Loc.diff_dec l l'). + rewrite !Locmap.gso; auto. + rewrite Locmap.gu_overlap; auto using Loc.diff_sym. +Qed. + +Lemma regfile_set_undef_lessdef: + forall ls1 ls2 r, + locmap_lessdef ls1 ls2 -> locmap_lessdef (Locmap.set (R r) Vundef ls1) ls2. +Proof. + intros; red; intros l. destruct l as [r'|], ls1, ls2; simpl. +- destruct (mreg_eq r r'). + + subst; rewrite Regfile.gss; auto. + destruct (Mreg.chunk_of r'); simpl; auto. + + rewrite Regfile.gso; auto. exact (H (R r')). +- exact (H (S sl pos q)). Qed. Lemma locmap_set_undef_lessdef: forall ls1 ls2 l, locmap_lessdef ls1 ls2 -> locmap_lessdef (Locmap.set l Vundef ls1) ls2. Proof. - intros; red; intros l'. unfold Locmap.set. destruct (Loc.eq l l'). -- destruct l; auto. destruct ty; auto. -- destruct (Loc.diff_dec l l'); auto. + intros; red; intros l'. + destruct (Loc.eq l l'). + rewrite e, Locmap.gss. destruct (chunk_of_type (Loc.type l')); auto. + destruct (Loc.diff_dec l l'). + rewrite Locmap.gso; auto. + rewrite Locmap.gu_overlap; auto using Loc.diff_sym. Qed. Lemma locmap_undef_regs_lessdef: forall rl ls1 ls2, locmap_lessdef ls1 ls2 -> locmap_lessdef (undef_regs rl ls1) (undef_regs rl ls2). Proof. - induction rl as [ | r rl]; intros; simpl. auto. apply locmap_set_lessdef; auto. + intros. destruct ls1, ls2. + rewrite !LTL_undef_regs_Regfile_undef_regs. + induction rl as [ | r rl]; simpl. auto. + apply regfile_set_lessdef; auto. Qed. Lemma locmap_undef_regs_lessdef_1: forall rl ls1 ls2, locmap_lessdef ls1 ls2 -> locmap_lessdef (undef_regs rl ls1) ls2. Proof. - induction rl as [ | r rl]; intros; simpl. auto. apply locmap_set_undef_lessdef; auto. + intros. destruct ls1, ls2. + rewrite !LTL_undef_regs_Regfile_undef_regs. + induction rl as [ | r rl]; simpl. auto. + fold (Locmap.set (R r) Vundef (Regfile.undef_regs rl t, t0)). + apply regfile_set_undef_lessdef; auto. Qed. (* @@ -324,21 +373,26 @@ Lemma locmap_setpair_lessdef: forall p ls1 ls2 v1 v2, locmap_lessdef ls1 ls2 -> Val.lessdef v1 v2 -> locmap_lessdef (Locmap.setpair p v1 ls1) (Locmap.setpair p v2 ls2). Proof. - intros; destruct p; simpl; auto using locmap_set_lessdef, Val.loword_lessdef, Val.hiword_lessdef. + intros; destruct p; unfold Locmap.setpair. + auto using locmap_set_lessdef. + auto using locmap_set_lessdef, Val.loword_lessdef, Val.hiword_lessdef. Qed. Lemma locmap_setres_lessdef: forall res ls1 ls2 v1 v2, locmap_lessdef ls1 ls2 -> Val.lessdef v1 v2 -> locmap_lessdef (Locmap.setres res v1 ls1) (Locmap.setres res v2 ls2). Proof. - induction res; intros; simpl; auto using locmap_set_lessdef, Val.loword_lessdef, Val.hiword_lessdef. + induction res; intros; unfold Locmap.setres; auto using locmap_set_lessdef, Val.loword_lessdef, Val.hiword_lessdef. Qed. Lemma locmap_undef_caller_save_regs_lessdef: forall ls1 ls2, locmap_lessdef ls1 ls2 -> locmap_lessdef (undef_caller_save_regs ls1) (undef_caller_save_regs ls2). Proof. - intros; red; intros. unfold undef_caller_save_regs. + intros; red; intros. rewrite !undef_caller_save_regs_correct. + unfold Locmap.get, undef_caller_save_regs_spec. + unfold locmap_lessdef, Locmap.get in H. + generalize (H l); intro L. destruct l. - destruct (Conventions1.is_callee_save r); auto. - destruct sl; auto. @@ -351,10 +405,10 @@ Lemma find_function_translated: find_function tge ros tls = Some (tunnel_fundef fd). Proof. intros. destruct ros; simpl in *. -- assert (E: tls (R m) = ls (R m)). - { exploit Genv.find_funct_inv; eauto. intros (b & EQ). - generalize (H (R m)). rewrite EQ. intros LD; inv LD. auto. } - rewrite E. apply functions_translated; auto. +- assert (E: tls @ (R m) = ls @ (R m)). + { exploit Genv.find_funct_inv; eauto. intros (b & EQ). destruct ls; simpl. + generalize (H (R m)). rewrite EQ. intros LD; inv LD. auto. destruct tls; congruence. } + fold (tls @ (R m)). rewrite E. apply functions_translated; auto. - rewrite symbols_preserved. destruct (Genv.find_symbol ge i); inv H0. apply function_ptr_translated; auto. Qed. @@ -362,7 +416,10 @@ Qed. Lemma call_regs_lessdef: forall ls1 ls2, locmap_lessdef ls1 ls2 -> locmap_lessdef (call_regs ls1) (call_regs ls2). Proof. - intros; red; intros. destruct l as [r | [] ofs ty]; simpl; auto. + intros; red; intros. rewrite !call_regs_correct. + destruct l as [r | [] ofs ty], ls1, ls2; simpl; auto. + exact (H (R r)). + exact (H (S Outgoing ofs ty)). Qed. Lemma return_regs_lessdef: @@ -371,9 +428,14 @@ Lemma return_regs_lessdef: locmap_lessdef callee1 callee2 -> locmap_lessdef (return_regs caller1 callee1) (return_regs caller2 callee2). Proof. - intros; red; intros. destruct l; simpl. -- destruct (Conventions1.is_callee_save r); auto. -- destruct sl; auto. + intros; red; intros. rewrite !return_regs_correct. + generalize (H l); intro Callers. + generalize (H0 l); intro Callees. + destruct l; simpl. + destruct caller1, callee1, caller2, callee2. + destruct (Conventions1.is_callee_save r); auto. + destruct caller1, caller2. + destruct sl; auto. Qed. (** To preserve non-terminating behaviours, we show that the transformed @@ -427,7 +489,7 @@ Proof. - (* Lop *) exploit eval_operation_lessdef. apply reglist_lessdef; eauto. eauto. eauto. intros (tv & EV & LD). - left; simpl; econstructor; split. + left; econstructor; split. eapply exec_Lop with (v := tv); eauto. rewrite <- EV. apply eval_operation_preserved. exact symbols_preserved. econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef. @@ -436,17 +498,17 @@ Proof. intros (ta & EV & LD). exploit Mem.loadv_extends. eauto. eauto. eexact LD. intros (tv & LOAD & LD'). - left; simpl; econstructor; split. + left; econstructor; split. eapply exec_Lload with (a := ta). rewrite <- EV. apply eval_addressing_preserved. exact symbols_preserved. eauto. eauto. econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef. - (* Lgetstack *) - left; simpl; econstructor; split. + left; econstructor; split. econstructor; eauto. econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef. - (* Lsetstack *) - left; simpl; econstructor; split. + left; econstructor; split. econstructor; eauto. econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef. - (* Lstore *) @@ -454,7 +516,7 @@ Proof. intros (ta & EV & LD). exploit Mem.storev_extends. eauto. eauto. eexact LD. apply LS. intros (tm' & STORE & MEM'). - left; simpl; econstructor; split. + left; econstructor; split. eapply exec_Lstore with (a := ta). rewrite <- EV. apply eval_addressing_preserved. exact symbols_preserved. eauto. eauto. @@ -475,7 +537,8 @@ Proof. apply sig_preserved. econstructor; eauto using return_regs_lessdef, match_parent_locset. - (* Lbuiltin *) - exploit eval_builtin_args_lessdef. eexact LS. eauto. eauto. intros (tvargs & EVA & LDA). + assert (LS': forall l: loc, Val.lessdef (Locmap.read rs l) (Locmap.read tls l)) by auto. + exploit eval_builtin_args_lessdef. eexact LS'. eauto. eauto. intros (tvargs & EVA & LDA). exploit external_call_mem_extends; eauto. intros (tvres & tm' & A & B & C & D). left; simpl; econstructor; split. eapply exec_Lbuiltin; eauto. @@ -503,7 +566,7 @@ Proof. destruct b; econstructor; eauto using locmap_undef_regs_lessdef. - (* Ljumptable *) - assert (tls (R arg) = Vint n). + assert (tls @ (R arg) = Vint n). { generalize (LS (R arg)); rewrite H; intros LD; inv LD; auto. } left; simpl; econstructor; split. eapply exec_Ljumptable. @@ -539,7 +602,7 @@ Lemma transf_initial_states: exists st2, initial_state tprog st2 /\ match_states st1 st2. Proof. intros. inversion H. - exists (Callstate nil (tunnel_fundef f) (Locmap.init Vundef) m0); split. + exists (Callstate nil (tunnel_fundef f) Locmap.init m0); split. econstructor; eauto. apply (Genv.init_mem_transf TRANSL); auto. rewrite (match_program_main TRANSL). diff --git a/backend/XTL.ml b/backend/XTL.ml index f10efeedbc..d39b38b8fd 100644 --- a/backend/XTL.ml +++ b/backend/XTL.ml @@ -142,10 +142,10 @@ let rec type_builtin_args al tyl = | a :: al, ty :: tyl -> type_builtin_arg a ty; type_builtin_args al tyl | _, _ -> raise Type_error -let rec type_builtin_res a ty = +let type_builtin_res a ty = match a with | BR v -> set_var_type v ty - | BR_splitlong(a1, a2) -> type_builtin_res a1 Tint; type_builtin_res a2 Tint + | BR_splitlong(a1, a2) -> set_var_type a1 Tint; set_var_type a2 Tint | _ -> () let type_instr = function diff --git a/common/AST.v b/common/AST.v index 145f49190a..7d1ff01da7 100644 --- a/common/AST.v +++ b/common/AST.v @@ -158,6 +158,8 @@ Global Opaque chunk_eq. Definition Mptr : memory_chunk := if Archi.ptr64 then Mint64 else Mint32. +Definition Mptr_any : memory_chunk := if Archi.ptr64 then Many64 else Many32. + (** The type (integer/pointer or float) of a chunk. *) Definition type_of_chunk (c: memory_chunk) : typ := @@ -193,6 +195,9 @@ Definition chunk_of_type (ty: typ) := Lemma chunk_of_Tptr: chunk_of_type Tptr = Mptr. Proof. unfold Mptr, Tptr; destruct Archi.ptr64; auto. Qed. +Lemma type_of_chunk_of_type: forall ty, type_of_chunk (chunk_of_type ty) = ty. +Proof. destruct ty; simpl; auto. Qed. + (** Initialization data for global variables. *) Inductive init_data: Type := @@ -633,7 +638,7 @@ Inductive builtin_arg (A: Type) : Type := Inductive builtin_res (A: Type) : Type := | BR (x: A) | BR_none - | BR_splitlong (hi lo: builtin_res A). + | BR_splitlong (hi lo: A). Fixpoint globals_of_builtin_arg (A: Type) (a: builtin_arg A) : list ident := match a with @@ -658,11 +663,11 @@ Fixpoint params_of_builtin_arg (A: Type) (a: builtin_arg A) : list A := Definition params_of_builtin_args (A: Type) (al: list (builtin_arg A)) : list A := List.fold_right (fun a l => params_of_builtin_arg a ++ l) nil al. -Fixpoint params_of_builtin_res (A: Type) (a: builtin_res A) : list A := +Definition params_of_builtin_res (A: Type) (a: builtin_res A) : list A := match a with | BR x => x :: nil | BR_none => nil - | BR_splitlong hi lo => params_of_builtin_res hi ++ params_of_builtin_res lo + | BR_splitlong hi lo => hi :: lo :: nil end. Fixpoint map_builtin_arg (A B: Type) (f: A -> B) (a: builtin_arg A) : builtin_arg B := @@ -682,12 +687,11 @@ Fixpoint map_builtin_arg (A B: Type) (f: A -> B) (a: builtin_arg A) : builtin_ar BA_addptr (map_builtin_arg f a1) (map_builtin_arg f a2) end. -Fixpoint map_builtin_res (A B: Type) (f: A -> B) (a: builtin_res A) : builtin_res B := +Definition map_builtin_res (A B: Type) (f: A -> B) (a: builtin_res A) : builtin_res B := match a with | BR x => BR (f x) | BR_none => BR_none - | BR_splitlong hi lo => - BR_splitlong (map_builtin_res f hi) (map_builtin_res f lo) + | BR_splitlong hi lo => BR_splitlong (f hi) (f lo) end. (** Which kinds of builtin arguments are supported by which external function. *) diff --git a/common/Memdata.v b/common/Memdata.v index a9ed48b445..e87e34166f 100644 --- a/common/Memdata.v +++ b/common/Memdata.v @@ -73,6 +73,11 @@ Proof. unfold Mptr; destruct Archi.ptr64; auto. Qed. +Lemma size_chunk_Mptr_any: size_chunk Mptr_any = if Archi.ptr64 then 8 else 4. +Proof. + unfold Mptr_any; destruct Archi.ptr64; auto. +Qed. + (** Memory reads and writes must respect alignment constraints: the byte offset of the location being addressed should be an exact multiple of the natural alignment for the chunk being addressed. @@ -108,6 +113,11 @@ Proof. unfold Mptr; destruct Archi.ptr64; auto. Qed. +Lemma align_chunk_Mptr_any: align_chunk Mptr_any = 4. +Proof. + unfold Mptr_any; destruct Archi.ptr64; auto. +Qed. + Lemma align_size_chunk_divides: forall chunk, (align_chunk chunk | size_chunk chunk). Proof. @@ -132,15 +142,93 @@ Definition quantity_eq (q1 q2: quantity) : {q1 = q2} + {q1 <> q2}. Proof. decide equality. Defined. Global Opaque quantity_eq. +Definition size_quantity (q: quantity) : Z := + match q with Q32 => 4 | Q64 => 8 end. + Definition size_quantity_nat (q: quantity) := match q with Q32 => 4%nat | Q64 => 8%nat end. +Lemma size_quantity_pos: + forall q, size_quantity q > 0. +Proof. + destruct q; simpl; omega. +Qed. + Lemma size_quantity_nat_pos: forall q, exists n, size_quantity_nat q = S n. Proof. intros. destruct q; [exists 3%nat | exists 7%nat]; auto. Qed. +(** The "most general" type that can accommodate every value that fits in the + given quantity. *) +Definition typ_of_quantity (q: quantity) : typ := + match q with + | Q32 => Tany32 + | Q64 => Tany64 + end. + +Definition quantity_of_typ (ty: typ) := + match ty with + | Tint | Tsingle | Tany32 => Q32 + | Tlong | Tfloat | Tany64 => Q64 + end. + +Lemma typ_of_quantity_of_typ: + forall ty, + ty = Tany32 \/ ty = Tany64 -> + typ_of_quantity (quantity_of_typ ty) = ty. +Proof. + intros; destruct H; subst; auto. +Qed. + +Lemma subtype_of_typ_of_quantity: + forall ty, + subtype ty (typ_of_quantity (quantity_of_typ ty)) = true. +Proof. + destruct ty; simpl; auto. +Qed. + +Lemma has_typ_of_quantity: + forall v ty, + Val.has_type v ty -> + Val.has_type v (typ_of_quantity (quantity_of_typ ty)). +Proof. + eauto using Val.has_subtype, subtype_of_typ_of_quantity. +Qed. + +Definition chunk_of_quantity (q: quantity) : memory_chunk := + match q with + | Q32 => Many32 + | Q64 => Many64 + end. + +Definition quantity_chunk (chunk: memory_chunk) := + match chunk with + | Mint64 | Mfloat64 | Many64 => Q64 + | _ => Q32 + end. + +Lemma chunk_of_typ_of_quantity: + forall q, chunk_of_type (typ_of_quantity q) = chunk_of_quantity q. +Proof. + destruct q; auto. +Qed. + +Lemma chunk_of_quantity_of_typ: + forall ty, + ty = Tany32 \/ ty = Tany64 -> + chunk_of_quantity (quantity_of_typ ty) = chunk_of_type ty. +Proof. + intros; destruct H; subst; auto. +Qed. + +Lemma chunk_of_quantity_of_Tptr: + chunk_of_quantity (quantity_of_typ Tptr) = Mptr_any. +Proof. + unfold Tptr, Mptr_any. destruct Archi.ptr64; auto. +Qed. + (** * Memory values *) (** A ``memory value'' is a byte-sized quantity that describes the current @@ -340,6 +428,98 @@ Proof. simpl. decEq. auto. Qed. +Lemma proj_bytes_fragment: + forall vl v q n, + In (Fragment v q n) vl -> + proj_bytes vl = None. +Proof. + intros; induction vl. + - tauto. + - destruct H. + + subst a. reflexivity. + + destruct a; simpl; auto. + rewrite IHvl; auto. +Qed. + +Lemma proj_bytes_undef: + forall vl, + In Undef vl -> + proj_bytes vl = None. +Proof. + intros; induction vl. + - tauto. + - destruct H. + + subst a. reflexivity. + + destruct a; simpl; auto. + rewrite IHvl; auto. +Qed. + +Definition fits_quantity (v: val) (q: quantity): Prop := + match v, q with + | Vundef, _ => True + | Vint _, _ => True + | Vlong _, Q64 => True + | Vfloat _, Q64 => True + | Vsingle _, _ => True + | Vptr _ _, Q32 => Archi.ptr64 = false + | _, Q64 => True + | _, _ => False + end. + +Remark fits_quantity_Q32: forall v, fits_quantity v Q32 <-> Val.has_type v Tany32. +Proof. + destruct v; simpl; tauto. +Qed. + +Remark fits_quantity_Q64: forall v, fits_quantity v Q64 <-> Val.has_type v Tany64. +Proof. + destruct v; simpl; tauto. +Qed. + +Remark fits_quantity_has_type: + forall v q, + fits_quantity v q <-> Val.has_type v (type_of_chunk (chunk_of_quantity q)). +Proof. + destruct q; simpl. apply fits_quantity_Q32. apply fits_quantity_Q64. +Qed. + +Corollary has_type_fits_quantity: + forall v ty, + Val.has_type v ty -> fits_quantity v (quantity_of_typ ty). +Proof. + intros. + destruct ty; simpl; (apply fits_quantity_Q32 || apply fits_quantity_Q64); + eapply Val.has_subtype; eauto; simpl; auto. +Qed. + +Remark fits_quantity_load_result: + forall v q, + fits_quantity v q -> Val.load_result (chunk_of_quantity q) v = v. +Proof. + intros. apply fits_quantity_has_type in H. + destruct q; simpl in *; auto. + destruct v; auto; inv H; rewrite H1; auto. +Qed. + +Remark any_fits_quantity_Q64: forall v, fits_quantity v Q64. +Proof. + destruct v; simpl; tauto. +Qed. + +Remark fits_quantity_undef: forall q, fits_quantity Vundef q. +Proof. + destruct q; simpl; auto. +Qed. + +Program Definition fits_quantity_dec (v: val) (q: quantity): { fits_quantity v q } + { ~ fits_quantity v q } := + match q with + | Q32 => if Val.has_type_dec v Tany32 then left _ else right _ + | Q64 => left _ + end. +Next Obligation. + destruct v; simpl; auto. +Qed. + Fixpoint inj_value_rec (n: nat) (v: val) (q: quantity) {struct n}: list memval := match n with | O => nil @@ -347,24 +527,277 @@ Fixpoint inj_value_rec (n: nat) (v: val) (q: quantity) {struct n}: list memval : end. Definition inj_value (q: quantity) (v: val): list memval := - inj_value_rec (size_quantity_nat q) v q. - -Fixpoint check_value (n: nat) (v: val) (q: quantity) (vl: list memval) - {struct n} : bool := + if Val.eq v Vundef then + list_repeat (size_quantity_nat q) Undef + else if fits_quantity_dec v q then + inj_value_rec (size_quantity_nat q) v q + else + list_repeat (size_quantity_nat q) Undef. + +Fixpoint check_value_rec (n: nat) (v: val) (q: quantity) (vl: list memval) + {struct n} : bool := match n, vl with | O, nil => true | S m, Fragment v' q' m' :: vl' => - Val.eq v v' && quantity_eq q q' && Nat.eqb m m' && check_value m v q vl' + Val.eq v v' && quantity_eq q q' && Nat.eqb m m' && check_value_rec m v q vl' | _, _ => false end. +Definition check_value (n: nat) (v: val) (q: quantity) (vl: list memval): bool := + fits_quantity_dec v q && check_value_rec n v q vl. + Definition proj_value (q: quantity) (vl: list memval) : val := match vl with | Fragment v q' n :: vl' => - if check_value (size_quantity_nat q) v q vl then v else Vundef + if check_value (size_quantity_nat q) v q vl + then v + else Vundef | _ => Vundef end. +Lemma check_value_rec_undef: + forall n q v vl, + In Undef vl -> check_value_rec n v q vl = false. +Proof. + induction n; intros; simpl. + destruct vl. elim H. auto. + destruct vl. auto. + destruct m; auto. simpl in H; destruct H. congruence. + rewrite IHn; auto. apply andb_false_r. +Qed. + +Corollary check_value_undef: + forall n q v vl, + In Undef vl -> check_value n v q vl = false. +Proof. + intros. unfold check_value. rewrite check_value_rec_undef, andb_false_r; auto. +Qed. + +Lemma proj_value_undef: + forall q vl, In Undef vl -> proj_value q vl = Vundef. +Proof. + intros; unfold proj_value. + destruct vl; auto. destruct m; auto. + rewrite check_value_undef. auto. auto. +Qed. + +Lemma inj_value_rec_length: + forall n q v, + length (inj_value_rec n v q) = n. +Proof. + induction n; simpl; auto. +Qed. + +Lemma inj_value_length: + forall q v, + length (inj_value q v) = size_quantity_nat q. +Proof. + intros. unfold inj_value. + destruct (Val.eq v Vundef), (fits_quantity_dec v q); auto using inj_value_rec_length, length_list_repeat. +Qed. + +Lemma firstn_inj_value: + forall q v, + firstn (size_quantity_nat q) (inj_value q v) = inj_value q v. +Proof. + intros. erewrite <- inj_value_length, firstn_all. reflexivity. +Qed. + +Lemma check_value_rec_length: + forall q v vl n, + check_value_rec n v q vl = true -> n = length vl. +Proof. + induction vl; intros. +- destruct n; auto. simpl in H; congruence. +- destruct n; simpl in H. congruence. + destruct a; try congruence. InvBooleans; subst. + simpl; auto. +Qed. + +Corollary check_value_length: + forall q v vl n, + check_value n v q vl = true -> n = length vl. +Proof. + unfold check_value. intros. InvBooleans. + eauto using check_value_rec_length. +Qed. + +Lemma check_value_length_mismatch: + forall q v vl, + size_quantity_nat q <> length vl -> check_value (size_quantity_nat q) v q vl = false. +Proof. + intros. apply not_true_is_false. contradict H. + apply check_value_length in H; auto. +Qed. + +Lemma proj_value_length_mismatch: + forall q vl, + size_quantity_nat q <> length vl -> proj_value q vl = Vundef. +Proof. + intros. unfold proj_value. + destruct vl; auto. destruct m; auto. + rewrite check_value_length_mismatch; auto. +Qed. + +Corollary proj_value_length: + forall q vl, + proj_value q vl <> Vundef -> size_quantity_nat q = length vl. +Proof. + intros. + destruct (Nat.eq_dec (size_quantity_nat q) (length vl)); auto. + contradict H. auto using proj_value_length_mismatch. +Qed. + +Lemma proj_value_check_value: + forall q vl v, + proj_value q vl = v -> + v <> Vundef -> + check_value (length vl) v q vl = true. +Proof. + intros. + destruct (check_value (size_quantity_nat q) v q vl) eqn:C. + - erewrite proj_value_length in C; eauto. congruence. + - unfold proj_value in H. + destruct vl; try congruence. + destruct m; try congruence. + destruct (check_value (size_quantity_nat q) v0 q (Fragment v0 q0 n :: vl)) eqn:E; congruence. +Qed. + +Lemma check_value_rec_fragment_length: + forall v q vl f, + check_value_rec (length vl) v q vl = true -> + In f vl -> + exists n', f = Fragment v q n'. +Proof. + induction vl; intros; inv H0. + destruct f; inv H; InvBooleans. exists n; congruence. + destruct a; inv H; InvBooleans. auto. +Qed. + +Corollary check_value_fragment_length: + forall v q vl f, + check_value_rec (length vl) v q vl = true -> + In f vl -> + exists n', f = Fragment v q n'. +Proof. + eauto using check_value_rec_fragment_length. +Qed. + +Lemma check_value_fragment: + forall n v q vl f, + check_value n v q vl = true -> + In f vl -> + exists n', f = Fragment v q n'. +Proof. + intros. unfold check_value in H; InvBooleans. + exploit check_value_rec_length; eauto; intros. + subst n; eauto using check_value_rec_fragment_length. +Qed. + +Lemma check_value_diff_q: + forall q q' v n n' vl, + q <> q' -> + In (Fragment v q n) vl -> check_value n' v q' vl = false. +Proof. + intros. destruct (check_value n' v q' vl) eqn:C; auto; exfalso. + eapply check_value_fragment in C; eauto. + destruct C; congruence. +Qed. + +Lemma proj_value_diff_q: + forall q q' v n vl, + q <> q' -> + In (Fragment v q n) vl -> proj_value q' vl = Vundef. +Proof. + intros; unfold proj_value. + destruct vl; auto. + inversion H0. subst. erewrite check_value_diff_q; eauto. + destruct m; eauto. + destruct (check_value (size_quantity_nat q') v0 q' (Fragment v0 q0 n0 :: vl)) eqn:C; eauto. + eapply check_value_fragment in C; eauto. destruct C; congruence. +Qed. + +Lemma check_value_byte: + forall b n v q vl, + In (Byte b) vl -> + check_value n v q vl = false. +Proof. + intros. destruct (check_value n v q vl) eqn:C; auto; exfalso. + eapply check_value_fragment in C; eauto. + destruct C; congruence. +Qed. + +Lemma proj_value_byte: + forall b q vl, + In (Byte b) vl -> + proj_value q vl = Vundef. +Proof. + intros. unfold proj_value. + destruct vl; auto. destruct m; auto. + erewrite check_value_byte; eauto. +Qed. + +Lemma check_value_fits_quantity: + forall n v q vl, + check_value n v q vl = true -> fits_quantity v q. +Proof. + intros. unfold check_value in H. InvBooleans; auto. +Qed. + +Lemma proj_value_fits_quantity: + forall q mvs, fits_quantity (proj_value q mvs) q. +Proof. + intros; unfold proj_value. + destruct mvs; simpl; auto. + destruct m; simpl; auto. + destruct (check_value _ _ _ _) eqn:CV; simpl; auto. + eauto using check_value_fits_quantity. +Qed. + +Lemma inj_bytes_encode_int: + forall sz x, + (0 <> sz)%nat -> + exists b, In (Byte b) (inj_bytes (encode_int sz x)). +Proof. + intros. unfold inj_bytes. + generalize (encode_int_length sz x); intro LEN. + destruct (encode_int sz x) eqn:E. + simpl in LEN; congruence. + exists i. simpl; auto. +Qed. + +Lemma proj_value_encode_int: + forall sz x q, + (0 <> sz)%nat -> + proj_value q (inj_bytes (encode_int sz x)) = Vundef. +Proof. + intros. + apply inj_bytes_encode_int with (x := x) in H. destruct H. + eauto using proj_value_byte. +Qed. + +Lemma proj_value_encode_int_app_l: + forall sz x q l, + (0 <> sz)%nat -> + proj_value q (inj_bytes (encode_int sz x) ++ l) = Vundef. +Proof. + intros. + exploit inj_bytes_encode_int; eauto; intros [b B]. + assert (In (Byte b) (inj_bytes (encode_int sz x) ++ l)) by (apply in_app; eauto). + eauto using proj_value_byte. +Qed. + +Lemma proj_value_encode_int_app_r: + forall sz x q l, + (0 <> sz)%nat -> + proj_value q (l ++ inj_bytes (encode_int sz x)) = Vundef. +Proof. + intros. + exploit inj_bytes_encode_int; eauto; intros [b B]. + assert (In (Byte b) (l ++ inj_bytes (encode_int sz x))) by (apply in_app; eauto). + eauto using proj_value_byte. +Qed. + Definition encode_val (chunk: memory_chunk) (v: val) : list memval := match v, chunk with | Vint n, (Mint8signed | Mint8unsigned) => inj_bytes (encode_int 1%nat (Int.unsigned n)) @@ -410,6 +843,7 @@ Ltac solve_encode_val_length := | [ |- length (inj_bytes _) = _ ] => rewrite length_inj_bytes; solve_encode_val_length | [ |- length (encode_int _ _) = _ ] => apply encode_int_length | [ |- length (if ?x then _ else _) = _ ] => destruct x eqn:?; solve_encode_val_length + | [ |- length (inj_value _ _) = _ ] => rewrite inj_value_length; solve_encode_val_length | _ => reflexivity end. @@ -419,38 +853,195 @@ Proof. intros. destruct v; simpl; destruct chunk; solve_encode_val_length. Qed. +Lemma encode_val_quantity_size: + forall q v, length (encode_val (chunk_of_quantity q) v) = size_quantity_nat q. +Proof. + intros. rewrite encode_val_length. destruct q; auto. +Qed. + +Lemma encode_val_q: + forall q v, + encode_val (chunk_of_quantity q) v = inj_value q v. +Proof. + destruct q, v; auto. +Qed. + +Lemma encode_val_any32: + forall v, + encode_val Many32 v = inj_value Q32 v. +Proof. + apply encode_val_q with (q := Q32). +Qed. + +Lemma encode_val_any64: + forall v, + encode_val Many64 v = inj_value Q64 v. +Proof. + apply encode_val_q with (q := Q64). +Qed. + Lemma check_inj_value: - forall v q n, check_value n v q (inj_value_rec n v q) = true. + forall v q n, fits_quantity v q -> check_value n v q (inj_value_rec n v q) = true. Proof. - induction n; simpl. auto. - unfold proj_sumbool. rewrite dec_eq_true. rewrite dec_eq_true. + unfold check_value. + induction n; simpl; intros; unfold proj_sumbool. + rewrite pred_dec_true; auto. + rewrite dec_eq_true. rewrite dec_eq_true. rewrite <- beq_nat_refl. simpl; auto. Qed. Lemma proj_inj_value: - forall q v, proj_value q (inj_value q v) = v. + forall q v, fits_quantity v q -> proj_value q (inj_value q v) = v. Proof. - intros. unfold proj_value, inj_value. destruct (size_quantity_nat_pos q) as [n EQ]. - rewrite EQ at 1. simpl. rewrite check_inj_value. auto. + intros. unfold proj_value, inj_value. + destruct (Val.eq v Vundef). destruct q; auto. + rewrite pred_dec_true by auto. + destruct (size_quantity_nat_pos q) as [n' EQ]. + rewrite EQ at 1. simpl. + rewrite check_inj_value; auto. +Qed. + +Remark proj_inj_undef: + forall q v, ~ fits_quantity v q -> proj_value q (inj_value q v) = Vundef. +Proof. + intros. unfold proj_value, inj_value. + rewrite pred_dec_false by (contradict H; subst; simpl; auto). + rewrite pred_dec_false by auto. + destruct q; simpl; auto. Qed. Remark in_inj_value: - forall mv v q, In mv (inj_value q v) -> exists n, mv = Fragment v q n. + forall mv v q, + v <> Vundef -> + fits_quantity v q -> + In mv (inj_value q v) -> + exists n, mv = Fragment v q n. Proof. Local Transparent inj_value. - unfold inj_value; intros until q. generalize (size_quantity_nat q). induction n; simpl; intros. + unfold inj_value; intros until q. intros D F. + rewrite pred_dec_false, pred_dec_true by auto. + generalize (size_quantity_nat q). induction n; simpl; intros. contradiction. destruct H. exists n; auto. eauto. Qed. +Remark inj_value_hd: + forall q v, + exists tl, + inj_value q v = Undef :: tl \/ inj_value q v = Fragment v q (pred (size_quantity_nat q)) :: tl. +Proof. + unfold inj_value; intros. + destruct (Val.eq v Vundef). destruct q; eexists; left; simpl; eauto. + destruct (fits_quantity_dec v q). destruct q; simpl; eauto. + destruct q; eexists; left; simpl; eauto. +Qed. + +Ltac destruct_nth_error_index := + match goal with + | [ |- nth_error (_ :: _) ?n = _ ] => + destruct n; simpl; auto + end. + +Remark inj_value_nth_error: + forall q v n, + (n < size_quantity_nat q)%nat -> + nth_error (inj_value q v) n = Some Undef \/ + nth_error (inj_value q v) n = Some (Fragment v q (pred (size_quantity_nat q) - n)%nat). +Proof. + unfold inj_value; intros. + destruct (Val.eq v Vundef). + - destruct q; simpl in H; left; simpl. + repeat destruct_nth_error_index; omega. + repeat destruct_nth_error_index; omega. + - destruct (fits_quantity_dec v q). + + destruct q; simpl in H; right; simpl. + repeat destruct_nth_error_index; omega. + repeat destruct_nth_error_index; omega. + + destruct q; simpl in H; left; simpl. + repeat destruct_nth_error_index; omega. + repeat destruct_nth_error_index; omega. +Qed. + +Lemma check_value_rec_nth_error: + forall vl n v q, + check_value_rec (length vl) v q vl = true -> + (n < length vl)%nat -> + nth_error vl n = Some (Fragment v q (pred (length vl) - n)%nat). +Proof. + induction vl; intros. + - simpl in H0; omega. + - simpl in H. destruct a; try congruence. InvBooleans. subst. + destruct n. + + rewrite Nat.sub_0_r; simpl. repeat f_equal. apply eq_sym, Nat.eqb_eq; auto. + + simpl. replace (length vl - S n)%nat with (pred (length vl) - n)%nat by omega. + apply IHvl. auto. simpl in H0. omega. +Qed. + +Lemma check_value_nth_error: + forall n v q vl, + check_value (size_quantity_nat q) v q vl = true -> + (n < size_quantity_nat q)%nat -> + nth_error vl n = Some (Fragment v q (pred (size_quantity_nat q) - n)%nat). +Proof. + intros until vl. intros CHK LEN. + exploit check_value_length; eauto; intro H. rewrite H in *. + unfold check_value in CHK. InvBooleans. + apply check_value_rec_nth_error; auto. +Qed. + +Lemma proj_value_misaligned: + forall vl n v q n', + nth_error vl n = Some (Fragment v q n') -> + n' <> (pred (size_quantity_nat q) - n)%nat -> + proj_value q vl = Vundef. +Proof. + intros. unfold proj_value. + destruct vl; auto. + destruct m; auto. + destruct (check_value _ _ _ _) eqn:CHK; auto. + exploit check_value_length; eauto; intro S. + exploit check_value_nth_error; eauto. + rewrite S in *. eapply nth_error_Some. erewrite H. congruence. + congruence. +Qed. + +Remark in_inj_value_cases: + forall q v, In Undef (inj_value q v) \/ In (Fragment v q (pred (size_quantity_nat q))) (inj_value q v). +Proof. + intros. destruct (inj_value_hd q v) as [tl [H | H]]; rewrite H. + left; constructor; auto. + right; constructor; auto. +Qed. + +Remark inj_value_charact: + forall q v v' q' n vl, + inj_value q v = Fragment v' q' n :: vl -> + q = q' /\ v = v' /\ S n = size_quantity_nat q /\ fits_quantity v q. +Proof. + intros. unfold inj_value in H. + destruct (Val.eq v Vundef). destruct q; simpl in H; congruence. + destruct (fits_quantity_dec v q); destruct q; inv H; auto. +Qed. + +Lemma proj_inj_value_undef: + forall q1 q2, proj_value q1 (inj_value q2 Vundef) = Vundef. +Proof. + destruct q1, q2; unfold inj_value; rewrite pred_dec_true; simpl; auto. +Qed. + Lemma proj_inj_value_mismatch: forall q1 q2 v, q1 <> q2 -> proj_value q1 (inj_value q2 v) = Vundef. Proof. - intros. unfold proj_value. destruct (inj_value q2 v) eqn:V. auto. destruct m; auto. - destruct (in_inj_value (Fragment v0 q n) v q2) as [n' EQ]. + intros. + destruct (Val.eq v Vundef) as [E | NE]. subst; auto using proj_inj_value_undef. + unfold proj_value. destruct (inj_value q2 v) eqn:V. auto. destruct m; auto. + destruct (inj_value_charact q2 v v0 q n l V) as (Q & V' & N & F). + destruct (in_inj_value (Fragment v0 q n) v q2) as [n' EQ]; auto. rewrite V; auto with coqlib. inv EQ. + unfold check_value. destruct (size_quantity_nat_pos q1) as [p EQ1]; rewrite EQ1; simpl. - unfold proj_sumbool. rewrite dec_eq_true. rewrite dec_eq_false by congruence. auto. + unfold proj_sumbool. rewrite dec_eq_true. rewrite dec_eq_false by congruence. + rewrite andb_false_r. auto. Qed. Definition decode_encode_val (v1: val) (chunk1 chunk2: memory_chunk) (v2: val) : Prop := @@ -493,16 +1084,25 @@ Definition decode_encode_val (v1: val) (chunk1 chunk2: memory_chunk) (v2: val) : | Vsingle f, _, _ => True (* nothing interesting to say about v2 *) end. -Remark decode_val_undef: +Remark decode_val_hd_undef: forall bl chunk, decode_val chunk (Undef :: bl) = Vundef. Proof. intros. unfold decode_val. simpl. destruct chunk, Archi.ptr64; auto. Qed. +Remark decode_val_undef: + forall bl chunk, In Undef bl -> decode_val chunk bl = Vundef. +Proof. + intros. unfold decode_val. + rewrite proj_bytes_undef by auto. + destruct chunk; auto; rewrite proj_value_undef; auto; destruct Archi.ptr64; auto. +Qed. + Remark proj_bytes_inj_value: forall q v, proj_bytes (inj_value q v) = None. Proof. - intros. destruct q; reflexivity. + intros. unfold inj_value. + destruct (Val.eq v Vundef), (fits_quantity_dec v q); destruct q; simpl; reflexivity. Qed. Ltac solve_decode_encode_val_general := @@ -511,7 +1111,7 @@ Ltac solve_decode_encode_val_general := | |- context [ if Archi.ptr64 then _ else _ ] => destruct Archi.ptr64 eqn:? | |- context [ proj_bytes (inj_bytes _) ] => rewrite proj_inj_bytes | |- context [ proj_bytes (inj_value _ _) ] => rewrite proj_bytes_inj_value - | |- context [ proj_value _ (inj_value _ _) ] => rewrite ?proj_inj_value, ?proj_inj_value_mismatch by congruence + | |- context [ proj_value _ (inj_value _ _) ] => rewrite ?proj_inj_value, ?proj_inj_value_mismatch, ?proj_inj_undef by (try congruence; simpl; auto) | |- context [ Int.repr(decode_int (encode_int 1 (Int.unsigned _))) ] => rewrite decode_encode_int_1 | |- context [ Int.repr(decode_int (encode_int 2 (Int.unsigned _))) ] => rewrite decode_encode_int_2 | |- context [ Int.repr(decode_int (encode_int 4 (Int.unsigned _))) ] => rewrite decode_encode_int_4 @@ -527,7 +1127,7 @@ Lemma decode_encode_val_general: Proof. Opaque inj_value. intros. - destruct v; destruct chunk1 eqn:C1; try (apply decode_val_undef); + destruct v; destruct chunk1 eqn:C1; try (apply decode_val_hd_undef); destruct chunk2 eqn:C2; unfold decode_encode_val, decode_val, encode_val, Val.load_result; repeat solve_decode_encode_val_general. - rewrite Float.of_to_bits; auto. @@ -546,6 +1146,23 @@ Proof. destruct v1; auto. Qed. +Lemma decode_encode_val_q: + forall q v, + decode_val (chunk_of_quantity q) (encode_val (chunk_of_quantity q) v) = + Val.load_result (chunk_of_quantity q) v. +Proof. + intros. + exploit decode_encode_val_similar; eauto. + apply decode_encode_val_general. +Qed. + +Lemma decode_encode_undef: + forall chunk1 chunk2, + decode_val chunk1 (encode_val chunk2 Vundef) = Vundef. +Proof. + intros. destruct chunk1, chunk2; simpl; auto; unfold decode_val; destruct Archi.ptr64; auto. +Qed. + Lemma decode_val_type: forall chunk cl, Val.has_type (decode_val chunk cl) (type_of_chunk chunk). @@ -558,6 +1175,15 @@ Local Opaque Val.load_result. (exact I || apply Val.load_result_type || destruct Archi.ptr64; (exact I || apply Val.load_result_type)). Qed. +Lemma decode_val_diff_q: + forall chunk q v n vl, + q <> quantity_chunk chunk -> + In (Fragment v q n) vl -> decode_val chunk vl = Vundef. +Proof. + intros. unfold decode_val. erewrite proj_bytes_fragment; eauto. + destruct chunk eqn:CHUNK, q eqn:Q, Archi.ptr64; auto; simpl in H; erewrite proj_value_diff_q; eauto. +Qed. + Lemma encode_val_int8_signed_unsigned: forall v, encode_val Mint8signed v = encode_val Mint8unsigned v. Proof. @@ -595,6 +1221,17 @@ Proof. intros; unfold encode_val. decEq. apply encode_int_16_mod. apply Int.eqmod_sign_ext'. compute; auto. Qed. +Lemma encode_val_inj_value: + forall v q chunk, + fits_quantity v q -> + chunk = chunk_of_type (typ_of_quantity q) -> + encode_val chunk v = inj_value q v. +Proof. + intros; destruct q; subst chunk; simpl. + destruct v; auto. + destruct v; auto. +Qed. + Lemma decode_val_cast: forall chunk l, let v := decode_val chunk l in @@ -615,12 +1252,6 @@ Qed. (** Pointers cannot be forged. *) -Definition quantity_chunk (chunk: memory_chunk) := - match chunk with - | Mint64 | Mfloat64 | Many64 => Q64 - | _ => Q32 - end. - Inductive shape_encoding (chunk: memory_chunk) (v: val): list memval -> Prop := | shape_encoding_f: forall q i mvl, (chunk = Mint32 \/ chunk = Many32 \/ chunk = Mint64 \/ chunk = Many64) -> @@ -655,8 +1286,11 @@ Proof. { Local Transparent inj_value. intros. unfold inj_value. destruct (size_quantity_nat_pos q) as [sz' EQ']. - rewrite EQ'. simpl. constructor; auto. - intros; eapply A; eauto. omega. + rewrite EQ'. + destruct (Val.eq v Vundef). simpl. constructor. eauto using in_list_repeat. + destruct (fits_quantity_dec v q). + - simpl. constructor; auto. intros; eapply A; eauto. omega. + - simpl. constructor. intros. eauto using in_list_repeat. } assert (C: forall bl, match v with Vint _ => True | Vlong _ => True | Vfloat _ => True | Vsingle _ => True | _ => False end -> @@ -709,11 +1343,13 @@ Proof. check_value n v q mvs = true -> In mv mvs -> (n < size_quantity_nat q)%nat -> exists j, mv = Fragment v q j /\ S j <> size_quantity_nat q). { - induction n; destruct mvs; simpl; intros; try discriminate. + unfold check_value. + induction n; destruct mvs; simpl; intros; InvBooleans; try discriminate. contradiction. - destruct m; try discriminate. InvBooleans. apply beq_nat_true in H4. subst. + destruct m; try discriminate. InvBooleans. apply beq_nat_true in H5. subst. destruct H0. subst mv. exists n0; split; auto. omega. - eapply IHn; eauto. omega. + eapply IHn; eauto. apply andb_true_intro; split; eauto using proj_sumbool_is_true. + omega. } assert (U: forall mvs, shape_decoding chunk mvs (Val.load_result chunk Vundef)). { @@ -724,16 +1360,20 @@ Proof. (chunk = Mint32 \/ chunk = Many32 \/ chunk = Mint64 \/ chunk = Many64) -> shape_decoding chunk (mv1 :: mvl) (Val.load_result chunk (proj_value q (mv1 :: mvl)))). { - intros. unfold proj_value. destruct mv1; auto. + intros. unfold proj_value. unfold check_value in *. destruct mv1; auto. destruct (size_quantity_nat_pos q) as [sz EQ]. rewrite EQ. - simpl. unfold proj_sumbool. rewrite dec_eq_true. - destruct (quantity_eq q q0); auto. - destruct (Nat.eqb sz n) eqn:EQN; auto. - destruct (check_value sz v q mvl) eqn:CHECK; auto. - simpl. apply beq_nat_true in EQN. subst n q0. constructor. auto. + simpl. unfold proj_sumbool. rewrite dec_eq_true; simpl. + destruct (quantity_eq q q0); try rewrite andb_false_r; auto. + destruct (Nat.eqb sz n) eqn:EQN; try rewrite andb_false_r; auto. + destruct (fits_quantity_dec v q); simpl; auto. + apply beq_nat_true in EQN. subst n q0. + destruct (check_value_rec sz v q mvl) eqn:CHECK; auto. + constructor. auto. destruct H0 as [E|[E|[E|E]]]; subst chunk; destruct q; auto || discriminate. congruence. - intros. eapply B; eauto. omega. + intros. eapply B; eauto. + rewrite proj_sumbool_is_true; simpl; eauto. + omega. } unfold decode_val. destruct (proj_bytes (mv1 :: mvl)) as [bl|] eqn:PB. @@ -762,6 +1402,17 @@ Proof. intros. inv H; econstructor. eapply val_inject_incr; eauto. Qed. +Lemma memval_inject_incr_list: + forall f f' vl vl', + list_forall2 (memval_inject f) vl vl' -> + inject_incr f f' -> + list_forall2 (memval_inject f') vl vl'. +Proof. + induction 1; intros. + apply list_forall2_nil. + apply list_forall2_cons; eauto using memval_inject_incr. +Qed. + (** [decode_val], applied to lists of memory values that are pairwise related by [memval_inject], returns values that are related by [Val.inject]. *) @@ -779,13 +1430,22 @@ Proof. congruence. Qed. -Lemma check_value_inject: +Lemma fits_quantity_inject: + forall f v v' q, + Val.inject f v v' -> v <> Vundef -> fits_quantity v q -> fits_quantity v' q. +Proof. + intros. destruct q. + inversion H; simpl; auto; subst; inv H1; congruence. + destruct v'; simpl; auto. +Qed. + +Lemma check_value_rec_inject: forall f vl vl', list_forall2 (memval_inject f) vl vl' -> forall v v' q n, - check_value n v q vl = true -> + check_value_rec n v q vl = true -> Val.inject f v v' -> v <> Vundef -> - check_value n v' q vl' = true. + check_value_rec n v' q vl' = true. Proof. induction 1; intros; destruct n; simpl in *; auto. inv H; auto. @@ -796,6 +1456,21 @@ Proof. discriminate. Qed. +Lemma check_value_inject: + forall f vl vl', + list_forall2 (memval_inject f) vl vl' -> + forall v v' q n, + check_value n v q vl = true -> + Val.inject f v v' -> v <> Vundef -> + check_value n v' q vl' = true. +Proof. + intros. unfold check_value in *; InvBooleans. + apply andb_true_intro; split. + unfold proj_sumbool. rewrite pred_dec_true; auto. + eauto using fits_quantity_inject. + eauto using check_value_rec_inject. +Qed. + Lemma proj_value_inject: forall f q vl1 vl2, list_forall2 (memval_inject f) vl1 vl2 -> @@ -822,25 +1497,6 @@ Proof. auto. Qed. -Lemma check_value_undef: - forall n q v vl, - In Undef vl -> check_value n v q vl = false. -Proof. - induction n; intros; simpl. - destruct vl. elim H. auto. - destruct vl. auto. - destruct m; auto. simpl in H; destruct H. congruence. - rewrite IHn; auto. apply andb_false_r. -Qed. - -Lemma proj_value_undef: - forall q vl, In Undef vl -> proj_value q vl = Vundef. -Proof. - intros; unfold proj_value. - destruct vl; auto. destruct m; auto. - rewrite check_value_undef. auto. auto. -Qed. - Theorem decode_val_inject: forall f vl1 vl2 chunk, list_forall2 (memval_inject f) vl1 vl2 -> @@ -894,15 +1550,45 @@ Proof. induction n; simpl; constructor; auto. constructor. Qed. -Lemma inj_value_inject: - forall f v1 v2 q, Val.inject f v1 v2 -> list_forall2 (memval_inject f) (inj_value q v1) (inj_value q v2). +Lemma inj_value_rec_inject: + forall f v1 v2 q n, + Val.inject f v1 v2 -> + list_forall2 (memval_inject f) (inj_value_rec n v1 q) (inj_value_rec n v2 q). Proof. intros. -Local Transparent inj_value. - unfold inj_value. generalize (size_quantity_nat q). induction n; simpl; constructor; auto. + generalize (size_quantity_nat q). induction n; simpl; constructor; auto. constructor; auto. Qed. +Lemma inj_value_inject_aux: + forall f v1 v2 q, + Val.inject f v1 v2 -> v1 <> Vundef -> + list_forall2 (memval_inject f) (inj_value q v1) (inj_value q v2). +Proof. + intros. +Local Transparent inj_value. + unfold inj_value. + rewrite dec_eq_false by auto. + destruct (Val.eq v2 Vundef). inv H; congruence. + destruct (fits_quantity_dec v1 q). + exploit fits_quantity_inject; eauto. intro. + rewrite pred_dec_true. apply inj_value_rec_inject; auto. auto. + destruct (fits_quantity_dec v2 q). + set (vl := inj_value_rec (size_quantity_nat q) v2 q). + assert (L: length vl = size_quantity_nat q) by apply inj_value_rec_length. + rewrite <- L; apply repeat_Undef_inject_any. + apply repeat_Undef_inject_self. +Qed. + +Lemma inj_value_inject: + forall f v1 v2 q, Val.inject f v1 v2 -> list_forall2 (memval_inject f) (inj_value q v1) (inj_value q v2). +Proof. + intros. destruct (Val.eq v1 Vundef). + - unfold inj_value at 1. subst v1; rewrite dec_eq_true. + erewrite <- inj_value_length. apply repeat_Undef_inject_any. + - auto using inj_value_inject_aux. +Qed. + Theorem encode_val_inject: forall f v1 v2 chunk, Val.inject f v1 v2 -> diff --git a/common/Memory.v b/common/Memory.v index 2cf1c3ab3b..adb64d628e 100644 --- a/common/Memory.v +++ b/common/Memory.v @@ -515,6 +515,14 @@ Proof. apply H. omega. apply IHn. intros. apply H. omega. Qed. +Remark getN_outside: + forall n p m c, + getN n (p + 1) (ZMap.set p m c) = getN n (p + 1) c. +Proof. + intros. apply getN_exten; intros i [L H]. + rewrite ZMap.gso; auto. apply Z.neq_sym. omega. +Qed. + Remark getN_setN_disjoint: forall vl q c n p, Intv.disjoint (p, p + Z.of_nat n) (q, q + Z.of_nat (length vl)) -> @@ -538,6 +546,68 @@ Proof. induction vl; simpl; intros. auto. rewrite IHvl. auto. Qed. +Remark getN_setN_prefix: + forall vl n p c, + (n <= length vl)%nat -> + getN n p (setN vl p c) = firstn n vl. +Proof. + induction vl; intros; simpl. + inv H; auto. + destruct n; simpl. auto. + decEq. + rewrite setN_outside, ZMap.gss; auto. omega. + apply IHvl. simpl in H; omega. +Qed. + +Lemma getN_concat: + forall c n1 n2 p, + getN (n1 + n2)%nat p c = getN n1 p c ++ getN n2 (p + Z.of_nat n1) c. +Proof. + induction n1; intros. + simpl. decEq. omega. + rewrite Nat2Z.inj_succ. simpl. decEq. + replace (p + Z.succ (Z.of_nat n1)) with ((p + 1) + Z.of_nat n1) by omega. + auto. +Qed. + +Lemma setN_concat: + forall bytes1 bytes2 ofs c, + setN (bytes1 ++ bytes2) ofs c = setN bytes2 (ofs + Z.of_nat (length bytes1)) (setN bytes1 ofs c). +Proof. + induction bytes1; intros. + simpl. decEq. omega. + simpl length. rewrite Nat2Z.inj_succ. simpl. rewrite IHbytes1. decEq. omega. +Qed. + +Remark getN_setN_concat: + forall n ln m lm p c, + (n = length ln)%nat -> + (m = length lm)%nat -> + getN (n + m)%nat p (setN (ln ++ lm) p c) = + getN n p (setN ln p c) ++ getN m (p + Z.of_nat n) (setN lm (p + Z.of_nat n) c). +Proof. + intros. + rewrite getN_concat, !setN_concat. + decEq. + rewrite getN_setN_outside, getN_setN_prefix; auto; omega. + rewrite <- H, !getN_setN_prefix; auto; omega. +Qed. + +Remark getN_setN_suffix: + forall vl n m p c, + (n + m = length vl)%nat -> + getN m (p + Z.of_nat n) (setN vl p c) = skipn n vl. +Proof. + intros. + assert (vl = firstn n vl ++ skipn n vl) by (rewrite firstn_skipn; auto). + assert (n = length (firstn n vl)) by (rewrite firstn_length_le; auto; omega). + assert (m = length (skipn n vl)). { rewrite H0, app_length, <- H1 in H. omega. } + rewrite H0, setN_concat. + rewrite <- H1, H2. + rewrite getN_setN_same. + rewrite <- H0; auto. +Qed. + (** [store chunk m b ofs v] perform a write in memory state [m]. Value [v] is stored at address [b] and offset [ofs]. Return the updated memory store, or [None] if the accessed bytes @@ -682,6 +752,14 @@ Proof. apply decode_val_type. Qed. +Theorem loadv_type: + forall m chunk addr v, + loadv chunk m addr = Some v -> + Val.has_type v (type_of_chunk chunk). +Proof. + unfold loadv; intros. destruct addr; try congruence. eauto using load_type. +Qed. + Theorem load_cast: forall m chunk b ofs v, load chunk m b ofs = Some v -> @@ -795,17 +873,6 @@ Proof. red; intros. omegaContradiction. Qed. -Lemma getN_concat: - forall c n1 n2 p, - getN (n1 + n2)%nat p c = getN n1 p c ++ getN n2 (p + Z.of_nat n1) c. -Proof. - induction n1; intros. - simpl. decEq. omega. - rewrite Nat2Z.inj_succ. simpl. decEq. - replace (p + Z.succ (Z.of_nat n1)) with ((p + 1) + Z.of_nat n1) by omega. - auto. -Qed. - Theorem loadbytes_concat: forall m b ofs n1 n2 bytes1 bytes2, loadbytes m b ofs n1 = Some bytes1 -> @@ -1150,6 +1217,23 @@ Proof. right. apply IHvl. omega. Qed. +Lemma setN_nth: + forall vl p q c m, + p <= q < p + Z.of_nat (length vl) -> + nth_error vl (Z.to_nat (q - p)) = Some m -> + ZMap.get q (setN vl p c) = m. +Proof. + induction vl; intros. + simpl in H. omegaContradiction. + simpl length in H. rewrite Nat2Z.inj_succ in H. simpl. + destruct (zeq p q). subst q. rewrite setN_outside. rewrite ZMap.gss. + rewrite Z.sub_diag in H0. simpl in H0; congruence. omega. + apply IHvl. omega. + assert (QP: q - p = Z.succ (q - (p + 1))) by omega. + rewrite QP, Z2Nat.inj_succ in H0 by omega. + simpl in H0. auto. +Qed. + Lemma getN_in: forall c q n p, p <= q < p + Z.of_nat n -> @@ -1579,15 +1663,6 @@ Qed. End STOREBYTES. -Lemma setN_concat: - forall bytes1 bytes2 ofs c, - setN (bytes1 ++ bytes2) ofs c = setN bytes2 (ofs + Z.of_nat (length bytes1)) (setN bytes1 ofs c). -Proof. - induction bytes1; intros. - simpl. decEq. omega. - simpl length. rewrite Nat2Z.inj_succ. simpl. rewrite IHbytes1. decEq. omega. -Qed. - Theorem storebytes_concat: forall m b ofs bytes1 m1 bytes2 m2, storebytes m b ofs bytes1 = Some m1 -> @@ -1829,7 +1904,7 @@ Proof. intros. exploit load_result; eauto. intro. rewrite H0. injection ALLOC; intros. rewrite <- H2; simpl. rewrite <- H1. rewrite PMap.gss. destruct (size_chunk_nat_pos chunk) as [n E]. rewrite E. simpl. - rewrite ZMap.gi. apply decode_val_undef. + rewrite ZMap.gi. apply decode_val_hd_undef. Qed. Theorem load_alloc_same': diff --git a/common/PrintAST.ml b/common/PrintAST.ml index e477957a47..b810e39d90 100644 --- a/common/PrintAST.ml +++ b/common/PrintAST.ml @@ -18,6 +18,7 @@ open Printf open Camlcoq open AST +open Memdata let name_of_type = function | Tint -> "int" @@ -27,6 +28,10 @@ let name_of_type = function | Tany32 -> "any32" | Tany64 -> "any64" +let name_of_quantity = function + | Q32 -> "Q32" + | Q64 -> "Q64" + let name_of_chunk = function | Mint8signed -> "int8s" | Mint8unsigned -> "int8u" @@ -83,10 +88,8 @@ let rec print_builtin_args px oc = function | a1 :: al -> fprintf oc "%a, %a" (print_builtin_arg px) a1 (print_builtin_args px) al -let rec print_builtin_res px oc = function +let print_builtin_res px oc = function | BR x -> px oc x | BR_none -> fprintf oc "_" - | BR_splitlong(hi, lo) -> - fprintf oc "splitlong(%a, %a)" - (print_builtin_res px) hi (print_builtin_res px) lo + | BR_splitlong(hi, lo) -> fprintf oc "splitlong(%a, %a)" px hi px lo diff --git a/common/Values.v b/common/Values.v index a20dd567f1..307160f7dc 100644 --- a/common/Values.v +++ b/common/Values.v @@ -89,6 +89,19 @@ Definition has_type (v: val) (t: typ) : Prop := | _, _ => False end. +Definition has_type_rpair {A} (v: val) (p: rpair A) (lof hif: val -> val) (tf: A -> typ) : Prop := + match p with + | One r => has_type v (tf r) + | Twolong hi lo => has_type (lof v) (tf lo) /\ has_type (hif v) (tf hi) + end. + +Definition has_type_builtin_res {A} (v: val) (r: builtin_res A) (lof hif: val -> val) (tf: A -> typ) : Prop := + match r with + | BR_none => True + | BR r => has_type v (tf r) + | BR_splitlong hi lo => has_type (lof v) (tf lo) /\ has_type (hif v) (tf hi) + end. + Fixpoint has_type_list (vl: list val) (tl: list typ) {struct vl} : Prop := match vl, tl with | nil, nil => True @@ -114,6 +127,12 @@ Proof. unfold has_type, Vnullptr, Tptr; destruct Archi.ptr64; reflexivity. Qed. +Lemma has_type_Tany64: + forall v, has_type v Tany64. +Proof. + destruct v; simpl; auto. +Qed. + Lemma has_subtype: forall ty1 ty2 v, subtype ty1 ty2 = true -> has_type v ty1 -> has_type v ty2. @@ -132,6 +151,35 @@ Proof. simpl in *. InvBooleans. destruct H0. split; auto. eapply has_subtype; eauto. Qed. +Definition has_type_b (v: val) (t: typ) : bool := + match v, t with + | Vundef, _ => true + | Vint _, Tint => true + | Vlong _, Tlong => true + | Vfloat _, Tfloat => true + | Vsingle _, Tsingle => true + | Vptr _ _, Tint => negb Archi.ptr64 + | Vptr _ _, Tlong => Archi.ptr64 + | (Vint _ | Vsingle _), Tany32 => true + | Vptr _ _, Tany32 => negb Archi.ptr64 + | _, Tany64 => true + | _, _ => false + end. + +Program Definition has_type_dec (v: val) (t: typ) : {has_type v t} + {~ has_type v t} := + match has_type_b v t with + | true => left _ + | false => right _ + end. +Next Obligation. + destruct v, t; simpl in *; auto; try congruence; destruct Archi.ptr64; auto. +Qed. +Next Obligation. + destruct v, t; simpl in *; auto; try congruence. + apply eq_sym, negb_false_iff in Heq_anonymous. congruence. + apply eq_sym, negb_false_iff in Heq_anonymous. congruence. +Qed. + (** Truth values. Non-zero integers are treated as [True]. The integer 0 (also used to represent the null pointer) is [False]. Other values are neither true nor false. *) @@ -938,6 +986,78 @@ Proof. destruct v; destruct ty; destruct Archi.ptr64; try contradiction; try discriminate; auto. Qed. +Lemma load_result_add: + forall c v1 v2, + c = Many32 \/ c = Many64 -> + load_result c (add v1 v2) = add v1 v2. +Proof. + intros. destruct H; subst c; destruct v1, v2; auto; simpl; destruct Archi.ptr64; auto. +Qed. + +Lemma load_result_or: + forall c v1 v2, + c = Many32 \/ c = Many64 -> + load_result c (or v1 v2) = or v1 v2. +Proof. + intros. destruct H; subst c; destruct v1, v2; auto. +Qed. + +Lemma load_result_and: + forall c v1 v2, + c = Many32 \/ c = Many64 -> + load_result c (and v1 v2) = and v1 v2. +Proof. + intros. destruct H; subst c; destruct v1, v2; auto. +Qed. + +Lemma load_result_shr: + forall c v1 v2, + c = Many32 \/ c = Many64 -> + load_result c (shr v1 v2) = shr v1 v2. +Proof. + intros. destruct H; subst c; destruct v1, v2; auto. + unfold shr. destruct (Int.ltu _ _); auto. +Qed. + +Lemma load_result_shru: + forall c v1 v2, + c = Many32 \/ c = Many64 -> + load_result c (shru v1 v2) = shru v1 v2. +Proof. + intros. destruct H; subst c; destruct v1, v2; auto. + unfold shru. destruct (Int.ltu _ _); auto. +Qed. + +Lemma load_result_of_optbool: + forall c ob, + c = Many32 \/ c = Many64 -> + load_result c (of_optbool ob) = of_optbool ob. +Proof. + intros. destruct H; subst c; destruct ob; try destruct b; auto. +Qed. + +Lemma load_result_intoffloat: + forall v v', + Val.intoffloat v = Some v' -> + load_result Many32 (Val.maketotal (Val.intoffloat v)) = v'. +Proof. + intros. destruct v; try inversion H. + rewrite H, <- H1. + destruct (Float.to_int f). simpl in *; congruence. + inversion H1. +Qed. + +Lemma load_result_intofsingle: + forall v v', + Val.intofsingle v = Some v' -> + load_result Many32 (Val.maketotal (Val.intofsingle v)) = v'. +Proof. + intros. destruct v; try inversion H. + rewrite H, <- H1. + destruct (Float32.to_int f). simpl in *; congruence. + inversion H1. +Qed. + (** Theorems on arithmetic operations. *) Theorem cast8unsigned_and: @@ -2044,6 +2164,13 @@ Proof. intros. destruct v; simpl; auto. f_equal. apply Ptrofs.add_assoc. Qed. +Lemma offset_ptr_type: + forall v d, has_type (offset_ptr v d) Tptr. +Proof. + intros. destruct v; simpl; auto. + unfold Tptr. destruct Archi.ptr64; auto. +Qed. + (** * Values and memory injections *) (** A memory injection [f] is a function from addresses to either [None] diff --git a/powerpc/Asm.v b/powerpc/Asm.v index aa15547bcf..f7a2f802b0 100644 --- a/powerpc/Asm.v +++ b/powerpc/Asm.v @@ -56,6 +56,30 @@ Proof. decide equality. Defined. Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}. Proof. decide equality. Defined. +Definition ireg_index (i: ireg): Z := + 2 + 2 * match i with + | GPR0 => 0 | GPR1 => 1 | GPR2 => 2 | GPR3 => 3 + | GPR4 => 4 | GPR5 => 5 | GPR6 => 6 | GPR7 => 7 + | GPR8 => 8 | GPR9 => 9 | GPR10 => 10 | GPR11 => 11 + | GPR12 => 12 | GPR13 => 13 | GPR14 => 14 | GPR15 => 15 + | GPR16 => 16 | GPR17 => 17 | GPR18 => 18 | GPR19 => 19 + | GPR20 => 20 | GPR21 => 21 | GPR22 => 22 | GPR23 => 23 + | GPR24 => 24 | GPR25 => 25 | GPR26 => 26 | GPR27 => 27 + | GPR28 => 28 | GPR29 => 29 | GPR30 => 30 | GPR31 => 31 + end. + +Definition freg_index (f: freg): Z := + 66 + 2 * match f with + | FPR0 => 0 | FPR1 => 1 | FPR2 => 2 | FPR3 => 3 + | FPR4 => 4 | FPR5 => 5 | FPR6 => 6 | FPR7 => 7 + | FPR8 => 8 | FPR9 => 9 | FPR10 => 10 | FPR11 => 11 + | FPR12 => 12 | FPR13 => 13 | FPR14 => 14 | FPR15 => 15 + | FPR16 => 16 | FPR17 => 17 | FPR18 => 18 | FPR19 => 19 + | FPR20 => 20 | FPR21 => 21 | FPR22 => 22 | FPR23 => 23 + | FPR24 => 24 | FPR25 => 25 | FPR26 => 26 | FPR27 => 27 + | FPR28 => 28 | FPR29 => 29 | FPR30 => 30 | FPR31 => 31 + end. + (** The PowerPC has a great many registers, some general-purpose, some very specific. We model only the following registers: *) @@ -78,12 +102,156 @@ Coercion FR: freg >-> preg. Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}. Proof. decide equality. apply ireg_eq. apply freg_eq. Defined. -Module PregEq. - Definition t := preg. - Definition eq := preg_eq. -End PregEq. +Definition preg_index (p: preg): Z := + match p with + | IR i => ireg_index i + | FR f => freg_index f + | PC => 132 + | LR => 134 + | CTR => 136 + | CARRY => 138 + | CR0_0 => 140 + | CR0_1 => 142 + | CR0_2 => 144 + | CR0_3 => 146 + | CR1_2 => 148 + end. -Module Pregmap := EMap(PregEq). +Lemma preg_index_bounds: + forall p: preg, + match p with + | IR _ => 0 < preg_index p /\ preg_index p <= 64 + | FR _ => 65 < preg_index p /\ preg_index p <= 128 + | _ => 131 < preg_index p + end. +Proof. + destruct p; unfold preg_index; try omega. + destruct i; unfold ireg_index; omega. + destruct f; unfold freg_index; omega. +Qed. + +Module Preg <: REGISTER_MODEL. + + Definition reg := preg. + Definition eq_dec := preg_eq. + + Definition type p := + match p with + | IR _ | PC | LR | CTR => if Archi.ppc64 then Tany64 else Tany32 + | FR _ => Tany64 + (* On PPC64 we sometimes want to write [Vlong] values into the carry + register, so type it accordingly. *) + | CARRY => if Archi.ppc64 then Tany64 else Tany32 + | _ => Tany32 (* condition bits *) + end. + + Definition quantity_of p := + match p with + | IR _ | PC | LR | CTR => if Archi.ppc64 then Q64 else Q32 + | FR _ => Q64 + | CARRY => if Archi.ppc64 then Q64 else Q32 + | _ => Q32 + end. + + Definition chunk_of p := + match p with + | IR _ | PC | LR | CTR => if Archi.ppc64 then Many64 else Many32 + | FR _ => Many64 + | CARRY => if Archi.ppc64 then Many64 else Many32 + | _ => Many32 + end. + + Lemma type_cases: forall r, type r = Tany32 \/ type r = Tany64. + Proof. + destruct r; simpl; auto; destruct Archi.ppc64; auto. + Qed. + + Definition ofs (r: preg): Z := + preg_index r. + + (* A register's address: The index of its first byte. *) + Definition addr (r: preg): Z := + preg_index r * 4. + + (* The address one byte past the end of register [r]. The next nonoverlapping + register may start here. *) + Definition next_addr (r: preg): Z := addr r + AST.typesize (type r). + + Remark addr_pos: forall p, addr p > 0. + Proof. + intros. unfold addr. + destruct p; simpl; try omega. + destruct i; simpl; omega. + destruct f; simpl; omega. + Qed. + + Lemma addr_compat: forall p, FragBlock.addr (ofs p) = addr p. + Proof. + reflexivity. + Qed. + + Lemma size_compat: + forall p, + AST.typesize (type p) = FragBlock.quantity_size (quantity_of p). + Proof. + intros. unfold quantity_of. + destruct p; simpl; auto; destruct Archi.ppc64; auto. + Qed. + + Lemma next_addr_compat: forall p, FragBlock.next_addr (ofs p) (quantity_of p) = next_addr p. + Proof. + unfold next_addr, addr, ofs, FragBlock.next_addr, FragBlock.addr; intros. + rewrite size_compat. auto. + Qed. + + Lemma quantity_of_compat: + forall p, + quantity_of p = quantity_of_typ (type p). + Proof. + destruct p; simpl; destruct Archi.ppc64; reflexivity. + Qed. + + Lemma chunk_of_reg_compat: + forall p, + chunk_of p = chunk_of_quantity (quantity_of p). + Proof. + destruct p; simpl; destruct Archi.ppc64; reflexivity. + Qed. + + Lemma chunk_of_reg_type: + forall p, + chunk_of p = chunk_of_type (type p). + Proof. + destruct p; simpl; destruct Archi.ppc64; reflexivity. + Qed. + + Lemma diff_outside_interval: + forall r s, r <> s -> next_addr r <= addr s \/ next_addr s <= addr r. + Proof. + intros. + unfold next_addr, addr. + assert (TS: forall p, AST.typesize (type p) = 4 \/ AST.typesize (type p) = 8). + { intro p. destruct (type_cases p) as [T | T]; rewrite T; auto. } + generalize (preg_index_bounds r), (preg_index_bounds s); intros RB SB. + destruct r eqn:R, s eqn:S; + try congruence; + try (destruct (TS r), (TS s); subst; omega); + simpl AST.typesize; + try (destruct Archi.ppc64; simpl; omega). + - (* two iregs *) + destruct Archi.ppc64; simpl; destruct i, i0; simpl; try omega; congruence. + - (* two fregs *) + destruct f, f0; simpl; try omega; congruence. + Qed. + +End Preg. + +Lemma pc_type: subtype Tptr (Preg.type PC) = true. +Proof. + simpl; destruct Archi.ppc64; auto. +Qed. + +Module Pregmap := Regfile(Preg). (** Conventional names for stack pointer ([SP]) and return address ([RA]) *) @@ -247,6 +415,8 @@ Inductive instruction : Type := | Plfdx_a: freg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) | Plfs: freg -> constant -> ireg -> instruction (**r load 32-bit float *) | Plfsx: freg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) + | Plfs_a: freg -> constant -> ireg -> instruction (**r load 32-bit float *) + | Plfsx_a: freg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) | Plha: ireg -> constant -> ireg -> instruction (**r load 16-bit signed int *) | Plhax: ireg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) | Plhbrx: ireg -> ireg -> ireg -> instruction (**r load 16-bit int and reverse endianness *) @@ -319,6 +489,8 @@ Inductive instruction : Type := | Pstfdx_a: freg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) | Pstfs: freg -> constant -> ireg -> instruction (**r store 32-bit float *) | Pstfsx: freg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) + | Pstfs_a: freg -> constant -> ireg -> instruction (**r store 32-bit float *) + | Pstfsx_a: freg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) | Psth: ireg -> constant -> ireg -> instruction (**r store 16-bit int *) | Psthx: ireg -> ireg -> ireg -> instruction (**r same, with 2 index regs *) | Psthbrx: ireg -> ireg -> ireg -> instruction (**r store 16-bit int with reverse endianness *) @@ -430,12 +602,14 @@ Definition program := AST.program fundef unit. and boolean registers ([CARRY], [CR0_0], etc) to either [Vzero] or [Vone]. *) -Definition regset := Pregmap.t val. +Definition regset := Pregmap.t. Definition genv := Genv.t fundef unit. -Notation "a # b" := (a b) (at level 1, only parsing) : asm. +Notation "a # b" := (Pregmap.get b a) (at level 1, only parsing) : asm. Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm. +Definition pregmap_read (rs: regset) := fun r => (Pregmap.get r rs). + Open Scope asm. (** Undefining some registers *) @@ -443,9 +617,26 @@ Open Scope asm. Fixpoint undef_regs (l: list preg) (rs: regset) : regset := match l with | nil => rs - | r :: l' => undef_regs l' (rs#r <- Vundef) + | r :: l' => (undef_regs l' rs)#r <- Vundef end. +Lemma undef_regs_other: + forall r rl rs, ~In r rl -> (undef_regs rl rs) # r = rs # r. +Proof. + induction rl; simpl; intros. auto. rewrite Pregmap.gso. apply IHrl. intuition. intuition. +Qed. + +Lemma undef_regs_same: + forall r rl rs, In r rl -> (undef_regs rl rs) # r = Vundef. +Proof. + induction rl; simpl; intros. tauto. + destruct H. + - subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + - destruct (Preg.eq_dec r a). + + subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + + rewrite Pregmap.gso; auto. +Qed. + (** Assigning a register pair *) Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := @@ -456,11 +647,11 @@ Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := (** Assigning the result of a builtin *) -Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := +Definition set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := match res with | BR r => rs#r <- v | BR_none => rs - | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) + | BR_splitlong hi lo => rs #hi <- (Val.hiword v) #lo <- (Val.loword v) end. Section RELSEM. @@ -710,7 +901,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pallocframe sz ofs _ => let (m1, stk) := Mem.alloc m 0 sz in let sp := Vptr stk Ptrofs.zero in - match Mem.storev Mint32 m1 (Val.offset_ptr sp ofs) rs#GPR1 with + match Mem.storev Many32 m1 (Val.offset_ptr sp ofs) rs#GPR1 with | None => Stuck | Some m2 => Next (nextinstr (rs#GPR1 <- sp #GPR0 <- Vundef)) m2 end @@ -757,7 +948,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | _ => Stuck end | Pbtbl r tbl => - match rs r with + match rs # r with | Vint n => match list_nth_z tbl (Int.unsigned n) with | None => Stuck @@ -802,7 +993,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pextzw rd r1 => Next (nextinstr (rs#rd <- (Val.longofintu rs#r1))) m | Pfreeframe sz ofs => - match Mem.loadv Mint32 m (Val.offset_ptr rs#GPR1 ofs) with + match Mem.loadv Many32 m (Val.offset_ptr rs#GPR1 ofs) with | None => Stuck | Some v => match rs#GPR1 with @@ -884,6 +1075,10 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out load1 Mfloat32 rd cst r1 rs m | Plfsx rd r1 r2 => load2 Mfloat32 rd r1 r2 rs m + | Plfs_a rd cst r1 => + load1 Many32 rd cst r1 rs m + | Plfsx_a rd r1 r2 => + load2 Many32 rd r1 r2 rs m | Plha rd cst r1 => load1 Mint16signed rd cst r1 rs m | Plhax rd r1 r2 => @@ -1003,6 +1198,10 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out store1 Mfloat32 rd cst r1 rs m | Pstfsx rd r1 r2 => store2 Mfloat32 rd r1 r2 rs m + | Pstfs_a rd cst r1 => + store1 Many32 rd cst r1 rs m + | Pstfsx_a rd r1 r2 => + store2 Many32 rd r1 r2 rs m | Psth rd cst r1 => store1 Mint16unsigned rd cst r1 rs m | Psthx rd r1 r2 => @@ -1132,25 +1331,57 @@ Definition preg_of (r: mreg) : preg := (** Undefine all registers except SP and callee-save registers *) -Definition undef_caller_save_regs (rs: regset) : regset := +Definition undef_caller_save_regs_spec (rs: preg -> val) : preg -> val := fun r => if preg_eq r SP || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) then rs r else Vundef. +Definition pregs_destroyed_at_call := + PC :: LR :: CTR :: CARRY :: CR0_0 :: CR0_1 :: CR0_2 :: CR0_3 :: CR1_2 + :: IR GPR0 :: IR GPR2 :: IR GPR13 + :: (map preg_of (filter (fun m => negb (is_callee_save m)) all_mregs)). + +Lemma pregs_destroyed_at_call_correct: + forall r, + preg_eq r SP + || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) + = + negb (In_dec preg_eq r pregs_destroyed_at_call). +Proof. + intros. + destruct r; auto. + destruct i; auto. + destruct f; auto. +Qed. + +Definition undef_caller_save_regs (rs: regset) : regset := + undef_regs pregs_destroyed_at_call rs. + +Lemma undef_caller_save_regs_correct: + forall rs r, + (undef_caller_save_regs rs) # r = undef_caller_save_regs_spec (fun r' => rs # r) r. +Proof. + intros. unfold undef_caller_save_regs, undef_caller_save_regs_spec. + rewrite pregs_destroyed_at_call_correct. + destruct (In_dec preg_eq r pregs_destroyed_at_call) as [IN | NOTIN]. + - rewrite undef_regs_same; auto. + - rewrite undef_regs_other; auto. +Qed. + (** Extract the values of the arguments of an external call. We exploit the calling conventions from module [Conventions], except that we use PPC registers instead of locations. *) Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop := | extcall_arg_reg: forall r, - extcall_arg rs m (R r) (rs (preg_of r)) - | extcall_arg_stack: forall ofs ty bofs v, + extcall_arg rs m (R r) (rs # (preg_of r)) + | extcall_arg_stack: forall ofs q bofs v, bofs = Stacklayout.fe_ofs_arg + 4 * ofs -> - Mem.loadv (chunk_of_type ty) m - (Val.offset_ptr (rs (IR GPR1)) (Ptrofs.repr bofs)) = Some v -> - extcall_arg rs m (S Outgoing ofs ty) v. + Mem.loadv (chunk_of_quantity q) m + (Val.offset_ptr (rs # (IR GPR1)) (Ptrofs.repr bofs)) = Some v -> + extcall_arg rs m (S Outgoing ofs q) v. Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop := | extcall_arg_one: forall l v, @@ -1176,17 +1407,17 @@ Inductive state: Type := Inductive step: state -> trace -> state -> Prop := | exec_step_internal: forall b ofs f i rs m rs' m', - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some i -> exec_instr f i rs m = Next rs' m' -> step (State rs m) E0 (State rs' m') | exec_step_builtin: forall b ofs f ef args res rs m vargs t vres rs' m', - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) -> - eval_builtin_args ge rs (rs GPR1) m args vargs -> + eval_builtin_args ge (fun r => rs#r) (rs # GPR1) m args vargs -> external_call ef ge vargs m t vres m' -> rs' = nextinstr (set_res res vres @@ -1194,11 +1425,11 @@ Inductive step: state -> trace -> state -> Prop := step (State rs m) t (State rs' m') | exec_step_external: forall b ef args res rs m t rs' m', - rs PC = Vptr b Ptrofs.zero -> + rs # PC = Vptr b Ptrofs.zero -> Genv.find_funct_ptr ge b = Some (External ef) -> external_call ef ge args m t res m' -> extcall_arguments rs m (ef_sig ef) args -> - rs' = (set_pair (loc_external_result (ef_sig ef)) res (undef_caller_save_regs rs)) #PC <- (rs RA) -> + rs' = (set_pair (loc_external_result (ef_sig ef)) res (undef_caller_save_regs rs)) #PC <- (rs#RA) -> step (State rs m) t (State rs' m'). End RELSEM. @@ -1210,7 +1441,7 @@ Inductive initial_state (p: program): state -> Prop := Genv.init_mem p = Some m0 -> let ge := Genv.globalenv p in let rs0 := - (Pregmap.init Vundef) + Pregmap.init # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero) # LR <- Vnullptr # GPR1 <- Vnullptr in diff --git a/powerpc/AsmToJSON.ml b/powerpc/AsmToJSON.ml index ee3eaca8a9..73598c58fd 100644 --- a/powerpc/AsmToJSON.ml +++ b/powerpc/AsmToJSON.ml @@ -243,8 +243,10 @@ let pp_instructions pp ic = | Plfd_a (fr,c,ir) -> instruction pp "Plfd" [Freg fr; Constant c; Ireg ir] | Plfdx (fr,ir1,ir2) | Plfdx_a (fr,ir1,ir2) -> instruction pp "Plfdx" [Freg fr; Ireg ir1; Ireg ir2] - | Plfs (fr,c,ir) -> instruction pp "Plfs" [Freg fr; Constant c; Ireg ir] - | Plfsx (fr,ir1,ir2) -> instruction pp "Plfsx" [Freg fr; Ireg ir1; Ireg ir2] + | Plfs (fr,c,ir) + | Plfs_a (fr,c,ir) -> instruction pp "Plfs" [Freg fr; Constant c; Ireg ir] + | Plfsx (fr,ir1,ir2) + | Plfsx_a (fr,ir1,ir2) -> instruction pp "Plfsx" [Freg fr; Ireg ir1; Ireg ir2] | Plha (ir1,c,ir2) -> instruction pp "Plha" [Ireg ir1; Constant c; Ireg ir2] | Plhax (ir1,ir2,ir3) -> instruction pp "Plhax" [Ireg ir1; Ireg ir2; Ireg ir3] | Plhbrx (ir1,ir2,ir3) -> instruction pp "Plhbrx" [Ireg ir1; Ireg ir2; Ireg ir3] @@ -315,8 +317,10 @@ let pp_instructions pp ic = | Pstfdu (fr,c,ir) -> instruction pp "Pstfdu" [Freg fr; Constant c; Ireg ir] | Pstfdx (fr,ir1,ir2) | Pstfdx_a (fr,ir1,ir2) -> instruction pp "Pstfdx" [Freg fr; Ireg ir1; Ireg ir2] - | Pstfs (fr,c,ir) -> instruction pp "Pstfs" [Freg fr; Constant c; Ireg ir] - | Pstfsx (fr,ir1,ir2) -> instruction pp "Pstfsx" [Freg fr; Ireg ir1; Ireg ir2] + | Pstfs (fr,c,ir) + | Pstfs_a (fr,c,ir) -> instruction pp "Pstfs" [Freg fr; Constant c; Ireg ir] + | Pstfsx (fr,ir1,ir2) + | Pstfsx_a (fr,ir1,ir2) -> instruction pp "Pstfsx" [Freg fr; Ireg ir1; Ireg ir2] | Psth (ir1,c,ir2) -> instruction pp "Psth" [Ireg ir1; Constant c; Ireg ir2] | Psthx (ir1,ir2,ir3) -> instruction pp "Psthx" [Ireg ir1; Ireg ir2; Ireg ir3] | Psthbrx (ir1,ir2,ir3) -> instruction pp "Psthbrx" [Ireg ir1; Ireg ir2; Ireg ir3] diff --git a/powerpc/Asmexpand.ml b/powerpc/Asmexpand.ml index f880e2d181..6927a523b1 100644 --- a/powerpc/Asmexpand.ml +++ b/powerpc/Asmexpand.ml @@ -275,7 +275,7 @@ let expand_builtin_vload_1 chunk addr res = (fun r c -> emit (Pld(res, c, r))) (fun r1 r2 -> emit (Pldx(res, r1, r2))) addr GPR11 - | Mint64, BR_splitlong(BR(IR hi), BR(IR lo)) -> + | Mint64, BR_splitlong(IR hi, IR lo) -> expand_volatile_access (fun r c -> match offset_constant c _4 with @@ -527,24 +527,24 @@ let expand_builtin_inline name args res = emit (Pcfi_adjust _m8) (* 64-bit integer arithmetic *) | "__builtin_negl", [BA_splitlong(BA(IR ah), BA(IR al))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah) rl (fun rl -> emit (Psubfic(rl, al, Cint _0)); emit (Psubfze(rh, ah))) | "__builtin_addl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah || rl = bh) rl (fun rl -> emit (Paddc(rl, al, bl)); emit (Padde(rh, ah, bh))) | "__builtin_subl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah || rl = bh) rl (fun rl -> emit (Psubfc(rl, bl, al)); emit (Psubfe(rh, bh, ah))) | "__builtin_mull", [BA(IR a); BA(IR b)], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = a || rl = b) rl (fun rl -> emit (Pmullw(rl, a, b)); emit (Pmulhwu(rh, a, b))) diff --git a/powerpc/Asmgen.v b/powerpc/Asmgen.v index 8c296f0ae2..2ad35ed536 100644 --- a/powerpc/Asmgen.v +++ b/powerpc/Asmgen.v @@ -181,28 +181,22 @@ Definition accessind {A: Type} then instr1 r (Cint ofs) base :: k else loadimm GPR0 ofs (instr2 r base GPR0 :: k). -Definition loadind (base: ireg) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) := - match ty, preg_of dst with - | Tint, IR r => OK(accessind Plwz Plwzx base ofs r k) - | Tany32, IR r => OK(accessind Plwz_a Plwzx_a base ofs r k) - | Tsingle, FR r => OK(accessind Plfs Plfsx base ofs r k) - | Tlong, IR r => OK(accessind Pld Pldx base ofs r k) - | Tfloat, FR r => OK(accessind Plfd Plfdx base ofs r k) - | Tany64, IR r => OK(accessind Pld_a Pldx_a base ofs r k) - | Tany64, FR r => OK(accessind Plfd_a Plfdx_a base ofs r k) - | _, _ => Error (msg "Asmgen.loadind") +Definition loadind (base: ireg) (ofs: ptrofs) (q: quantity) (dst: mreg) (k: code) := + match q, preg_of dst, Archi.ppc64 with + | Q32, IR r, false => OK(accessind Plwz_a Plwzx_a base ofs r k) + | Q64, IR r, true => OK(accessind Pld_a Pldx_a base ofs r k) + | Q32, FR r, _ => OK(accessind Plfs_a Plfsx_a base ofs r k) + | Q64, FR r, _ => OK(accessind Plfd_a Plfdx_a base ofs r k) + | _, _, _ => Error (msg "Asmgen.loadind") end. -Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) := - match ty, preg_of src with - | Tint, IR r => OK(accessind Pstw Pstwx base ofs r k) - | Tany32, IR r => OK(accessind Pstw_a Pstwx_a base ofs r k) - | Tsingle, FR r => OK(accessind Pstfs Pstfsx base ofs r k) - | Tlong, IR r => OK(accessind Pstd Pstdx base ofs r k) - | Tfloat, FR r => OK(accessind Pstfd Pstfdx base ofs r k) - | Tany64, IR r => OK(accessind Pstd_a Pstdx_a base ofs r k) - | Tany64, FR r => OK(accessind Pstfd_a Pstfdx_a base ofs r k) - | _, _ => Error (msg "Asmgen.storeind") +Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (q: quantity) (k: code) := + match q, preg_of src, Archi.ppc64 with + | Q32, IR r, false => OK(accessind Pstw_a Pstwx_a base ofs r k) + | Q64, IR r, true => OK(accessind Pstd_a Pstdx_a base ofs r k) + | Q32, FR r, _ => OK(accessind Pstfs_a Pstfsx_a base ofs r k) + | Q64, FR r, _ => OK(accessind Pstfd_a Pstfdx_a base ofs r k) + | _, _, _ => Error (msg "Asmgen.storeind") end. (** Constructor for a floating-point comparison. The PowerPC has @@ -588,6 +582,7 @@ Definition transl_op do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- freg_of res; OK (Pfmake r r1 r2 :: k) | Omakelong, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Plmake r r1 r2 :: k) | Olowlong, a1 :: nil => assertion (mreg_eq a1 res); @@ -598,83 +593,110 @@ Definition transl_op transl_cond_op cmp args res k (*c PPC64 operations *) | Olongconst n, nil => + assertion Archi.ppc64; do r <- ireg_of res; OK (loadimm64 r n k) | Ocast32signed, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pextsw r r1 :: k) | Ocast32unsigned, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pextzw r r1 :: k) | Oaddl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Padd64 r r1 r2 :: k) | Oaddlimm n, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (addimm64 r r1 n k) | Osubl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Psubfc64 r r2 r1 :: k) | Onegl, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Psubfic64 r r1 Int64.zero :: k) | Omull, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pmulld r r1 r2 :: k) | Omullhs, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pmulhd r r1 r2 :: k) | Omullhu, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pmulhdu r r1 r2 :: k) | Odivl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pdivd r r1 r2 :: k) | Odivlu, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pdivdu r r1 r2 :: k) | Oandl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pand_64 r r1 r2 :: k) | Oandlimm n, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (andimm64 r r1 n k) | Oorl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Por64 r r1 r2 :: k) | Oorlimm n, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (orimm64 r r1 n k) | Oxorl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Pxor64 r r1 r2 :: k) | Oxorlimm n, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (xorimm64 r r1 n k) | Onotl, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pnor64 r r1 r1 :: k) | Oshll, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Psld r r1 r2 :: k) | Oshrl, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Psrad r r1 r2 :: k) | Oshrlimm n, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Psradi r r1 n :: k) | Oshrlu, a1 :: a2 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; do r <- ireg_of res; OK (Psrd r r1 r2 :: k) | Orolml amount mask, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (rolm64 r r1 amount mask k) | Oshrxlimm n, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Psradi r r1 n :: Paddze64 r r :: k) | Olongoffloat, a1 :: nil => + assertion Archi.ppc64; do r1 <- freg_of a1; do r <- ireg_of res; OK (Pfctid r r1 :: k) | Ofloatoflong, a1 :: nil => + assertion Archi.ppc64; do r1 <- ireg_of a1; do r <- freg_of res; OK (Pfcfl r r1 :: k) | _, _ => @@ -755,6 +777,7 @@ Definition transl_load (chunk: memory_chunk) (addr: addressing) do r <- ireg_of dst; transl_memory_access (Plwz r) (Plwzx r) addr args GPR12 k | Mint64 => + assertion Archi.ppc64; do r <- ireg_of dst; transl_memory_access (Pld r) (Pldx r) addr args GPR12 k | Mfloat32 => @@ -781,6 +804,7 @@ Definition transl_store (chunk: memory_chunk) (addr: addressing) do r <- ireg_of src; transl_memory_access (Pstw r) (Pstwx r) addr args temp k | Mint64 => + assertion Archi.ppc64; do r <- ireg_of src; transl_memory_access (Pstd r) (Pstdx r) addr args temp k | Mfloat32 => @@ -801,7 +825,7 @@ Definition transl_epilogue (f: Mach.function) (k: code) := if is_leaf_function f then Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k else - Plwz GPR0 (Cint (Ptrofs.to_int f.(fn_retaddr_ofs))) GPR1 :: + Plwz_a GPR0 (Cint (Ptrofs.to_int f.(fn_retaddr_ofs))) GPR1 :: Pmtlr GPR0 :: Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k. @@ -810,16 +834,16 @@ Definition transl_epilogue (f: Mach.function) (k: code) := Definition transl_instr (f: Mach.function) (i: Mach.instruction) (r11_is_parent: bool) (k: code) := match i with - | Mgetstack ofs ty dst => - loadind GPR1 ofs ty dst k - | Msetstack src ofs ty => - storeind src GPR1 ofs ty k - | Mgetparam ofs ty dst => + | Mgetstack ofs q dst => + loadind GPR1 ofs q dst k + | Msetstack src ofs q => + storeind src GPR1 ofs q k + | Mgetparam ofs q dst => if r11_is_parent then - loadind GPR11 ofs ty dst k + loadind GPR11 ofs q dst k else - (do k1 <- loadind GPR11 ofs ty dst k; - loadind GPR1 f.(fn_link_ofs) Tint R11 k1) + (do k1 <- loadind GPR11 ofs q dst k; + loadind GPR1 f.(fn_link_ofs) Q32 R11 k1) | Mop op args res => transl_op op args res k | Mload chunk addr args dst => @@ -857,8 +881,8 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction) Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool := match i with - | Msetstack src ofs ty => before - | Mgetparam ofs ty dst => negb (mreg_eq dst R11) + | Msetstack src ofs q => before + | Mgetparam ofs q dst => negb (mreg_eq dst R11) | Mop Omove args res => before && negb (mreg_eq res R11) | _ => false end. @@ -899,7 +923,7 @@ Definition transl_function (f: Mach.function) := OK (mkfunction f.(Mach.fn_sig) (Pallocframe f.(fn_stacksize) f.(fn_link_ofs) f.(fn_retaddr_ofs) :: Pmflr GPR0 :: - Pstw GPR0 (Cint (Ptrofs.to_int f.(fn_retaddr_ofs))) GPR1 :: + Pstw_a GPR0 (Cint (Ptrofs.to_int f.(fn_retaddr_ofs))) GPR1 :: Pcfi_rel_offset (Ptrofs.to_int f.(fn_retaddr_ofs)) :: c)). Definition transf_function (f: Mach.function) : res Asm.function := diff --git a/powerpc/Asmgenproof.v b/powerpc/Asmgenproof.v index 8ad28aea03..fde2097c04 100644 --- a/powerpc/Asmgenproof.v +++ b/powerpc/Asmgenproof.v @@ -74,7 +74,7 @@ Qed. Lemma exec_straight_exec: forall fb f c ep tf tc c' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs # PC) fb f c ep tf tc -> exec_straight tge tf tc rs m c' rs' m' -> plus step tge (State rs m) E0 (State rs' m'). Proof. @@ -86,10 +86,10 @@ Qed. Lemma exec_straight_at: forall fb f c ep tf tc c' ep' tc' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs # PC) fb f c ep tf tc -> transl_code f c' ep' = OK tc' -> exec_straight tge tf tc rs m tc' rs' m' -> - transl_code_at_pc ge (rs' PC) fb f c' ep' tf tc'. + transl_code_at_pc ge (rs' # PC) fb f c' ep' tf tc'. Proof. intros. inv H. exploit exec_straight_steps_2; eauto. @@ -353,11 +353,11 @@ Lemma find_label_goto_label: forall f tf lbl rs m c' b ofs, Genv.find_funct_ptr ge b = Some (Internal f) -> transf_function f = OK tf -> - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> exists tc', exists rs', goto_label tf lbl rs m = Next rs' m - /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc' + /\ transl_code_at_pc ge (rs' # PC) b f c' false tf tc' /\ forall r, r <> PC -> rs'#r = rs#r. Proof. intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2. @@ -366,7 +366,9 @@ Proof. intros [pos' [P [Q R]]]. exists tc; exists (rs#PC <- (Vptr b (Ptrofs.repr pos'))). split. unfold goto_label. rewrite P. rewrite H1. auto. - split. rewrite Pregmap.gss. constructor; auto. + split. rewrite Pregmap.gss. + rewrite Preg.chunk_of_reg_type, Val.load_result_same by (simpl; destruct Archi.ppc64; auto). + constructor; auto. rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q. auto. omega. generalize (transf_function_no_overflow _ _ H0). omega. @@ -413,7 +415,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (FIND: Genv.find_funct_ptr ge fb = Some (Internal f)) (MEXT: Mem.extends m m') - (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc) + (AT: transl_code_at_pc ge (rs#PC) fb f c ep tf tc) (AG: agree ms sp rs) (DXP: ep = true -> rs#GPR11 = parent_sp s) (LEAF: is_leaf_function f = true -> rs#LR = parent_ra s), @@ -424,8 +426,8 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = Vptr fb Ptrofs.zero) - (ATLR: rs RA = parent_ra s), + (ATPC: rs#PC = Vptr fb Ptrofs.zero) + (ATLR: rs#RA = parent_ra s), match_states (Mach.Callstate s fb ms m) (Asm.State rs m') | match_states_return: @@ -433,7 +435,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = parent_ra s), + (ATPC: rs#PC = parent_ra s), match_states (Mach.Returnstate s ms m) (Asm.State rs m'). @@ -442,7 +444,7 @@ Lemma exec_straight_steps: match_stack ge s -> Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> (is_leaf_function f = true -> rs1#LR = parent_ra s) -> (forall k c (TR: transl_instr f i ep k = OK c), exists rs2, @@ -467,7 +469,7 @@ Lemma exec_straight_steps_goto: Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> it1_is_parent ep i = false -> (is_leaf_function f = true -> rs1#LR = parent_ra s) -> (forall k c (TR: transl_instr f i ep k = OK c), @@ -534,10 +536,10 @@ Proof. - (* Mlabel *) left; eapply exec_straight_steps; eauto; intros. - monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto. + monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. split. apply agree_nextinstr; auto. split. simpl; congruence. - auto with asmgen. + Simpl. - (* Mgetstack *) unfold load_stack in H. @@ -552,7 +554,7 @@ Proof. - (* Msetstack *) unfold store_stack in H. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (regmap_get src rs) (rs0#(preg_of src))). eapply preg_val; eauto. 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. @@ -591,7 +593,10 @@ Opaque loadind. split. eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto. instantiate (1 := rs1#GPR11 <- (rs2#GPR11)). intros. rewrite Pregmap.gso; auto with asmgen. - congruence. intros. unfold Pregmap.set. destruct (PregEq.eq r' GPR11). congruence. auto with asmgen. + congruence. intros. + destruct (Preg.eq_dec r' GPR11). subst; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto using Pregmap.get_has_type. + rewrite Pregmap.gso; auto with asmgen. split. simpl; intros. rewrite U; auto with asmgen. apply preg_of_not_GPR11; auto. rewrite U; auto with asmgen. @@ -629,7 +634,7 @@ Opaque loadind. rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1. intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (regmap_get src rs) (rs0#(preg_of src))). 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]]. @@ -645,10 +650,10 @@ Opaque loadind. eapply transf_function_no_overflow; eauto. destruct ros as [rf|fid]; simpl in H; monadInv H5. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0#x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto. generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. generalize (code_tail_next_int _ _ _ _ NOOV CT1). intro CT2. @@ -665,9 +670,10 @@ Opaque loadind. traceEq. econstructor; eauto. econstructor; eauto. - eapply agree_sp_def; eauto. + eapply agree_sp_def; eauto. eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simpl. - Simpl. rewrite <- H2. auto. + Simpl. simpl. destruct Archi.ppc64; rewrite H7; simpl; auto. + Simpl. rewrite <- H2. simpl. destruct Archi.ppc64; auto. + (* Direct call *) generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x). @@ -679,9 +685,10 @@ Opaque loadind. simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. eauto. econstructor; eauto. econstructor; eauto. - eapply agree_sp_def; eauto. + eapply agree_sp_def; eauto. eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simpl. - Simpl. rewrite <- H2. auto. + Simpl. simpl. destruct Archi.ppc64; auto. + Simpl. rewrite <- H2. simpl. destruct Archi.ppc64; auto. - (* Mtailcall *) assert (f0 = f) by congruence. subst f0. @@ -690,24 +697,24 @@ Opaque loadind. eapply transf_function_no_overflow; eauto. destruct ros as [rf|fid]; simpl in H; monadInv H7. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0#x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H7; intros LD; inv LD; auto. set (rs1 := nextinstr (rs0#CTR <- (Vptr f' Ptrofs.zero))). assert (AG1: agree rs (Vptr stk soff) rs1). { apply agree_nextinstr. apply agree_set_other; auto. } - exploit transl_epilogue_correct; eauto. + exploit transl_epilogue_correct; eauto. unfold rs1; Simpl. intros (rs2 & m2' & A & B & C & D & E & F). assert (A': exec_straight tge tf (Pmtctr x0 :: transl_epilogue f (Pbctr sig :: x)) rs0 m'0 (Pbctr sig :: x) rs2 m2'). - { apply exec_straight_step with rs1 m'0. simpl. rewrite H9. reflexivity. reflexivity. eexact A. } + { apply exec_straight_step with rs1 m'0. simpl. rewrite H9. reflexivity. unfold rs1; Simpl. eexact A. } clear A. exploit exec_straight_steps_2; eauto using functions_transl. intros (ofs' & U & V). - set (rs3 := rs2#PC <- (rs2 CTR)). + set (rs3 := rs2#PC <- (rs2#CTR)). left; exists (State rs3 m2'); split. (* execution *) eapply plus_right'. eapply exec_straight_exec; eauto. @@ -715,7 +722,8 @@ Opaque loadind. traceEq. (* match states *) econstructor; eauto. apply agree_set_other; auto. - unfold rs3; Simpl. rewrite F by congruence. reflexivity. + unfold rs3. rewrite F by congruence. unfold rs1. Simpl. simpl. destruct Archi.ppc64; auto. + unfold rs3. rewrite F by congruence. unfold rs1. Simpl. + (* Direct call *) exploit transl_epilogue_correct; eauto. intros (rs2 & m2' & A & B & C & D & E & F). @@ -730,11 +738,13 @@ Opaque loadind. traceEq. (* match states *) econstructor; eauto. apply agree_set_other; auto. + unfold rs3. Simpl. simpl. destruct Archi.ppc64; auto. + unfold rs3. Simpl. - (* Mbuiltin *) - inv AT. monadInv H4. + inv AT. monadInv H5. exploit functions_transl; eauto. intro FN. - generalize (transf_function_no_overflow _ _ H3); intro NOOV. + generalize (transf_function_no_overflow _ _ H4); intro NOOV. exploit builtin_args_match; eauto. intros [vargs' [P Q]]. exploit external_call_mem_extends; eauto. intros [vres' [m2' [A [B [C D]]]]]. @@ -749,12 +759,19 @@ Opaque loadind. instantiate (2 := tf); instantiate (1 := x). unfold nextinstr. rewrite Pregmap.gss. rewrite set_res_other. rewrite undef_regs_other_2. - rewrite <- H1. simpl. econstructor; eauto. - eapply code_tail_next_int; eauto. + rewrite <- H2. simpl. + destruct Archi.ppc64; econstructor; eauto; eapply code_tail_next_int; eauto. rewrite preg_notin_charact. intros. auto with asmgen. auto with asmgen. apply agree_nextinstr. eapply agree_set_res; auto. eapply agree_undef_regs; eauto. intros; apply undef_regs_other_2; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + destruct res; simpl in H1; InvBooleans; simpl; auto. + eapply Val.has_subtype; eauto. + split. + destruct vres'; simpl; auto; destruct (mreg_type lo); auto; congruence. + destruct vres'; simpl; auto; destruct (mreg_type hi); auto; congruence. congruence. intros. Simpl. rewrite set_res_other by auto. rewrite undef_regs_other_2; auto with asmgen. @@ -779,7 +796,7 @@ Opaque loadind. left; eapply exec_straight_steps_goto; eauto. intros. simpl in TR. destruct (transl_cond_correct_1 tge tf cond args _ rs0 m' _ TR) as [rs' [A [B C]]]. - rewrite EC in B. + unfold pregmap_read in EC. rewrite EC in B. destruct (snd (crbit_for_cond cond)). (* Pbt, taken *) econstructor; econstructor; econstructor; split. eexact A. @@ -796,12 +813,12 @@ Opaque loadind. exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC. left; eapply exec_straight_steps; eauto. intros. simpl in TR. destruct (transl_cond_correct_1 tge tf cond args _ rs0 m' _ TR) as [rs' [A [B C]]]. - rewrite EC in B. + unfold pregmap_read in EC. rewrite EC in B. econstructor; split. eapply exec_straight_trans. eexact A. destruct (snd (crbit_for_cond cond)). - apply exec_straight_one. simpl. rewrite B. reflexivity. auto. - apply exec_straight_one. simpl. rewrite B. reflexivity. auto. + apply exec_straight_one. simpl. rewrite B. reflexivity. Simpl. + apply exec_straight_one. simpl. rewrite B. reflexivity. Simpl. split. eapply agree_undef_regs; eauto with asmgen. intros; Simpl. split. simpl. congruence. @@ -848,6 +865,11 @@ Local Transparent destroyed_by_jumptable. traceEq. (* match states *) econstructor; eauto. apply agree_set_other; auto. + unfold rs3; Simpl. + exploit parent_ra_type; eauto; intro. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tptr); auto. + simpl; destruct Archi.ppc64; auto. - (* internal function *) exploit functions_translated; eauto. intros [tf [A B]]. monadInv B. @@ -866,7 +888,7 @@ Local Transparent destroyed_by_jumptable. set (tfbody := Pallocframe (fn_stacksize f) (fn_link_ofs f) (fn_retaddr_ofs f) :: Pmflr GPR0 - :: Pstw GPR0 (Cint (Ptrofs.to_int (fn_retaddr_ofs f))) + :: Pstw_a GPR0 (Cint (Ptrofs.to_int (fn_retaddr_ofs f))) GPR1 :: Pcfi_rel_offset (Ptrofs.to_int (fn_retaddr_ofs f)) :: x0) in *. @@ -875,6 +897,7 @@ Local Transparent destroyed_by_jumptable. set (rs3 := nextinstr (rs2#GPR0 <- (rs0#LR))). set (rs4 := nextinstr rs3). set (rs5 := nextinstr rs4). + simpl chunk_of_quantity in *. assert (EXEC_PROLOGUE: exec_straight tge tf tf.(fn_code) rs0 m' @@ -882,30 +905,41 @@ Local Transparent destroyed_by_jumptable. change (fn_code tf) with tfbody; unfold tfbody. apply exec_straight_step with rs2 m2'. unfold exec_instr. rewrite C. fold sp. - rewrite <- (sp_val _ _ _ AG). rewrite F. auto. auto. + rewrite <- (sp_val _ _ _ AG). rewrite F. auto. unfold rs2; Simpl. apply exec_straight_step with rs3 m2'. - simpl. auto. auto. + simpl. unfold rs2; Simpl. unfold rs3; Simpl. apply exec_straight_two with rs4 m3'. - simpl. unfold store1. rewrite gpr_or_zero_not_zero. - change (rs3 GPR1) with sp. change (rs3 GPR0) with (rs0 LR). - simpl const_low. rewrite ATLR. erewrite storev_offset_ptr by eexact P. auto. congruence. - auto. auto. auto. + simpl. unfold store1. rewrite gpr_or_zero_not_zero by congruence. + replace (rs3#GPR1) with sp. replace (rs3#GPR0) with (rs0#LR). + simpl const_low. rewrite ATLR. erewrite storev_offset_ptr by eexact P. auto. + unfold rs3; Simpl. + replace (Preg.chunk_of GPR0) with (chunk_of_type (Preg.type LR)) by (simpl; destruct Archi.ppc64; auto). + rewrite Val.load_result_same; auto using Pregmap.get_has_type. + unfold rs3, rs2, sp; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + simpl; destruct Archi.ppc64; auto. + auto. + unfold rs4; Simpl. + unfold rs5; Simpl. left; exists (State rs5 m3'); split. eapply exec_straight_steps_1; eauto. omega. constructor. econstructor; eauto. - change (rs5 PC) with (Val.offset_ptr (Val.offset_ptr (Val.offset_ptr (Val.offset_ptr (rs0 PC) Ptrofs.one) Ptrofs.one) Ptrofs.one) Ptrofs.one). + replace (rs5#PC) with (Val.offset_ptr (Val.offset_ptr (Val.offset_ptr (Val.offset_ptr (rs0#PC) Ptrofs.one) Ptrofs.one) Ptrofs.one) Ptrofs.one). rewrite ATPC. simpl. constructor; eauto. eapply code_tail_next_int. omega. eapply code_tail_next_int. omega. eapply code_tail_next_int. omega. eapply code_tail_next_int. omega. constructor. + unfold rs5, rs4, rs3, rs2; Simpl. unfold rs5, rs4, rs3, rs2. apply agree_nextinstr. apply agree_nextinstr. apply agree_set_other; auto. apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto. eapply agree_change_sp; eauto. unfold sp; congruence. + apply Val.Vptr_has_type. simpl. destruct Archi.ppc64; auto. congruence. + intro. rewrite <- ATLR. unfold rs5, rs4, rs3, rs2; Simpl. - (* external function *) exploit functions_translated; eauto. @@ -921,6 +955,21 @@ Local Transparent destroyed_by_jumptable. unfold loc_external_result. apply agree_set_other; auto. apply agree_set_pair; auto. apply agree_undef_caller_save_regs; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + unfold Val.has_type_rpair in H3. + unfold loc_result, proj_sig_res, loc_result_64, loc_result_32 in *. + destruct Archi.ptr64, (sig_res (ef_sig ef)). + destruct t0; simpl in *; auto. + simpl in *; auto. + destruct t0; simpl in *; auto. + simpl in *; auto. + rewrite Pregmap.gss. rewrite ATLR. + exploit parent_ra_type; eauto; intro. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tptr); auto. + simpl; destruct Archi.ppc64; auto. + - (* return *) inv STACKS. simpl in *. right. split. omega. split. auto. @@ -943,7 +992,13 @@ Proof. econstructor; eauto. constructor. apply Mem.extends_refl. - split. auto. simpl. unfold Vnullptr; simpl; congruence. intros. rewrite Regmap.gi. auto. + split. auto. simpl. unfold Vnullptr; simpl. + Simpl. simpl. destruct Archi.ppc64; auto. + simpl. discriminate. + simpl. auto. + intros. unfold regmap_get, Regmap.get. rewrite FragBlock.gi; auto. + Simpl. simpl. destruct Archi.ppc64; auto. + Simpl. simpl. destruct Archi.ppc64; auto. unfold Genv.symbol_address. rewrite (match_program_main TRANSF). rewrite symbols_preserved. diff --git a/powerpc/Asmgenproof1.v b/powerpc/Asmgenproof1.v index 460fa6709b..3a4c3e6ef7 100644 --- a/powerpc/Asmgenproof1.v +++ b/powerpc/Asmgenproof1.v @@ -110,6 +110,27 @@ Proof. intros. rewrite Val.add_assoc. rewrite low_high_half. apply add_zero_symbol_address. Qed. +Corollary low_high_half_of_ireg: + forall (r: ireg) rs ge id ofs, + Val.load_result (Preg.chunk_of r) + (Val.add + (Val.load_result (Preg.chunk_of r) (Val.add (gpr_or_zero rs GPR0) (high_half ge id ofs))) + (low_half ge id ofs)) = Genv.symbol_address ge id ofs. +Proof. + intros. + unfold gpr_or_zero; rewrite dec_eq_true. + generalize (low_high_half ge id ofs); intro H. + generalize (low_high_half_zero ge id ofs); intro ADD. + simpl Preg.chunk_of. + destruct Archi.ppc64. + apply low_high_half_zero. + rewrite <- low_high_half. + destruct (high_half ge id ofs) eqn:HH, (low_half ge id ofs) eqn:LH; inv H; simpl; auto. + rewrite Int.add_zero_l; auto. + rewrite Int.add_zero_l; auto. + simpl in *; congruence. +Qed. + (** * Useful properties of registers *) (** [important_preg] extends [data_preg] by tracking the LR register as well. @@ -208,15 +229,25 @@ Proof. Qed. Hint Resolve preg_of_not_LR preg_notin_LR: asmgen. - + +Lemma load_result_compare: + forall c, + Val.load_result Many32 (Val.of_optbool c) = Val.of_optbool c. +Proof. + intros. destruct c; auto. destruct b; auto. +Qed. + (** Useful simplification tactic *) Ltac Simplif := ((rewrite nextinstr_inv by eauto with asmgen) || (rewrite nextinstr_inv2 by eauto with asmgen) - || (rewrite Pregmap.gss) || (rewrite nextinstr_pc) - || (rewrite Pregmap.gso by eauto with asmgen)); auto with asmgen. + || (rewrite Pregmap.gss) + || (rewrite Pregmap.gso by eauto with asmgen) + || (rewrite load_result_compare) + || (rewrite load_result_offset_pc) + ); auto with asmgen. Ltac Simpl := repeat Simplif. @@ -255,11 +286,11 @@ Lemma compare_float_spec: /\ rs1#CR0_2 = Val.cmpf Ceq v1 v2 /\ forall r', r' <> CR0_0 -> r' <> CR0_1 -> r' <> CR0_2 -> r' <> CR0_3 -> r' <> PC -> rs1#r' = rs#r'. Proof. - intros. unfold rs1. - split. reflexivity. - split. reflexivity. - split. reflexivity. - intros. unfold compare_float. Simpl. + intros. unfold rs1, compare_float, Val.cmpf. + split. Simpl. + split. Simpl. + split. Simpl. + intros. Simpl. Qed. Lemma compare_sint_spec: @@ -270,10 +301,10 @@ Lemma compare_sint_spec: /\ rs1#CR0_2 = Val.cmp Ceq v1 v2 /\ forall r', r' <> CR0_0 -> r' <> CR0_1 -> r' <> CR0_2 -> r' <> CR0_3 -> r' <> PC -> rs1#r' = rs#r'. Proof. - intros. unfold rs1. - split. reflexivity. - split. reflexivity. - split. reflexivity. + intros. unfold rs1, compare_sint, Val.cmp. + split. Simpl. + split. Simpl. + split. Simpl. intros. unfold compare_sint. Simpl. Qed. @@ -285,10 +316,10 @@ Lemma compare_uint_spec: /\ rs1#CR0_2 = Val.cmpu (Mem.valid_pointer m) Ceq v1 v2 /\ forall r', r' <> CR0_0 -> r' <> CR0_1 -> r' <> CR0_2 -> r' <> CR0_3 -> r' <> PC -> rs1#r' = rs#r'. Proof. - intros. unfold rs1. - split. reflexivity. - split. reflexivity. - split. reflexivity. + intros. unfold rs1, compare_uint, Val.cmpu. + split. Simpl. + split. Simpl. + split. Simpl. intros. unfold compare_uint. Simpl. Qed. @@ -300,7 +331,7 @@ Lemma compare_slong_spec: /\ rs1#CR0_2 = Val.of_optbool (Val.cmpl_bool Ceq v1 v2) /\ forall r', r' <> CR0_0 -> r' <> CR0_1 -> r' <> CR0_2 -> r' <> CR0_3 -> r' <> PC -> rs1#r' = rs#r'. Proof. - intros. unfold rs1. split. reflexivity. split. reflexivity. split. reflexivity. + intros. unfold rs1, compare_slong. split. Simpl. split. Simpl. split. Simpl. intros. unfold compare_slong. Simpl. Qed. @@ -312,7 +343,7 @@ Lemma compare_ulong_spec: /\ rs1#CR0_2 = Val.of_optbool (Val.cmplu_bool (Mem.valid_pointer m) Ceq v1 v2) /\ forall r', r' <> CR0_0 -> r' <> CR0_1 -> r' <> CR0_2 -> r' <> CR0_3 -> r' <> PC -> rs1#r' = rs#r'. Proof. - intros. unfold rs1. split. reflexivity. split. reflexivity. split. reflexivity. + intros. unfold rs1, compare_ulong. split. Simpl. split. Simpl. split. Simpl. intros. unfold compare_ulong. Simpl. Qed. @@ -328,17 +359,20 @@ Proof. intros. unfold loadimm. case (Int.eq (high_s n) Int.zero). (* addi *) - econstructor; split. apply exec_straight_one. simpl; eauto. auto. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. rewrite Int.add_zero_l. intuition Simpl. + simpl. destruct (Archi.ppc64); auto. (* addis *) generalize (Int.eq_spec (low_s n) Int.zero); case (Int.eq (low_s n) Int.zero); intro. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. rewrite <- H. rewrite Int.add_commut. rewrite low_high_s. intuition Simpl. + simpl. destruct (Archi.ppc64); auto. (* addis + ori *) econstructor; split. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split. Simpl. rewrite Int.add_zero_l. unfold Val.or. rewrite low_high_u. auto. + simpl; eauto. simpl; eauto. Simpl. Simpl. + split. Simpl. rewrite Int.add_zero_l. simpl. + destruct (Archi.ppc64); simpl; rewrite low_high_u; auto. intros; Simpl. Qed. @@ -358,21 +392,34 @@ Proof. case (Int.eq (high_s n) Int.zero). econstructor; split. apply exec_straight_one. simpl. rewrite gpr_or_zero_not_zero; eauto. - reflexivity. + Simpl. intuition Simpl. + simpl. destruct (Archi.ppc64); simpl; auto. + destruct (Pregmap.get r2 rs); simpl; auto. (* addis *) generalize (Int.eq_spec (low_s n) Int.zero); case (Int.eq (low_s n) Int.zero); intro. econstructor; split. apply exec_straight_one. - simpl. rewrite gpr_or_zero_not_zero; auto. auto. + simpl. rewrite gpr_or_zero_not_zero; auto. Simpl. split. Simpl. - generalize (low_high_s n). rewrite H1. rewrite Int.add_zero. congruence. + generalize (low_high_s n). rewrite H1. rewrite Int.add_zero. + intro SHL; rewrite SHL. + change (Preg.chunk_of r1) with (Preg.chunk_of r2). + rewrite Preg.chunk_of_reg_type. apply Val.load_result_same. + destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. intros; Simpl. (* addis + addi *) econstructor; split. eapply exec_straight_two. simpl. rewrite gpr_or_zero_not_zero; eauto. simpl. rewrite gpr_or_zero_not_zero; eauto. - auto. auto. - split. Simpl. rewrite Val.add_assoc. simpl. rewrite low_high_s. auto. + Simpl. Simpl. + split. Simpl. + rewrite Preg.chunk_of_reg_type. + assert (T: forall v1 v2, Val.has_type (Val.add v1 v2) (Preg.type r1)). + { destruct (Preg.type_cases r1) as [T | T]; rewrite T. + destruct v1, v2; simpl; auto. + destruct v1, v2; simpl; auto. } + rewrite !Val.load_result_same by apply T. + rewrite Val.add_assoc. simpl. rewrite low_high_s. auto. intros; Simpl. Qed. @@ -394,8 +441,9 @@ Proof. exists (nextinstr (compare_sint (rs#r1 <- v) v Vzero)). generalize (compare_sint_spec (rs#r1 <- v) v Vzero). intros [A [B [C D]]]. - split. apply exec_straight_one. reflexivity. reflexivity. + split. apply exec_straight_one. reflexivity. unfold compare_sint. Simpl. split. rewrite D; auto with asmgen. Simpl. + subst v. unfold Val.and. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. split. auto. intros. rewrite D; auto with asmgen. Simpl. (* andis *) @@ -406,8 +454,9 @@ Proof. intros [A [B [C D]]]. split. apply exec_straight_one. simpl. generalize (low_high_u n). rewrite H0. rewrite Int.or_zero. - intro. rewrite H1. reflexivity. reflexivity. + intro. rewrite H1. reflexivity. unfold compare_sint. Simpl. split. rewrite D; auto with asmgen. Simpl. + subst v. unfold Val.and. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. split. auto. intros. rewrite D; auto with asmgen. Simpl. (* loadimm + and *) @@ -419,8 +468,9 @@ Proof. split. eapply exec_straight_trans. eexact EX1. apply exec_straight_one. simpl. rewrite RES1. rewrite (OTHER1 r2). reflexivity. congruence. congruence. - reflexivity. + unfold compare_sint. Simpl. split. rewrite D; auto with asmgen. Simpl. + subst v. unfold Val.and. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. split. auto. intros. rewrite D; auto with asmgen. Simpl. Qed. @@ -436,8 +486,9 @@ Proof. intros. unfold andimm. destruct (is_rlw_mask n). (* turned into rlw *) econstructor; split. eapply exec_straight_one. - simpl. rewrite Val.rolm_zero. eauto. auto. + simpl. rewrite Val.rolm_zero. eauto. Simpl. intuition Simpl. + destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. (* andimm_base *) destruct (andimm_base_correct r1 r2 n k rs m) as [rs' [A [B [C D]]]]; auto. exists rs'; auto. @@ -457,20 +508,29 @@ Proof. case (Int.eq (high_u n) Int.zero). (* ori *) exists (nextinstr (rs#r1 <- v)). - split. apply exec_straight_one. reflexivity. reflexivity. + split. apply exec_straight_one. reflexivity. Simpl. intuition Simpl. + subst v. unfold Val.or. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. (* oris *) generalize (Int.eq_spec (low_u n) Int.zero); case (Int.eq (low_u n) Int.zero); intro. exists (nextinstr (rs#r1 <- v)). split. apply exec_straight_one. simpl. generalize (low_high_u n). rewrite H. rewrite Int.or_zero. - intro. rewrite H0. reflexivity. reflexivity. + intro. rewrite H0. reflexivity. Simpl. intuition Simpl. + subst v. unfold Val.or. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. (* oris + ori *) - econstructor; split. eapply exec_straight_two; simpl; reflexivity. + econstructor; split. eapply exec_straight_two; simpl; try reflexivity. + Simpl. Simpl. intuition Simpl. - rewrite Val.or_assoc. simpl. rewrite low_high_u. reflexivity. + rewrite Preg.chunk_of_reg_type. + assert (T: forall v1 v2, Val.has_type (Val.or v1 v2) (Preg.type r1)). + { destruct (Preg.type_cases r1) as [T | T]; rewrite T. + destruct v1, v2; simpl; auto. + destruct v1, v2; simpl; auto. } + rewrite !Val.load_result_same by apply T. + rewrite Val.or_assoc. simpl. rewrite low_high_u. auto. Qed. (** Xor integer immediate. *) @@ -487,19 +547,28 @@ Proof. case (Int.eq (high_u n) Int.zero). (* xori *) exists (nextinstr (rs#r1 <- v)). - split. apply exec_straight_one. reflexivity. reflexivity. + split. apply exec_straight_one. reflexivity. Simpl. intuition Simpl. + subst v. unfold Val.xor. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. (* xoris *) generalize (Int.eq_spec (low_u n) Int.zero); case (Int.eq (low_u n) Int.zero); intro. exists (nextinstr (rs#r1 <- v)). split. apply exec_straight_one. simpl. generalize (low_high_u n). rewrite H. rewrite Int.or_zero. - intro. rewrite H0. reflexivity. reflexivity. + intro. rewrite H0. reflexivity. Simpl. intuition Simpl. + subst v. unfold Val.xor. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. (* xoris + xori *) - econstructor; split. eapply exec_straight_two; simpl; reflexivity. + econstructor; split. eapply exec_straight_two; simpl; try reflexivity. + Simpl. Simpl. intuition Simpl. + rewrite Preg.chunk_of_reg_type. + assert (T: forall v1 v2, Val.has_type (Val.xor v1 v2) (Preg.type r1)). + { destruct (Preg.type_cases r1) as [T | T]; rewrite T. + destruct v1, v2; simpl; auto. + destruct v1, v2; simpl; auto. } + rewrite !Val.load_result_same by apply T. rewrite Val.xor_assoc. simpl. rewrite low_high_u_xor. reflexivity. Qed. @@ -517,15 +586,27 @@ Proof. (* rlwinm *) econstructor; split. eapply exec_straight_one; simpl; eauto. intuition Simpl. + intuition Simpl. + unfold Val.rolm. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. (* rlwinm ; andimm *) set (rs1 := nextinstr (rs#r1 <- (Val.rolm rs#r2 amount Int.mone))). destruct (andimm_base_correct r1 r1 mask k rs1 m) as [rs' [A [B [C D]]]]; auto. exists rs'. - split. eapply exec_straight_step; eauto. auto. auto. + split. eapply exec_straight_step; eauto. auto. + subst rs1. Simpl. split. rewrite B. unfold rs1. rewrite nextinstr_inv; auto with asmgen. - rewrite Pregmap.gss. destruct (rs r2); simpl; auto. + rewrite Pregmap.gss. + unfold Val.rolm. destruct (Pregmap.get r2 rs); simpl; destruct Archi.ppc64; auto. + change Many64 with (chunk_of_type Tany64). + rewrite Val.load_result_same. simpl. + unfold Int.rolm. rewrite Int.and_assoc. + decEq; decEq; decEq. rewrite Int.and_commut. apply Int.and_mone. + simpl; auto. + change Many32 with (chunk_of_type Tany32). + rewrite Val.load_result_same. simpl. unfold Int.rolm. rewrite Int.and_assoc. decEq; decEq; decEq. rewrite Int.and_commut. apply Int.and_mone. + simpl; auto. intros. rewrite D; auto. unfold rs1; Simpl. Qed. @@ -535,7 +616,7 @@ Lemma loadimm64_correct: forall r n k rs m, exists rs', exec_straight ge fn (loadimm64 r n k) rs m k rs' m - /\ rs'#r = Vlong n + /\ rs'#r = Val.load_result (Preg.chunk_of r) (Vlong n) /\ forall r': preg, r' <> r -> r' <> GPR12 -> r' <> PC -> rs'#r' = rs#r'. Proof. intros. unfold loadimm64. @@ -543,15 +624,17 @@ Proof. set (hi' := Int64.shl hi_s (Int64.repr 16)). destruct (Int64.eq n (low64_s n)). (* addi *) - econstructor; split. apply exec_straight_one. simpl; eauto. auto. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. rewrite Int64.add_zero_l. intuition Simpl. + (*simpl. rewrite H; simpl; auto.*) (* addis + ori *) predSpec Int64.eq Int64.eq_spec n (Int64.or hi' (low64_u n)). - econstructor; split. eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split. Simpl. rewrite Int64.add_zero_l. simpl; f_equal; auto. + econstructor; split. eapply exec_straight_two. simpl; eauto. simpl; eauto. Simpl. Simpl. + split. Simpl. rewrite Int64.add_zero_l. + simpl. destruct Archi.ppc64. rewrite H at 2; f_equal. simpl; auto. intros. Simpl. (* ldi *) - econstructor; split. apply exec_straight_one. simpl; eauto. auto. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. intuition Simpl. Qed. @@ -559,6 +642,7 @@ Qed. Lemma addimm64_correct: forall r1 r2 n k rs m, + Archi.ppc64 = true -> r2 <> GPR0 -> exists rs', exec_straight ge fn (addimm64 r1 r2 n k) rs m k rs' m @@ -569,19 +653,22 @@ Proof. - (* addi *) econstructor; split. apply exec_straight_one. simpl. rewrite gpr_or_zero_l_not_zero; eauto. - reflexivity. + Simpl. intuition Simpl. + simpl. rewrite H; simpl; auto. - (* move-loadimm-add *) subst r2. - edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_step. simpl; reflexivity. Simpl. + eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. split. rewrite B, C by eauto with asmgen. Simpl. + simpl. rewrite H; simpl; auto. intros. Simpl. rewrite C by auto. Simpl. - (* loadimm-add *) - edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. split. rewrite B, C by eauto with asmgen. Simpl. + simpl. rewrite H; simpl; auto. intros. Simpl. Qed. @@ -589,6 +676,7 @@ Qed. Lemma orimm64_correct: forall r1 r2 n k rs m, + Archi.ppc64 = true -> r2 <> GPR0 -> exists rs', exec_straight ge fn (orimm64 r1 r2 n k) rs m k rs' m @@ -597,19 +685,22 @@ Lemma orimm64_correct: Proof. intros. unfold orimm64, opimm64. destruct (Int64.eq n (low64_u n)); [|destruct (ireg_eq r2 GPR12)]. - (* ori *) - econstructor; split. apply exec_straight_one. simpl; eauto. reflexivity. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. intuition Simpl. + simpl. rewrite H; simpl; auto. - (* move-loadimm-or *) subst r2. - edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_step. simpl; reflexivity. Simpl. + eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. split. rewrite B, C by eauto with asmgen. Simpl. + simpl. rewrite H; simpl; auto. intros. Simpl. rewrite C by auto. Simpl. - (* loadimm-or *) - edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. split. rewrite B, C by eauto with asmgen. Simpl. + simpl. rewrite H; simpl; auto. intros. Simpl. Qed. @@ -617,6 +708,7 @@ Qed. Lemma xorimm64_correct: forall r1 r2 n k rs m, + Archi.ppc64 = true -> r2 <> GPR0 -> exists rs', exec_straight ge fn (xorimm64 r1 r2 n k) rs m k rs' m @@ -625,19 +717,22 @@ Lemma xorimm64_correct: Proof. intros. unfold xorimm64, opimm64. destruct (Int64.eq n (low64_u n)); [|destruct (ireg_eq r2 GPR12)]. - (* xori *) - econstructor; split. apply exec_straight_one. simpl; eauto. reflexivity. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. intuition Simpl. + simpl. rewrite H; simpl; auto. - (* move-loadimm-xor *) subst r2. - edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_step. simpl; reflexivity. Simpl. + eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. split. rewrite B, C by eauto with asmgen. Simpl. + simpl. rewrite H; simpl; auto. intros. Simpl. rewrite C by auto. Simpl. - (* loadimm-xor *) - edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. split. rewrite B, C by eauto with asmgen. Simpl. + simpl. rewrite H; simpl; auto. intros. Simpl. Qed. @@ -645,6 +740,7 @@ Qed. Lemma andimm64_base_correct: forall r1 r2 n k rs m, + Archi.ppc64 = true -> r2 <> GPR0 -> exists rs', exec_straight ge fn (andimm64_base r1 r2 n k) rs m k rs' m @@ -653,24 +749,31 @@ Lemma andimm64_base_correct: Proof. intros. unfold andimm64_base, opimm64. destruct (Int64.eq n (low64_u n)); [|destruct (ireg_eq r2 GPR12)]. - (* andi *) - econstructor; split. apply exec_straight_one. simpl; eauto. reflexivity. + econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. + unfold compare_slong; intuition Simpl. unfold compare_slong; intuition Simpl. + simpl. rewrite H; simpl; auto. - (* move-loadimm-and *) subst r2. - edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR12 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_step. simpl; reflexivity. Simpl. + eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. + unfold compare_slong; intuition Simpl. split. rewrite B, C by eauto with asmgen. unfold compare_slong; Simpl. + simpl. rewrite H; simpl; auto. intros. unfold compare_slong; Simpl. rewrite C by auto with asmgen. Simpl. - (* loadimm-xor *) - edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). - econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. auto. + edestruct (loadimm64_correct GPR0 n) as (rs2 & A & B & C). auto. + econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; reflexivity. Simpl. + unfold compare_slong; Simpl. split. rewrite B, C by eauto with asmgen. unfold compare_slong; Simpl. + simpl. rewrite H; simpl; auto. intros. unfold compare_slong; Simpl. Qed. Lemma andimm64_correct: forall r1 r2 n k rs m, + Archi.ppc64 = true -> r2 <> GPR0 -> exists rs', exec_straight ge fn (andimm64 r1 r2 n k) rs m k rs' m @@ -678,8 +781,8 @@ Lemma andimm64_correct: /\ forall r': preg, r' <> r1 -> r' <> GPR12 -> important_preg r' = true -> rs'#r' = rs#r'. Proof. intros. unfold andimm64. destruct (is_rldl_mask n || is_rldr_mask n). -- econstructor; split. eapply exec_straight_one. simpl; reflexivity. reflexivity. - split. Simpl. destruct (rs r2); simpl; auto. rewrite Int64.rolm_zero. auto. +- econstructor; split. eapply exec_straight_one. simpl; reflexivity. Simpl. + split. Simpl. destruct (rs # r2); simpl; rewrite H; auto. rewrite Int64.rolm_zero. auto. intros; Simpl. - apply andimm64_base_correct; auto. Qed. @@ -688,6 +791,7 @@ Qed. Lemma rolm64_correct: forall r1 r2 amount mask k rs m, + Archi.ppc64 = true -> r1 <> GPR0 -> exists rs', exec_straight ge fn (rolm64 r1 r2 amount mask k) rs m k rs' m @@ -696,13 +800,17 @@ Lemma rolm64_correct: Proof. intros. unfold rolm64. destruct (is_rldl_mask mask || is_rldr_mask mask || is_rldl_mask (Int64.shru' mask amount)). -- econstructor; split. eapply exec_straight_one. simpl; reflexivity. reflexivity. +- econstructor; split. eapply exec_straight_one. simpl; reflexivity. Simpl. intuition Simpl. + change (Preg.chunk_of r1) with (Preg.chunk_of r2). + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + destruct (Pregmap.get r2 rs); simpl; try rewrite H; auto. - edestruct (andimm64_base_correct r1 r1 mask k) as (rs2 & A & B & C); auto. econstructor; split. - eapply exec_straight_step. simpl; reflexivity. reflexivity. eexact A. - split. rewrite B. Simpl. destruct (rs r2); simpl; auto. unfold Int64.rolm. - rewrite Int64.and_assoc, Int64.and_mone_l; auto. + eapply exec_straight_step. simpl; reflexivity. Simpl. eexact A. + split. rewrite B. Simpl. destruct (rs # r2); simpl; rewrite H; auto. unfold Int64.rolm. + change Many64 with (chunk_of_type Tany64). rewrite Val.load_result_same by (simpl; auto). + unfold Val.andl. rewrite Int64.and_assoc, Int64.and_mone_l; auto. intros; Simpl. rewrite C by auto. Simpl. Qed. @@ -719,46 +827,55 @@ Lemma accessind_load_correct: exec_instr ge fn (instr2 r1 r2 r3) rs m = load2 chunk (inj r1) r2 r3 rs m) -> Mem.loadv chunk m (Val.offset_ptr rs#base ofs) = Some v -> base <> GPR0 -> inj rx <> PC -> + Val.has_type v (Preg.type (inj rx)) -> exists rs', exec_straight ge fn (accessind instr1 instr2 base ofs rx k) rs m k rs' m /\ rs'#(inj rx) = v /\ forall r, r <> PC -> r <> inj rx -> r <> GPR0 -> rs'#r = rs#r. Proof. intros. unfold accessind. set (ofs' := Ptrofs.to_int ofs) in *. - assert (LD: Mem.loadv chunk m (Val.add (rs base) (Vint ofs')) = Some v) + assert (LD: Mem.loadv chunk m (Val.add (rs # base) (Vint ofs')) = Some v) by (apply loadv_offset_ptr; auto). destruct (Int.eq (high_s ofs') Int.zero). - econstructor; split. apply exec_straight_one. rewrite H. unfold load1. rewrite gpr_or_zero_not_zero by auto. simpl. rewrite LD. eauto. unfold nextinstr. repeat Simplif. split. unfold nextinstr. repeat Simplif. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. intros. repeat Simplif. - exploit (loadimm_correct GPR0 ofs'); eauto. intros [rs' [P [Q R]]]. econstructor; split. eapply exec_straight_trans. eexact P. apply exec_straight_one. rewrite H0. unfold load2. rewrite Q, R by auto with asmgen. rewrite LD. reflexivity. unfold nextinstr. repeat Simplif. split. repeat Simplif. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. intros. repeat Simplif. Qed. Lemma loadind_correct: - forall (base: ireg) ofs ty dst k (rs: regset) m v c, - loadind base ofs ty dst k = OK c -> - Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v -> + forall (base: ireg) ofs q dst k (rs: regset) m v c, + loadind base ofs q dst k = OK c -> + Mem.loadv (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) = Some v -> base <> GPR0 -> exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of dst) = v /\ forall r, r <> PC -> r <> preg_of dst -> r <> GPR0 -> rs'#r = rs#r. Proof. - unfold loadind; intros. destruct ty; try discriminate; destruct (preg_of dst); inv H; simpl in H0. - apply accessind_load_correct with (inj := IR) (chunk := Mint32); auto with asmgen. - apply accessind_load_correct with (inj := FR) (chunk := Mfloat64); auto with asmgen. - apply accessind_load_correct with (inj := IR) (chunk := Mint64); auto with asmgen. - apply accessind_load_correct with (inj := FR) (chunk := Mfloat32); auto with asmgen. - apply accessind_load_correct with (inj := IR) (chunk := Many32); auto with asmgen. - apply accessind_load_correct with (inj := IR) (chunk := Many64); auto with asmgen. - apply accessind_load_correct with (inj := FR) (chunk := Many64); auto with asmgen. + unfold loadind; intros. + destruct q, Archi.ppc64 eqn:PPC64; destruct (preg_of dst); inv H; simpl in H0. +- apply accessind_load_correct with (inj := FR) (chunk := Many32); auto with asmgen. + simpl; apply Val.has_type_Tany64. +- apply accessind_load_correct with (inj := IR) (chunk := Many32); auto with asmgen. + simpl; rewrite PPC64; apply Mem.loadv_type in H0; auto. +- apply accessind_load_correct with (inj := FR) (chunk := Many32); auto with asmgen. + simpl; apply Val.has_type_Tany64. +- apply accessind_load_correct with (inj := IR) (chunk := Many64); auto with asmgen. + simpl; rewrite PPC64; apply Val.has_type_Tany64. +- apply accessind_load_correct with (inj := FR) (chunk := Many64); auto with asmgen. + simpl; apply Val.has_type_Tany64. +- apply accessind_load_correct with (inj := FR) (chunk := Many64); auto with asmgen. + simpl; apply Val.has_type_Tany64. Qed. (** Indexed memory stores. *) @@ -772,14 +889,14 @@ Lemma accessind_store_correct: exec_instr ge fn (instr1 r1 cst r2) rs m = store1 ge chunk (inj r1) cst r2 rs m) -> (forall rs m r1 r2 r3, exec_instr ge fn (instr2 r1 r2 r3) rs m = store2 chunk (inj r1) r2 r3 rs m) -> - Mem.storev chunk m (Val.offset_ptr rs#base ofs) (rs (inj rx)) = Some m' -> + Mem.storev chunk m (Val.offset_ptr rs#base ofs) (rs # (inj rx)) = Some m' -> base <> GPR0 -> inj rx <> PC -> inj rx <> GPR0 -> exists rs', exec_straight ge fn (accessind instr1 instr2 base ofs rx k) rs m k rs' m' /\ forall r, r <> PC -> r <> GPR0 -> rs'#r = rs#r. Proof. intros. unfold accessind. set (ofs' := Ptrofs.to_int ofs) in *. - assert (ST: Mem.storev chunk m (Val.add (rs base) (Vint ofs')) (rs (inj rx)) = Some m') + assert (ST: Mem.storev chunk m (Val.add (rs # base) (Vint ofs')) (rs # (inj rx)) = Some m') by (apply storev_offset_ptr; auto). destruct (Int.eq (high_s ofs') Int.zero). - econstructor; split. apply exec_straight_one. @@ -795,9 +912,9 @@ Proof. Qed. Lemma storeind_correct: - forall (base: ireg) ofs ty src k (rs: regset) m m' c, - storeind src base ofs ty k = OK c -> - Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' -> + forall (base: ireg) ofs q src k (rs: regset) m m' c, + storeind src base ofs q k = OK c -> + Mem.storev (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' -> base <> GPR0 -> exists rs', exec_straight ge fn c rs m k rs' m' @@ -805,14 +922,13 @@ Lemma storeind_correct: Proof. unfold storeind; intros. assert (preg_of src <> GPR0) by auto with asmgen. - destruct ty; try discriminate; destruct (preg_of src) ; inv H; simpl in H0. - apply accessind_store_correct with (inj := IR) (chunk := Mint32); auto with asmgen. - apply accessind_store_correct with (inj := FR) (chunk := Mfloat64); auto with asmgen. - apply accessind_store_correct with (inj := IR) (chunk := Mint64); auto with asmgen. - apply accessind_store_correct with (inj := FR) (chunk := Mfloat32); auto with asmgen. + destruct q, Archi.ppc64; try discriminate; destruct (preg_of src) ; inv H; simpl in H0. + apply accessind_store_correct with (inj := FR) (chunk := Many32); auto with asmgen. apply accessind_store_correct with (inj := IR) (chunk := Many32); auto with asmgen. + apply accessind_store_correct with (inj := FR) (chunk := Many32); auto with asmgen. apply accessind_store_correct with (inj := IR) (chunk := Many64); auto with asmgen. apply accessind_store_correct with (inj := FR) (chunk := Many64); auto with asmgen. + apply accessind_store_correct with (inj := FR) (chunk := Many64); auto with asmgen. Qed. (** Float comparisons. *) @@ -838,8 +954,10 @@ Proof. case cmp; tauto. unfold floatcomp. elim H; intro; clear H. exists rs1. - split. destruct H0 as [EQ|[EQ|[EQ|EQ]]]; subst cmp; - apply exec_straight_one; reflexivity. + split. + destruct H0 as [EQ|[EQ|[EQ|EQ]]]; subst cmp; + apply exec_straight_one; try reflexivity; + subst rs1; unfold compare_float; Simpl. split. destruct H0 as [EQ|[EQ|[EQ|EQ]]]; subst cmp; simpl; auto. rewrite Val.negate_cmpf_eq. auto. @@ -850,19 +968,56 @@ Proof. apply exec_straight_two with rs1 m. reflexivity. simpl. rewrite C; rewrite A. rewrite Val.or_commut. rewrite <- Val.cmpf_le. - reflexivity. reflexivity. reflexivity. + reflexivity. subst rs1; unfold compare_float; Simpl. subst rs1; unfold compare_float; Simpl. apply exec_straight_two with rs1 m. reflexivity. simpl. rewrite C; rewrite B. rewrite Val.or_commut. rewrite <- Val.cmpf_ge. - reflexivity. reflexivity. reflexivity. + reflexivity. subst rs1; unfold compare_float; Simpl. subst rs1; unfold compare_float; Simpl. split. elim H0; intro; subst cmp; simpl. - reflexivity. - reflexivity. + subst rs1; unfold compare_float, Val.cmpf; Simpl. + subst rs1; unfold compare_float, Val.cmpf; Simpl. intros. Simpl. Qed. (** Translation of conditions. *) +Lemma load_result_long_32: + forall (i: ireg) l, + Archi.ppc64 = false -> + Val.load_result (Preg.chunk_of i) (Vlong l) = Vundef. +Proof. + intros. simpl. rewrite H; reflexivity. +Qed. + +Lemma load_result_long_64: + forall (i: ireg) v, + Archi.ppc64 = true -> + Val.load_result (Preg.chunk_of i) v = v. +Proof. + intros. rewrite Preg.chunk_of_reg_type. simpl. rewrite H; reflexivity. +Qed. + +Lemma compare_long_32: + forall c (i: ireg) rs v, + Archi.ppc64 = false -> + Val.cmpl_bool c (Pregmap.get i rs) v = None. +Proof. + intros. generalize (Pregmap.get_has_type i rs); intro T. + destruct (Pregmap.get i rs); auto. + simpl in T. rewrite H in T. contradiction. +Qed. + +Lemma compare_ulong_32: + forall p c (i: ireg) rs v, + Archi.ppc64 = false -> + Val.cmplu_bool p c (Pregmap.get i rs) v = None. +Proof. + intros. generalize (Pregmap.get_has_type i rs); intro T. + destruct (Pregmap.get i rs); auto. + simpl in T. rewrite H in T. contradiction. + unfold Val.cmplu_bool. simpl. destruct v; auto. +Qed. + Ltac ArgsInv := repeat (match goal with | [ H: Error _ = OK _ |- _ ] => discriminate @@ -884,76 +1039,76 @@ Lemma transl_cond_correct_1: exec_straight ge fn c rs m k rs' m /\ rs'#(reg_of_crbit (fst (crbit_for_cond cond))) = (if snd (crbit_for_cond cond) - then Val.of_optbool (eval_condition cond (map rs (map preg_of args)) m) - else Val.notbool (Val.of_optbool (eval_condition cond (map rs (map preg_of args)) m))) + then Val.of_optbool (eval_condition cond (map (fun r => rs#r) (map preg_of args)) m) + else Val.notbool (Val.of_optbool (eval_condition cond (map (fun r => rs#r) (map preg_of args)) m))) /\ forall r, important_preg r = true -> preg_notin r (destroyed_by_cond cond) -> rs'#r = rs#r. Proof. intros. Opaque Int.eq. unfold transl_cond in H; destruct cond; ArgsInv; simpl. (* Ccomp *) - fold (Val.cmp c0 (rs x) (rs x0)). - destruct (compare_sint_spec rs (rs x) (rs x0)) as [A [B [C D]]]. + fold (Val.cmp c0 (rs # x) (rs # x0)). + destruct (compare_sint_spec rs (rs # x) (rs # x0)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_sint; Simpl. split. case c0; simpl; auto; rewrite <- Val.negate_cmp; simpl; auto. auto with asmgen. (* Ccompu *) - fold (Val.cmpu (Mem.valid_pointer m) c0 (rs x) (rs x0)). - destruct (compare_uint_spec rs m (rs x) (rs x0)) as [A [B [C D]]]. + fold (Val.cmpu (Mem.valid_pointer m) c0 (rs # x) (rs # x0)). + destruct (compare_uint_spec rs m (rs # x) (rs # x0)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_uint; Simpl. split. case c0; simpl; auto; rewrite <- Val.negate_cmpu; simpl; auto. auto with asmgen. (* Ccompimm *) - fold (Val.cmp c0 (rs x) (Vint i)). + fold (Val.cmp c0 (rs # x) (Vint i)). destruct (Int.eq (high_s i) Int.zero); inv EQ0. - destruct (compare_sint_spec rs (rs x) (Vint i)) as [A [B [C D]]]. + destruct (compare_sint_spec rs (rs # x) (Vint i)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_sint; Simpl. split. case c0; simpl; auto; rewrite <- Val.negate_cmp; simpl; auto. auto with asmgen. destruct (loadimm_correct GPR0 i (Pcmpw x GPR0 :: k) rs m) as [rs1 [EX1 [RES1 OTH1]]]. - destruct (compare_sint_spec rs1 (rs x) (Vint i)) as [A [B [C D]]]. - assert (SAME: rs1 x = rs x) by (apply OTH1; eauto with asmgen). - exists (nextinstr (compare_sint rs1 (rs1 x) (Vint i))). + destruct (compare_sint_spec rs1 (rs # x) (Vint i)) as [A [B [C D]]]. + assert (SAME: rs1 # x = rs # x) by (apply OTH1; eauto with asmgen). + exists (nextinstr (compare_sint rs1 (rs1 # x) (Vint i))). split. eapply exec_straight_trans. eexact EX1. apply exec_straight_one. simpl. rewrite RES1; rewrite SAME; auto. - reflexivity. + unfold compare_sint; Simpl. split. rewrite SAME. case c0; simpl; auto; rewrite <- Val.negate_cmp; simpl; auto. intros. rewrite SAME; rewrite D; auto with asmgen. (* Ccompuimm *) - fold (Val.cmpu (Mem.valid_pointer m) c0 (rs x) (Vint i)). + fold (Val.cmpu (Mem.valid_pointer m) c0 (rs # x) (Vint i)). destruct (Int.eq (high_u i) Int.zero); inv EQ0. - destruct (compare_uint_spec rs m (rs x) (Vint i)) as [A [B [C D]]]. + destruct (compare_uint_spec rs m (rs # x) (Vint i)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_uint; Simpl. split. case c0; simpl; auto; rewrite <- Val.negate_cmpu; simpl; auto. auto with asmgen. destruct (loadimm_correct GPR0 i (Pcmplw x GPR0 :: k) rs m) as [rs1 [EX1 [RES1 OTH1]]]. - destruct (compare_uint_spec rs1 m (rs x) (Vint i)) as [A [B [C D]]]. - assert (SAME: rs1 x = rs x) by (apply OTH1; eauto with asmgen). - exists (nextinstr (compare_uint rs1 m (rs1 x) (Vint i))). + destruct (compare_uint_spec rs1 m (rs # x) (Vint i)) as [A [B [C D]]]. + assert (SAME: rs1 # x = rs # x) by (apply OTH1; eauto with asmgen). + exists (nextinstr (compare_uint rs1 m (rs1 # x) (Vint i))). split. eapply exec_straight_trans. eexact EX1. apply exec_straight_one. simpl. rewrite RES1; rewrite SAME; auto. - reflexivity. + unfold compare_uint; Simpl. split. rewrite SAME. case c0; simpl; auto; rewrite <- Val.negate_cmpu; simpl; auto. intros. rewrite SAME; rewrite D; auto with asmgen. (* Ccompf *) - fold (Val.cmpf c0 (rs x) (rs x0)). + fold (Val.cmpf c0 (rs # x) (rs # x0)). destruct (floatcomp_correct c0 x x0 k rs m) as [rs' [EX [RES OTH]]]. exists rs'. split. auto. split. apply RES. auto with asmgen. (* Cnotcompf *) rewrite Val.notbool_negb_3. rewrite Val.notbool_idem4. - fold (Val.cmpf c0 (rs x) (rs x0)). + fold (Val.cmpf c0 (rs # x) (rs # x0)). destruct (floatcomp_correct c0 x x0 k rs m) as [rs' [EX [RES OTH]]]. exists rs'. split. auto. split. rewrite RES. destruct (snd (crbit_for_fcmp c0)); auto. @@ -962,80 +1117,126 @@ Opaque Int.eq. destruct (andimm_base_correct GPR0 x i k rs m) as [rs' [A [B [C D]]]]. eauto with asmgen. exists rs'. split. assumption. - split. rewrite C. destruct (rs x); auto. + split. rewrite C. destruct (rs # x); auto. auto with asmgen. (* Cmasknotzero *) destruct (andimm_base_correct GPR0 x i k rs m) as [rs' [A [B [C D]]]]. eauto with asmgen. exists rs'. split. assumption. - split. rewrite C. destruct (rs x); auto. + split. rewrite C. destruct (rs # x); auto. fold (option_map negb (Some (Int.eq (Int.and i0 i) Int.zero))). rewrite Val.notbool_negb_3. rewrite Val.notbool_idem4. auto. auto with asmgen. - (* Ccompl *) - destruct (compare_slong_spec rs (rs x) (rs x0)) as [A [B [C D]]]. + destruct (compare_slong_spec rs (rs # x) (rs # x0)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_slong; Simpl. rewrite <- Val.notbool_negb_3. rewrite <- Val.negate_cmpl_bool. split. case c0; simpl; auto. auto with asmgen. - (* Ccomplu *) - destruct (compare_ulong_spec rs m (rs x) (rs x0)) as [A [B [C D]]]. + destruct (compare_ulong_spec rs m (rs # x) (rs # x0)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_ulong; Simpl. rewrite <- Val.notbool_negb_3. rewrite <- Val.negate_cmplu_bool. split. case c0; simpl; auto. auto with asmgen. - (* Ccomplimm *) rewrite <- Val.notbool_negb_3. rewrite <- Val.negate_cmpl_bool. destruct (Int64.eq i (low64_s i)); [|destruct (ireg_eq x GPR12)]; inv EQ0. -+ destruct (compare_slong_spec rs (rs x) (Vlong i)) as [A [B [C D]]]. ++ destruct (compare_slong_spec rs (rs # x) (Vlong i)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_slong; Simpl. split. case c0; simpl; auto. auto with asmgen. + destruct (loadimm64_correct GPR12 i (Pcmpd GPR0 GPR12 :: k) (nextinstr (rs#GPR0 <- (rs#GPR12))) m) as [rs1 [EX1 [RES1 OTH1]]]. - destruct (compare_slong_spec rs1 (rs GPR12) (Vlong i)) as [A [B [C D]]]. - assert (SAME: rs1 GPR0 = rs GPR12) by (apply OTH1; eauto with asmgen). + destruct (compare_slong_spec rs1 (rs # GPR12) (Vlong i)) as [A [B [C D]]]. + assert (SAME: rs1 # GPR0 = rs # GPR12). + { rewrite OTH1; auto with asmgen. Simpl. + change (Preg.chunk_of GPR0) with (Preg.chunk_of GPR12). + rewrite Preg.chunk_of_reg_type. apply Val.load_result_same, Pregmap.get_has_type. } econstructor; split. - eapply exec_straight_step. simpl;reflexivity. reflexivity. - eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. reflexivity. - split. rewrite RES1, SAME. destruct c0; simpl; auto. - simpl; intros. rewrite RES1, SAME. rewrite D by eauto with asmgen. rewrite OTH1 by eauto with asmgen. Simpl. + eapply exec_straight_step. simpl;reflexivity. Simpl. + eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. + unfold compare_slong; Simpl. + split; intros; rewrite RES1, SAME. + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. destruct c0; simpl; auto. } + { rewrite load_result_long_32 by auto. + unfold compare_slong in *. rewrite !compare_long_32 in * by auto. + destruct c0; simpl; repeat Simplif; destruct (Pregmap.get GPR12 rs); auto. } + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. rewrite D, OTH1 by eauto with asmgen. Simpl. } + { rewrite load_result_long_32 by auto. + unfold compare_slong in *. rewrite !compare_long_32 in * by auto. + rewrite D, OTH1 by eauto with asmgen. Simpl. } + destruct (loadimm64_correct GPR0 i (Pcmpd x GPR0 :: k) rs m) as [rs1 [EX1 [RES1 OTH1]]]. - destruct (compare_slong_spec rs1 (rs x) (Vlong i)) as [A [B [C D]]]. - assert (SAME: rs1 x = rs x) by (apply OTH1; eauto with asmgen). + destruct (compare_slong_spec rs1 (rs # x) (Vlong i)) as [A [B [C D]]]. + assert (SAME: rs1 # x = rs # x) by (apply OTH1; eauto with asmgen). econstructor; split. - eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. reflexivity. - split. rewrite RES1, SAME. destruct c0; simpl; auto. - simpl; intros. rewrite RES1, SAME. rewrite D; eauto with asmgen. + eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. + unfold compare_slong; Simpl. + split; intros; rewrite RES1, SAME. + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. destruct c0; simpl; auto. } + { rewrite load_result_long_32 by auto. + unfold compare_slong in *. rewrite !compare_long_32 in * by auto. + destruct c0; simpl; repeat Simplif; destruct (Pregmap.get GPR12 rs); auto. } + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. rewrite D, OTH1 by eauto with asmgen. reflexivity. } + { rewrite load_result_long_32 by auto. + unfold compare_slong in *. rewrite !compare_long_32 in * by auto. + rewrite D, OTH1 by eauto with asmgen. reflexivity. } - (* Ccompluimm *) rewrite <- Val.notbool_negb_3. rewrite <- Val.negate_cmplu_bool. destruct (Int64.eq i (low64_u i)); [|destruct (ireg_eq x GPR12)]; inv EQ0. -+ destruct (compare_ulong_spec rs m (rs x) (Vlong i)) as [A [B [C D]]]. ++ destruct (compare_ulong_spec rs m (rs # x) (Vlong i)) as [A [B [C D]]]. econstructor; split. - apply exec_straight_one. simpl; reflexivity. reflexivity. + apply exec_straight_one. simpl; reflexivity. unfold compare_ulong; Simpl. split. case c0; simpl; auto. auto with asmgen. + destruct (loadimm64_correct GPR12 i (Pcmpld GPR0 GPR12 :: k) (nextinstr (rs#GPR0 <- (rs#GPR12))) m) as [rs1 [EX1 [RES1 OTH1]]]. - destruct (compare_ulong_spec rs1 m (rs GPR12) (Vlong i)) as [A [B [C D]]]. - assert (SAME: rs1 GPR0 = rs GPR12) by (apply OTH1; eauto with asmgen). + destruct (compare_ulong_spec rs1 m (rs # GPR12) (Vlong i)) as [A [B [C D]]]. + assert (SAME: rs1 # GPR0 = rs # GPR12). + { rewrite OTH1; auto with asmgen. Simpl. + change (Preg.chunk_of GPR0) with (Preg.chunk_of GPR12). + rewrite Preg.chunk_of_reg_type. apply Val.load_result_same, Pregmap.get_has_type. } econstructor; split. - eapply exec_straight_step. simpl;reflexivity. reflexivity. - eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. reflexivity. - split. rewrite RES1, SAME. destruct c0; simpl; auto. - simpl; intros. rewrite RES1, SAME. rewrite D by eauto with asmgen. rewrite OTH1 by eauto with asmgen. Simpl. + eapply exec_straight_step. simpl;reflexivity. Simpl. + eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. + unfold compare_ulong; Simpl. + split; intros; rewrite RES1, SAME. + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. destruct c0; simpl; auto. } + { rewrite load_result_long_32 by auto. + unfold compare_ulong in *. rewrite !compare_ulong_32 in * by auto. + destruct c0; simpl; repeat Simplif; destruct (Pregmap.get GPR12 rs); auto. } + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. rewrite D, OTH1 by eauto with asmgen. Simpl. } + { rewrite load_result_long_32 by auto. + unfold compare_ulong in *. rewrite !compare_ulong_32 in * by auto. + rewrite D, OTH1 by eauto with asmgen. Simpl. } + destruct (loadimm64_correct GPR0 i (Pcmpld x GPR0 :: k) rs m) as [rs1 [EX1 [RES1 OTH1]]]. - destruct (compare_ulong_spec rs1 m (rs x) (Vlong i)) as [A [B [C D]]]. - assert (SAME: rs1 x = rs x) by (apply OTH1; eauto with asmgen). + destruct (compare_ulong_spec rs1 m (rs # x) (Vlong i)) as [A [B [C D]]]. + assert (SAME: rs1 # x = rs # x) by (apply OTH1; eauto with asmgen). econstructor; split. - eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. reflexivity. - split. rewrite RES1, SAME. destruct c0; simpl; auto. - simpl; intros. rewrite RES1, SAME. rewrite D; eauto with asmgen. + eapply exec_straight_trans. eexact EX1. eapply exec_straight_one. simpl;reflexivity. + unfold compare_ulong; Simpl. + split; intros; rewrite RES1, SAME. + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. destruct c0; simpl; auto. } + { rewrite load_result_long_32 by auto. + unfold compare_ulong in *. rewrite !compare_ulong_32 in * by auto. + destruct c0; simpl; repeat Simplif; destruct (Pregmap.get GPR12 rs); auto. } + destruct Archi.ppc64 eqn:PPC64. + { rewrite load_result_long_64 by auto. rewrite D, OTH1 by eauto with asmgen. reflexivity. } + { rewrite load_result_long_32 by auto. + unfold compare_ulong in *. rewrite !compare_ulong_32 in * by auto. + rewrite D, OTH1 by eauto with asmgen. reflexivity. } Qed. Lemma transl_cond_correct_2: forall cond args k rs m b c, transl_cond cond args k = OK c -> - eval_condition cond (map rs (map preg_of args)) m = Some b -> + eval_condition cond (map (fun r => rs#r) (map preg_of args)) m = Some b -> exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(reg_of_crbit (fst (crbit_for_cond cond))) = @@ -1046,7 +1247,7 @@ Lemma transl_cond_correct_2: Proof. intros. replace (Val.of_bool b) - with (Val.of_optbool (eval_condition cond rs ## (preg_of ## args) m)). + with (Val.of_optbool (eval_condition cond (map (fun r => rs#r) (map preg_of args)) m)). eapply transl_cond_correct_1; eauto. rewrite H0; auto. Qed. @@ -1055,7 +1256,7 @@ Lemma transl_cond_correct_3: forall cond args k ms sp rs m b m' c, transl_cond cond args k = OK c -> agree ms sp rs -> - eval_condition cond (map ms args) m = Some b -> + eval_condition cond (map (fun r => Regmap.get r ms) args) m = Some b -> Mem.extends m m' -> exists rs', exec_straight ge fn c rs m' k rs' m' @@ -1126,37 +1327,48 @@ Lemma transl_cond_op_correct: transl_cond_op cond args r k = OK c -> exists rs', exec_straight ge fn c rs m k rs' m - /\ rs'#(preg_of r) = Val.of_optbool (eval_condition cond (map rs (map preg_of args)) m) + /\ rs'#(preg_of r) = Val.of_optbool (eval_condition cond (map (fun r => rs#r) (map preg_of args)) m) /\ forall r', important_preg r' = true -> preg_notin r' (destroyed_by_cond cond) -> r' <> preg_of r -> rs'#r' = rs#r'. Proof. intros until args. unfold transl_cond_op. destruct (classify_condition cond args); intros; monadInv H; simpl; erewrite ! ireg_of_eq; eauto. (* eq 0 *) - econstructor; split. - eapply exec_straight_two; simpl; reflexivity. - split. Simpl. destruct (rs x0); simpl; auto. +- econstructor; split. ++ eapply exec_straight_two; try reflexivity; try rewrite nextinstr_pc; Simpl. ++ Simpl. split. + destruct (rs # x0); simpl; destruct Archi.ppc64; auto. + apply add_carry_eq0. apply add_carry_eq0. intros; Simpl. (* ne 0 *) - econstructor; split. - eapply exec_straight_two; simpl; reflexivity. +- econstructor; split. + eapply exec_straight_two; simpl; try reflexivity; try rewrite nextinstr_pc; Simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. - split. Simpl. destruct (rs x0); simpl; auto. + split. Simpl. destruct (rs # x0); simpl; destruct Archi.ppc64; auto. + apply add_carry_ne0. apply add_carry_ne0. intros; Simpl. (* ge 0 *) - econstructor; split. - eapply exec_straight_two; simpl; reflexivity. - split. Simpl. rewrite Val.rolm_ge_zero. auto. +- econstructor; split. + eapply exec_straight_two; simpl; try reflexivity; try rewrite nextinstr_pc; Simpl. + split. Simpl. simpl. + destruct Archi.ppc64. + simpl; rewrite Val.rolm_ge_zero; auto. + destruct (rs # x0); simpl; auto. + change (Vint (Int.xor (Int.rolm i Int.one Int.one) Int.one)) + with (Val.xor (Val.rolm (Vint i) Int.one Int.one) (Vint Int.one)). + rewrite Val.rolm_ge_zero. auto. intros; Simpl. (* lt 0 *) - econstructor; split. - apply exec_straight_one; simpl; reflexivity. - split. Simpl. rewrite Val.rolm_lt_zero. auto. +- econstructor; split. + apply exec_straight_one; simpl; try reflexivity; try rewrite nextinstr_pc; Simpl. + split. Simpl. rewrite Val.rolm_lt_zero. simpl. + destruct Archi.ppc64. + simpl; reflexivity. unfold Val.cmp; apply load_result_compare. intros; Simpl. (* default *) - set (bit := fst (crbit_for_cond c)) in *. +- set (bit := fst (crbit_for_cond c)) in *. set (isset := snd (crbit_for_cond c)) in *. set (k1 := Pmfcrbit x bit :: @@ -1168,39 +1380,80 @@ Proof. intros [rs1 [EX1 [RES1 AG1]]]. destruct isset. (* bit set *) - econstructor; split. eapply exec_straight_trans. eexact EX1. - unfold k1. apply exec_straight_one; simpl; reflexivity. ++ econstructor; split. eapply exec_straight_trans. eexact EX1. + unfold k1. apply exec_straight_one; simpl; try reflexivity; try rewrite nextinstr_pc; Simpl. intuition Simpl. + rewrite RES1. simpl; destruct Archi.ppc64; auto using load_result_compare. + (* bit clear *) - econstructor; split. eapply exec_straight_trans. eexact EX1. - unfold k1. eapply exec_straight_two; simpl; reflexivity. ++ econstructor; split. eapply exec_straight_trans. eexact EX1. + unfold k1. eapply exec_straight_two; simpl; try reflexivity; try rewrite nextinstr_pc; Simpl. intuition Simpl. - rewrite RES1. destruct (eval_condition c rs ## (preg_of ## rl) m). destruct b; auto. auto. + rewrite RES1. simpl Preg.chunk_of. + destruct (eval_condition c (map (fun r => rs#r) (map preg_of rl)) m), Archi.ppc64; auto. + destruct b; auto. destruct b; auto. Qed. (** Translation of arithmetic operations. *) +Ltac invert_load_result := + match goal with + | [ |- ?f ?arg1 ?arg2 ?arg3 = Val.load_result ?c (?f ?arg1 ?arg2 ?arg3) ] => + destruct arg1, arg2, arg3; simpl; auto + | [ |- ?f ?arg1 ?arg2 = Val.load_result ?c (?f ?arg1 ?arg2) ] => + destruct arg1, arg2; simpl; destruct Archi.ppc64; auto; + match goal with + | [ |- (if ?cond then _ else _) = Val.load_result _ (if ?cond then _ else _) ] => + destruct cond; auto + | _ => + idtac + end + | [ |- ?f ?arg1 = Val.load_result ?c (?f ?arg1) ] => + destruct arg1; simpl; destruct Archi.ppc64; auto + | [ |- match ?x with _ => _ end = Val.load_result _ (match ?x with _ => _ end) ] => + destruct x; simpl; destruct Archi.ppc64; auto + end. + +Ltac ResolveNextinstrPC := + match goal with + | [ |- Pregmap.get PC (nextinstr _) = Val.offset_ptr (Pregmap.get PC _) _ ] => + rewrite nextinstr_pc; Simpl; reflexivity + | _ => reflexivity + end. + Ltac TranslOpSimpl := econstructor; split; - [ apply exec_straight_one; [simpl; eauto | reflexivity] - | split; [ apply Val.lessdef_same; Simpl; fail | intros; Simpl; fail ] ]. + [ apply exec_straight_one; [simpl; eauto | ResolveNextinstrPC] + | split; [ apply Val.lessdef_same; Simpl; try invert_load_result; fail | intros; Simpl; fail ] ]. + +Ltac TranslOpSimpl64 := + econstructor; split; + [ eapply exec_straight_one; [simpl; eauto | ResolveNextinstrPC] + | split; [Simpl; simpl; destruct Archi.ppc64; auto; congruence | intros; Simpl] ]. Lemma transl_op_correct_aux: forall op args res k (rs: regset) m v c, transl_op op args res k = OK c -> - eval_operation ge (rs#GPR1) op (map rs (map preg_of args)) m = Some v -> + eval_operation ge (rs#GPR1) op (map (fun r => rs#r) (map preg_of args)) m = Some v -> exists rs', exec_straight ge fn c rs m k rs' m /\ Val.lessdef v rs'#(preg_of res) - /\ forall r, important_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r. + /\ forall r, important_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs'#r = rs#r. Proof. assert (SAME: forall v1 v2, v1 = v2 -> Val.lessdef v2 v1). { intros; subst; auto. } Opaque Int.eq. - intros. unfold transl_op in H; destruct op; ArgsInv; simpl in H0; try (inv H0); try TranslOpSimpl. + intros. + unfold transl_op in H; destruct op; ArgsInv; simpl in H0; + try (inv H0); try TranslOpSimpl; try TranslOpSimpl64. - (* Omove *) destruct (preg_of res) eqn:RES; destruct (preg_of m0) eqn:ARG; inv H. - TranslOpSimpl. - TranslOpSimpl. + (* ireg *) + econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. intuition Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + change (Preg.type i) with (Preg.type i0). + apply Pregmap.get_has_type. + (* freg *) + econstructor; split. apply exec_straight_one; simpl; eauto. intuition Simpl. intuition Simpl. - (* Ointconst *) destruct (loadimm_correct x i k rs m) as [rs' [A [B C]]]. exists rs'. rewrite B. auto with asmgen. @@ -1209,23 +1462,25 @@ Opaque Int.eq. destruct (symbol_is_small_data i i0) eqn:SD; [ | destruct (symbol_is_rel_data i i0) ]. + (* small data *) Opaque Val.add. - econstructor; split. apply exec_straight_one; simpl; reflexivity. - split. apply SAME. Simpl. rewrite small_data_area_addressing by auto. apply add_zero_symbol_address. + econstructor; split. apply exec_straight_one; simpl. intuition Simpl. Simpl. + split. apply SAME. Simpl. rewrite small_data_area_addressing by auto. rewrite add_zero_symbol_address. + unfold Genv.symbol_address in *. simpl Preg.chunk_of. + destruct (Genv.find_symbol ge i), Archi.ppc64; auto. intros; Simpl. + (* relative data *) - econstructor; split. eapply exec_straight_two; simpl; reflexivity. + econstructor; split. eapply exec_straight_two; simpl; intuition Simpl. split. apply SAME. Simpl. rewrite gpr_or_zero_not_zero by eauto with asmgen. Simpl. - apply low_high_half_zero. + apply low_high_half_of_ireg. intros; Simpl. + (* absolute data *) - econstructor; split. eapply exec_straight_two; simpl; reflexivity. + econstructor; split. eapply exec_straight_two; simpl; intuition Simpl. split. apply SAME. Simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. Simpl. - apply low_high_half_zero. + apply low_high_half_of_ireg. intros; Simpl. - (* Oaddrstack *) destruct (addimm_correct x GPR1 (Ptrofs.to_int i) k rs m) as [rs' [EX [RES OTH]]]; eauto with asmgen. exists rs'; split. auto. split; auto with asmgen. - rewrite RES. destruct (rs GPR1); simpl; auto. + rewrite RES. destruct (rs # GPR1); simpl; auto. Transparent Val.add. simpl. rewrite Ptrofs.of_int_to_int; auto. Opaque Val.add. @@ -1235,53 +1490,74 @@ Opaque Val.add. - (* Oaddsymbol *) destruct (symbol_is_small_data i i0) eqn:SD; [ | destruct (symbol_is_rel_data i i0) ]. + (* small data *) - econstructor; split. eapply exec_straight_two; simpl; reflexivity. - split. apply SAME. Simpl. rewrite (Val.add_commut (rs x)). f_equal. - rewrite small_data_area_addressing by auto. apply add_zero_symbol_address. + econstructor; split. eapply exec_straight_two; simpl; intuition Simpl. + split. apply SAME. Simpl. rewrite (Val.add_commut (rs # x)). + rewrite small_data_area_addressing, add_zero_symbol_address by auto. + unfold Genv.symbol_address. simpl Preg.chunk_of. + destruct (Genv.find_symbol ge i), Archi.ppc64; auto. + destruct (Pregmap.get x rs); simpl; auto. intros; Simpl. + (* relative data *) econstructor; split. eapply exec_straight_trans. - eapply exec_straight_two; simpl; reflexivity. - eapply exec_straight_two; simpl; reflexivity. + eapply exec_straight_two; simpl; eauto; rewrite !nextinstr_pc; Simpl. + eapply exec_straight_two; simpl; eauto; rewrite !nextinstr_pc; Simpl. split. assert (GPR0 <> x0) by (apply not_eq_sym; eauto with asmgen). Simpl. rewrite ! gpr_or_zero_zero. rewrite ! gpr_or_zero_not_zero by eauto with asmgen. Simpl. - rewrite low_high_half_zero. auto. + rewrite !Val.load_result_add by (simpl; destruct Archi.ppc64; auto). + rewrite low_high_half_zero. + change (Preg.chunk_of GPR0) with (Preg.chunk_of x). + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto using Pregmap.get_has_type. intros; Simpl. + (* absolute data *) - econstructor; split. eapply exec_straight_two; simpl; reflexivity. + econstructor; split. eapply exec_straight_two; simpl; eauto; Simpl. split. Simpl. rewrite ! gpr_or_zero_not_zero by (eauto with asmgen). Simpl. - rewrite Val.add_assoc. rewrite (Val.add_commut (rs x)). rewrite low_high_half. auto. + rewrite !Val.load_result_add by (simpl; destruct Archi.ppc64; auto). + rewrite Val.add_assoc. rewrite (Val.add_commut (rs # x)). rewrite low_high_half. auto. intros; Simpl. - (* Osubimm *) case (Int.eq (high_s i) Int.zero). TranslOpSimpl. destruct (loadimm_correct GPR0 i (Psubfc x0 x GPR0 :: k) rs m) as [rs1 [EX [RES OTH]]]. econstructor; split. - eapply exec_straight_trans. eexact EX. apply exec_straight_one; simpl; reflexivity. - split. Simpl. rewrite RES. rewrite OTH; eauto with asmgen. + eapply exec_straight_trans. eexact EX. apply exec_straight_one; simpl; eauto; Simpl. + split. Simpl. rewrite RES. rewrite OTH; destruct (rs # x); eauto with asmgen. + simpl. destruct Archi.ppc64; auto. intros; Simpl. - (* Omulimm *) case (Int.eq (high_s i) Int.zero). TranslOpSimpl. destruct (loadimm_correct GPR0 i (Pmullw x0 x GPR0 :: k) rs m) as [rs1 [EX [RES OTH]]]. econstructor; split. - eapply exec_straight_trans. eexact EX. apply exec_straight_one; simpl; reflexivity. + eapply exec_straight_trans. eexact EX. apply exec_straight_one; simpl; eauto; Simpl. split. Simpl. rewrite RES. rewrite OTH; eauto with asmgen. + simpl. destruct (Pregmap.get x rs), Archi.ppc64; auto. intros; Simpl. - (* Odivs *) - replace v with (Val.maketotal (Val.divs (rs x) (rs x0))). - TranslOpSimpl. + replace v with (Val.maketotal (Val.divs (rs # x) (rs # x0))). + econstructor; split. + apply exec_straight_one; [simpl; eauto | ResolveNextinstrPC]. + split. apply Val.lessdef_same; Simpl. + destruct (Pregmap.get x rs), (Pregmap.get x0 rs); simpl; destruct Archi.ppc64; auto. + destruct (Int.eq _ _ || _); auto. + intros; Simpl. rewrite H1; auto. - (* Odivu *) - replace v with (Val.maketotal (Val.divu (rs x) (rs x0))). - TranslOpSimpl. + replace v with (Val.maketotal (Val.divu (rs # x) (rs # x0))). + econstructor; split. + apply exec_straight_one; [simpl; eauto | ResolveNextinstrPC]. + split. apply Val.lessdef_same; Simpl. + destruct (Pregmap.get x rs), (Pregmap.get x0 rs); simpl; destruct Archi.ppc64; auto. + destruct (Int.eq _ _); auto. + intros; Simpl. rewrite H1; auto. - (* Oand *) - set (v' := Val.and (rs x) (rs x0)) in *. + set (v' := Val.and (rs # x) (rs # x0)) in *. pose (rs1 := rs#x1 <- v'). destruct (compare_sint_spec rs1 v' Vzero) as [A [B [C D]]]. - econstructor; split. apply exec_straight_one; simpl; reflexivity. - split. rewrite D; auto with asmgen. unfold rs1; Simpl. + econstructor; split. apply exec_straight_one; simpl; eauto. + rewrite nextinstr_pc. f_equal. unfold compare_sint; Simpl. + split. rewrite D; auto with asmgen. + unfold rs1. destruct (rs # x), (rs # x0); Simpl. simpl; destruct Archi.ppc64; auto. intros. rewrite D; auto with asmgen. unfold rs1; Simpl. - (* Oandimm *) destruct (andimm_correct x0 x i k rs m) as [rs' [A [B C]]]. eauto with asmgen. @@ -1292,87 +1568,106 @@ Opaque Val.add. - (* Oxorimm *) destruct (xorimm_correct x0 x i k rs m) as [rs' [A [B C]]]. exists rs'; auto with asmgen. -- (* Onor *) - replace (Val.notint (rs x)) - with (Val.notint (Val.or (rs x) (rs x))). - TranslOpSimpl. - destruct (rs x); simpl; auto. rewrite Int.or_idem. auto. +- (* Onot *) + replace (Val.notint (rs # x)) + with (Val.notint (Val.or (rs # x) (rs # x))). + econstructor; split. + apply exec_straight_one; [simpl; eauto | ResolveNextinstrPC]. + split. apply Val.lessdef_same; Simpl. + destruct (rs # x); simpl; destruct Archi.ppc64; auto. + intros; Simpl. + destruct (rs # x); simpl; auto. rewrite Int.or_idem. auto. - (* Oshrximm *) econstructor; split. - eapply exec_straight_two; simpl; reflexivity. - split. Simpl. apply SAME. apply Val.shrx_carry. auto. + eapply exec_straight_two; simpl; eauto; Simpl. + split. Simpl. apply SAME. + apply Val.shrx_carry in H1. simpl Preg.chunk_of. + destruct Archi.ppc64, (rs # x); simpl in *; try destruct (Int.ltu _ _); auto. intros; Simpl. - (* Orolm *) destruct (rolm_correct x0 x i i0 k rs m) as [rs' [A [B C]]]. eauto with asmgen. exists rs'; auto. - (* Olongconst *) destruct (loadimm64_correct x i k rs m) as [rs' [A [B C]]]. - exists rs'; auto with asmgen. + econstructor; split; eauto with asmgen. + split; eauto with asmgen. + rewrite B; simpl; rewrite Heqb; auto. - (* Oaddlimm *) - destruct (addimm64_correct x0 x i k rs m) as [rs' [A [B C]]]. eauto with asmgen. + destruct (addimm64_correct x0 x i k rs m) as [rs' [A [B C]]]; eauto with asmgen. exists rs'; auto with asmgen. - (* Odivl *) - replace v with (Val.maketotal (Val.divls (rs x) (rs x0))). - TranslOpSimpl. + replace v with (Val.maketotal (Val.divls (rs # x) (rs # x0))). + TranslOpSimpl64. rewrite H1; auto. - (* Odivlu *) - replace v with (Val.maketotal (Val.divlu (rs x) (rs x0))). - TranslOpSimpl. + replace v with (Val.maketotal (Val.divlu (rs # x) (rs # x0))). + TranslOpSimpl64. rewrite H1; auto. - (* Oandl *) - set (v' := Val.andl (rs x) (rs x0)) in *. + set (v' := Val.andl (rs # x) (rs # x0)) in *. pose (rs1 := rs#x1 <- v'). destruct (compare_slong_spec rs1 v' (Vlong Int64.zero)) as [A [B [C D]]]. - econstructor; split. apply exec_straight_one; simpl; reflexivity. - split. rewrite D; auto with asmgen. unfold rs1; Simpl. + econstructor; split. apply exec_straight_one; simpl; eauto. unfold compare_slong; Simpl. + split. rewrite D; auto with asmgen. unfold rs1; Simpl. simpl; rewrite Heqb; auto. intros. rewrite D; auto with asmgen. unfold rs1; Simpl. - (* Oandlimm *) - destruct (andimm64_correct x0 x i k rs m) as [rs' [A [B C]]]. eauto with asmgen. + destruct (andimm64_correct x0 x i k rs m) as [rs' [A [B C]]]; eauto with asmgen. exists rs'; auto with asmgen. - (* Oorlimm *) - destruct (orimm64_correct x0 x i k rs m) as [rs' [A [B C]]]. eauto with asmgen. + destruct (orimm64_correct x0 x i k rs m) as [rs' [A [B C]]]; eauto with asmgen. exists rs'; auto with asmgen. - (* Oxorlimm *) - destruct (xorimm64_correct x0 x i k rs m) as [rs' [A [B C]]]. eauto with asmgen. + destruct (xorimm64_correct x0 x i k rs m) as [rs' [A [B C]]]; eauto with asmgen. exists rs'; auto with asmgen. - (* Onotl *) - econstructor; split. eapply exec_straight_one; simpl; reflexivity. - split. Simpl. destruct (rs x); simpl; auto. rewrite Int64.or_idem; auto. + econstructor; split. eapply exec_straight_one; simpl; eauto. Simpl. + split. Simpl. destruct (rs # x); simpl; auto. rewrite Int64.or_idem, Heqb; auto. intros; Simpl. - (* Oshrxlimm *) econstructor; split. - eapply exec_straight_two; simpl; reflexivity. - split. Simpl. apply SAME. apply Val.shrxl_carry. auto. + eapply exec_straight_two; simpl; eauto; Simpl. + split. Simpl. apply SAME. + simpl Preg.chunk_of; rewrite Heqb; simpl. apply Val.shrxl_carry; auto. intros; Simpl. - (* Orolml *) - destruct (rolm64_correct x0 x i i0 k rs m) as [rs' [A [B C]]]. eauto with asmgen. + destruct (rolm64_correct x0 x i i0 k rs m) as [rs' [A [B C]]]; eauto with asmgen. exists rs'; auto with asmgen. - (* Olongoffloat *) - replace v with (Val.maketotal (Val.longoffloat (rs x))). - TranslOpSimpl. + replace v with (Val.maketotal (Val.longoffloat (rs # x))). + TranslOpSimpl64. rewrite H1; auto. - (* Ofloatoflong *) - replace v with (Val.maketotal (Val.floatoflong (rs x))). + replace v with (Val.maketotal (Val.floatoflong (rs # x))). TranslOpSimpl. rewrite H1; auto. - (* Ointoffloat *) -- replace v with (Val.maketotal (Val.intoffloat (rs x))). - TranslOpSimpl. +- (* Ointoffloat *) + replace v with (Val.maketotal (Val.intoffloat (rs # x))). + econstructor; split. + eapply exec_straight_one; [simpl; eauto | ResolveNextinstrPC]. + split. + apply Val.lessdef_same; Simpl; simpl. destruct Archi.ppc64; auto. + destruct (rs # x); simpl; auto. destruct (Float.to_int f); simpl; auto. + intros; Simpl. rewrite H1; auto. - (* Ointuoffloat *) -- replace v with (Val.maketotal (Val.intuoffloat (rs x))). - TranslOpSimpl. +- (* Ointuoffloat *) + replace v with (Val.maketotal (Val.intuoffloat (rs # x))). + econstructor; split. + eapply exec_straight_one; [simpl; eauto | ResolveNextinstrPC]. + split. + apply Val.lessdef_same; Simpl; simpl. destruct Archi.ppc64; auto. + destruct (rs # x); simpl; auto. destruct (Float.to_intu f); simpl; auto. + intros; Simpl. rewrite H1; auto. - (* Ofloatofint *) -- replace v with (Val.maketotal (Val.floatofint (rs x))). +- (* Ofloatofint *) + replace v with (Val.maketotal (Val.floatofint (rs # x))). TranslOpSimpl. rewrite H1; auto. - (* Ofloatofintu *) -- replace v with (Val.maketotal (Val.floatofintu (rs x))). +- (* Ofloatofintu *) + replace v with (Val.maketotal (Val.floatofintu (rs # x))). TranslOpSimpl. rewrite H1; auto. - (* Ocmp *) -- destruct (transl_cond_op_correct c0 args res k rs m c) as [rs' [A [B C]]]; auto. +- (* Ocmp *) + destruct (transl_cond_op_correct c0 args res k rs m c) as [rs' [A [B C]]]; auto. exists rs'; auto with asmgen. Qed. @@ -1380,12 +1675,12 @@ Lemma transl_op_correct: forall op args res k ms sp rs m v m' c, transl_op op args res k = OK c -> agree ms sp rs -> - eval_operation ge sp op (map ms args) m = Some v -> + eval_operation ge sp op (map (fun r => Regmap.get r ms) args) m = Some v -> Mem.extends m m' -> exists rs', exec_straight ge fn c rs m' k rs' m' /\ agree (Regmap.set res v (Mach.undef_regs (destroyed_by_op op) ms)) sp rs' - /\ forall r, important_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r. + /\ forall r, important_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs'#r = rs#r. Proof. intros. exploit eval_operation_lessdef. eapply preg_vals; eauto. eauto. eauto. @@ -1401,16 +1696,16 @@ Qed. Lemma transl_memory_access_correct: forall (P: regset -> Prop) mk1 mk2 addr args temp k c (rs: regset) a m m', transl_memory_access mk1 mk2 addr args temp k = OK c -> - eval_addressing ge (rs#GPR1) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#GPR1) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> temp <> GPR0 -> (forall cst (r1: ireg) (rs1: regset) k, Val.lessdef a (Val.add (gpr_or_zero rs1 r1) (const_low ge cst)) -> - (forall r, r <> PC -> r <> temp -> r <> GPR0 -> rs1 r = rs r) -> + (forall r, r <> PC -> r <> temp -> r <> GPR0 -> rs1#r = rs#r) -> exists rs', exec_straight ge fn (mk1 cst r1 :: k) rs1 m k rs' m' /\ P rs') -> (forall (r1 r2: ireg) (rs1: regset) k, Val.lessdef a (Val.add rs1#r1 rs1#r2) -> - (forall r, r <> PC -> r <> temp -> r <> GPR0 -> rs1 r = rs r) -> + (forall r, r <> PC -> r <> temp -> r <> GPR0 -> rs1#r = rs#r) -> exists rs', exec_straight ge fn (mk2 r1 r2 :: k) rs1 m k rs' m' /\ P rs') -> exists rs', @@ -1418,106 +1713,127 @@ Lemma transl_memory_access_correct: Proof. intros until m'; intros TR ADDR TEMP MK1 MK2. unfold transl_memory_access in TR; destruct addr; ArgsInv; simpl in ADDR; inv ADDR. - (* Aindexed *) +- (* Aindexed *) case (Int.eq (high_s i) Int.zero). - (* Aindexed short *) ++ (* Aindexed short *) apply MK1. rewrite gpr_or_zero_not_zero; eauto with asmgen. auto. - (* Aindexed long *) - set (rs1 := nextinstr (rs#temp <- (Val.add (rs x) (Vint (Int.shl (high_s i) (Int.repr 16)))))). ++ (* Aindexed long *) + set (rs1 := nextinstr (rs#temp <- (Val.add (rs#x) (Vint (Int.shl (high_s i) (Int.repr 16)))))). exploit (MK1 (Cint (low_s i)) temp rs1 k). simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. - unfold rs1; Simpl. rewrite Val.add_assoc. + unfold rs1; Simpl. rewrite Val.load_result_add by (simpl; destruct Archi.ppc64; auto). + rewrite Val.add_assoc. Transparent Val.add. simpl. rewrite low_high_s. auto. intros; unfold rs1; Simpl. intros [rs' [EX' AG']]. exists rs'. split. apply exec_straight_step with rs1 m. - simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. auto. + simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. + subst rs1. Simpl. auto. auto. - (* Aindexed2 *) +- (* Aindexed2 *) apply MK2; auto. - (* Aglobal *) +- (* Aglobal *) destruct (symbol_is_small_data i i0) eqn:SISD; [ | destruct (symbol_is_rel_data i i0) ]; inv TR. - (* Aglobal from small data *) ++ (* Aglobal from small data *) apply MK1. unfold const_low. rewrite small_data_area_addressing by auto. rewrite add_zero_symbol_address. auto. auto. - (* Aglobal from relative data *) ++ (* Aglobal from relative data *) set (rs1 := nextinstr (rs#temp <- (Val.add Vzero (high_half ge i i0)))). set (rs2 := nextinstr (rs1#temp <- (Genv.symbol_address ge i i0))). exploit (MK1 (Cint Int.zero) temp rs2). simpl. rewrite gpr_or_zero_not_zero by eauto with asmgen. - unfold rs2. Simpl. rewrite Val.add_commut. rewrite add_zero_symbol_address. auto. + unfold rs2. Simpl. rewrite Val.add_commut. + rewrite Preg.chunk_of_reg_type, Val.load_result_same. + rewrite add_zero_symbol_address. auto. + unfold Genv.symbol_address; simpl. destruct (Genv.find_symbol ge i), Archi.ppc64; simpl; auto. intros; unfold rs2, rs1; Simpl. intros [rs' [EX' AG']]. exists rs'; split. apply exec_straight_step with rs1 m; auto. + unfold rs1; Simpl. apply exec_straight_step with rs2 m; auto. simpl. unfold rs2. rewrite gpr_or_zero_not_zero by eauto with asmgen. f_equal. f_equal. f_equal. - unfold rs1; Simpl. apply low_high_half_zero. + unfold rs1; Simpl. rewrite Val.load_result_add. apply low_high_half_zero. + simpl; destruct Archi.ppc64; auto. + unfold rs2; Simpl. eexact EX'. auto. - (* Aglobal from absolute data *) ++ (* Aglobal from absolute data *) set (rs1 := nextinstr (rs#temp <- (Val.add Vzero (high_half ge i i0)))). exploit (MK1 (Csymbol_low i i0) temp rs1). simpl. rewrite gpr_or_zero_not_zero by eauto with asmgen. - unfold rs1. Simpl. rewrite low_high_half_zero. auto. + unfold rs1. Simpl. rewrite Val.load_result_add, low_high_half_zero. auto. + simpl; destruct Archi.ppc64; auto. intros; unfold rs1; Simpl. intros [rs' [EX' AG']]. exists rs'; split. apply exec_straight_step with rs1 m; auto. + unfold rs1; Simpl. eexact EX'. auto. - (* Abased *) +- (* Abased *) destruct (symbol_is_small_data i i0) eqn:SISD; [ | destruct (symbol_is_rel_data i i0) ]. - (* Abased from small data *) ++ (* Abased from small data *) set (rs1 := nextinstr (rs#GPR0 <- (Genv.symbol_address ge i i0))). exploit (MK2 x GPR0 rs1 k). - unfold rs1; Simpl. rewrite Val.add_commut. auto. + unfold rs1; Simpl. rewrite Val.add_commut. + rewrite Preg.chunk_of_reg_type, Val.load_result_same. + auto. + unfold Genv.symbol_address; simpl. destruct (Genv.find_symbol ge i), Archi.ppc64; simpl; auto. intros. unfold rs1; Simpl. intros [rs' [EX' AG']]. exists rs'; split. apply exec_straight_step with rs1 m. unfold exec_instr. rewrite gpr_or_zero_zero. f_equal. unfold rs1. f_equal. f_equal. unfold const_low. rewrite small_data_area_addressing; auto. apply add_zero_symbol_address. - reflexivity. + unfold rs1; Simpl. assumption. assumption. - (* Abased from relative data *) ++ (* Abased from relative data *) set (rs1 := nextinstr (rs#GPR0 <- (rs#x))). set (rs2 := nextinstr (rs1#temp <- (Val.add Vzero (high_half ge i i0)))). set (rs3 := nextinstr (rs2#temp <- (Genv.symbol_address ge i i0))). exploit (MK2 temp GPR0 rs3). f_equal. unfold rs3; Simpl. unfold rs3, rs2, rs1; Simpl. + simpl. destruct Archi.ppc64; auto. + unfold Genv.symbol_address; simpl. + destruct (Genv.find_symbol ge i), (rs # x); simpl; auto. intros. unfold rs3, rs2, rs1; Simpl. intros [rs' [EX' AG']]. exists rs'. split. eapply exec_straight_trans with (rs2 := rs3) (m2 := m). apply exec_straight_three with rs1 m rs2 m; auto. simpl. unfold rs3. f_equal. f_equal. f_equal. rewrite gpr_or_zero_not_zero by auto. - unfold rs2; Simpl. apply low_high_half_zero. + unfold rs2; Simpl. rewrite Val.load_result_add. apply low_high_half_zero. + simpl; destruct Archi.ppc64; auto. + unfold rs1; Simpl. + unfold rs2; Simpl. + unfold rs3; Simpl. eexact EX'. auto. - (* Abased absolute *) - set (rs1 := nextinstr (rs#temp <- (Val.add (rs x) (high_half ge i i0)))). ++ (* Abased absolute *) + set (rs1 := nextinstr (rs#temp <- (Val.add (rs#x) (high_half ge i i0)))). exploit (MK1 (Csymbol_low i i0) temp rs1 k). simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. - unfold rs1. Simpl. + unfold rs1. Simpl. rewrite Val.load_result_add by (simpl; destruct Archi.ppc64; auto). rewrite Val.add_assoc. rewrite low_high_half. rewrite Val.add_commut. auto. intros; unfold rs1; Simpl. intros [rs' [EX' AG']]. exists rs'. split. apply exec_straight_step with rs1 m. - unfold exec_instr. rewrite gpr_or_zero_not_zero; eauto with asmgen. auto. + unfold exec_instr. rewrite gpr_or_zero_not_zero; eauto with asmgen. unfold rs1; Simpl. assumption. assumption. - (* Ainstack *) +- (* Ainstack *) set (ofs := Ptrofs.to_int i) in *. - assert (L: Val.lessdef (Val.offset_ptr (rs GPR1) i) (Val.add (rs GPR1) (Vint ofs))). - { destruct (rs GPR1); simpl; auto. unfold ofs; rewrite Ptrofs.of_int_to_int; auto. } + assert (L: Val.lessdef (Val.offset_ptr (rs#GPR1) i) (Val.add (rs#GPR1) (Vint ofs))). + { destruct (rs#GPR1); simpl; auto. unfold ofs; rewrite Ptrofs.of_int_to_int; auto. } destruct (Int.eq (high_s ofs) Int.zero); inv TR. apply MK1. simpl. rewrite gpr_or_zero_not_zero; eauto with asmgen. auto. set (rs1 := nextinstr (rs#temp <- (Val.add rs#GPR1 (Vint (Int.shl (high_s ofs) (Int.repr 16)))))). exploit (MK1 (Cint (low_s ofs)) temp rs1 k). simpl. rewrite gpr_or_zero_not_zero; auto. unfold rs1. rewrite nextinstr_inv. rewrite Pregmap.gss. + rewrite Val.load_result_add by (simpl; destruct Archi.ppc64; auto). rewrite Val.add_assoc. simpl. rewrite low_high_s. auto. congruence. intros. unfold rs1. rewrite nextinstr_inv; auto. apply Pregmap.gso; auto. intros [rs' [EX' AG']]. exists rs'. split. apply exec_straight_step with rs1 m. - unfold exec_instr. rewrite gpr_or_zero_not_zero; eauto with asmgen. auto. + unfold exec_instr. rewrite gpr_or_zero_not_zero; eauto with asmgen. unfold rs1; Simpl. assumption. assumption. Qed. @@ -1526,12 +1842,13 @@ Qed. 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 -> - eval_addressing ge (rs#GPR1) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#GPR1) addr (map (fun r => rs # r) (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> + (*Val.has_type v (Preg.type (preg_of dst)) ->*) exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of dst) = v - /\ forall r, r <> PC -> r <> GPR12 -> r <> GPR0 -> r <> preg_of dst -> rs' r = rs r. + /\ forall r, r <> PC -> r <> GPR12 -> r <> GPR0 -> r <> preg_of dst -> rs' # r = rs # r. Proof. intros. assert (LD: forall v, Val.lessdef a v -> v = a). @@ -1539,6 +1856,7 @@ Proof. assert (BASE: forall mk1 mk2 k' chunk' v', transl_memory_access mk1 mk2 addr args GPR12 k' = OK c -> Mem.loadv chunk' m a = Some v' -> + Val.has_type v' (Preg.type (preg_of dst)) -> (forall cst (r1: ireg) (rs1: regset), exec_instr ge fn (mk1 cst r1) rs1 m = load1 ge chunk' (preg_of dst) cst r1 rs1 m) -> @@ -1548,17 +1866,19 @@ Proof. exists rs', exec_straight ge fn c rs m k' rs' m /\ rs'#(preg_of dst) = v' - /\ forall r, r <> PC -> r <> GPR12 -> r <> GPR0 -> r <> preg_of dst -> rs' r = rs r). + /\ forall r, r <> PC -> r <> GPR12 -> r <> GPR0 -> r <> preg_of dst -> rs' # r = rs # r). { intros. eapply transl_memory_access_correct; eauto. congruence. intros. econstructor; split. apply exec_straight_one. - rewrite H4. unfold load1. apply LD in H6. rewrite H6. rewrite H3. eauto. + rewrite H5. unfold load1. apply LD in H7. rewrite H7. rewrite H3. eauto. unfold nextinstr. rewrite Pregmap.gss. rewrite Pregmap.gso; auto with asmgen. - intuition Simpl. + simpl. destruct (rs1 # PC), Archi.ppc64; simpl; auto. + intuition Simpl. rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. intros. econstructor; split. apply exec_straight_one. - rewrite H5. unfold load2. apply LD in H6. rewrite H6. rewrite H3. eauto. + rewrite H6. unfold load2. apply LD in H7. rewrite H7. rewrite H3. eauto. unfold nextinstr. rewrite Pregmap.gss. rewrite Pregmap.gso; auto with asmgen. - intuition Simpl. + simpl. destruct (rs1 # PC), Archi.ppc64; simpl; auto. + intuition Simpl. rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. } destruct chunk; monadInv H. - (* Mint8signed *) @@ -1571,24 +1891,40 @@ Proof. } destruct H as [v1 [LD' SG]]. clear H1. exploit BASE; eauto; erewrite ireg_of_eq by eauto; auto. + apply Mem.loadv_type in LD'. + simpl. destruct Archi.ppc64; apply Val.has_subtype with (ty1 := Tint); eauto. intros [rs1 [A [B C]]]. econstructor; split. - eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto. - split. Simpl. congruence. intros. Simpl. + eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. Simpl. + split. Simpl. + subst. simpl. destruct (rs1 # x), Archi.ppc64; simpl; auto. + intros. Simpl. - (* Mint8unsigned *) eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. + exploit Mem.loadv_type; eauto; intro. + exploit Val.has_subtype; eauto. destruct (Preg.type_cases x) as [T|T]; rewrite T; auto. - (* Mint816signed *) eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. + exploit Mem.loadv_type; eauto; intro. + exploit Val.has_subtype; eauto. destruct (Preg.type_cases x) as [T|T]; rewrite T; auto. - (* Mint16unsigned *) eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. + exploit Mem.loadv_type; eauto; intro. + exploit Val.has_subtype; eauto. destruct (Preg.type_cases x) as [T|T]; rewrite T; auto. - (* Mint32 *) eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. + exploit Mem.loadv_type; eauto; intro. + exploit Val.has_subtype; eauto. destruct (Preg.type_cases x) as [T|T]; rewrite T; auto. - (* Mint64 *) eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. + exploit Mem.loadv_type; eauto; intro. + exploit Val.has_subtype; eauto. simpl Preg.type. rewrite Heqb; auto. - (* Mfloat32 *) eapply BASE; eauto; erewrite freg_of_eq by eauto; auto. + auto using Val.has_type_Tany64. - (* Mfloat64 *) eapply BASE; eauto; erewrite freg_of_eq by eauto; auto. + auto using Val.has_type_Tany64. Qed. (** Translation of stores *) @@ -1596,11 +1932,11 @@ Qed. Lemma transl_store_correct: forall chunk addr args src k c (rs: regset) m a m', transl_store chunk addr args src k = OK c -> - eval_addressing ge (rs#GPR1) addr (map rs (map preg_of args)) = Some a -> - Mem.storev chunk m a (rs (preg_of src)) = Some m' -> + eval_addressing ge (rs#GPR1) addr (map (fun r => rs # r) (map preg_of args)) = Some a -> + Mem.storev chunk m a (rs # (preg_of src)) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' - /\ forall r, r <> PC -> r <> GPR0 -> preg_notin r (destroyed_by_store chunk addr) -> rs' r = rs r. + /\ forall r, r <> PC -> r <> GPR0 -> preg_notin r (destroyed_by_store chunk addr) -> rs' # r = rs # r. Proof. Local Transparent destroyed_by_store. intros. @@ -1617,7 +1953,7 @@ Local Transparent destroyed_by_store. eapply preg_of_injective; eauto. assert (BASE: forall mk1 mk2 chunk', transl_memory_access mk1 mk2 addr args (int_temp_for src) k = OK c -> - Mem.storev chunk' m a (rs (preg_of src)) = Some m' -> + Mem.storev chunk' m a (rs # (preg_of src)) = Some m' -> (forall cst (r1: ireg) (rs1: regset), exec_instr ge fn (mk1 cst r1) rs1 m = store1 ge chunk' (preg_of src) cst r1 rs1 m) -> @@ -1626,25 +1962,25 @@ Local Transparent destroyed_by_store. store2 chunk' (preg_of src) r1 r2 rs1 m) -> exists rs', exec_straight ge fn c rs m k rs' m' - /\ forall r, r <> PC -> r <> GPR0 -> r <> GPR11 /\ r <> GPR12 -> rs' r = rs r). + /\ forall r, r <> PC -> r <> GPR0 -> r <> GPR11 /\ r <> GPR12 -> rs' # r = rs # r). { intros. eapply transl_memory_access_correct; eauto. intros. econstructor; split. apply exec_straight_one. - rewrite H4. unfold store1. apply LD in H6. rewrite H6. rewrite H7; auto with asmgen. rewrite H3. eauto. auto. + rewrite H4. unfold store1. apply LD in H6. rewrite H6. rewrite H7; auto with asmgen. rewrite H3. eauto. Simpl. intros; Simpl. apply H7; auto. destruct TEMP0; destruct H10; congruence. intros. econstructor; split. apply exec_straight_one. - rewrite H5. unfold store2. apply LD in H6. rewrite H6. rewrite H7; auto with asmgen. rewrite H3. eauto. auto. + rewrite H5. unfold store2. apply LD in H6. rewrite H6. rewrite H7; auto with asmgen. rewrite H3. eauto. Simpl. intros; Simpl. apply H7; auto. destruct TEMP0; destruct H10; congruence. } destruct chunk; monadInv H. - (* Mint8signed *) - assert (Mem.storev Mint8unsigned m a (rs (preg_of src)) = Some m'). + assert (Mem.storev Mint8unsigned m a (rs # (preg_of src)) = Some m'). rewrite <- H1. destruct a; simpl; auto. symmetry. apply Mem.store_signed_unsigned_8. clear H1. eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. - (* Mint8unsigned *) eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. - (* Mint16signed *) - assert (Mem.storev Mint16unsigned m a (rs (preg_of src)) = Some m'). + assert (Mem.storev Mint16unsigned m a (rs # (preg_of src)) = Some m'). rewrite <- H1. destruct a; simpl; auto. symmetry. apply Mem.store_signed_unsigned_16. clear H1. eapply BASE; eauto; erewrite ireg_of_eq by eauto; auto. - (* Mint16unsigned *) @@ -1663,8 +1999,8 @@ Qed. Lemma transl_epilogue_correct: forall ge0 f m stk soff cs m' ms rs k tm, - load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) -> - load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs) -> + load_stack m (Vptr stk soff) (quantity_of_typ Tptr) f.(fn_link_ofs) = Some (parent_sp cs) -> + load_stack m (Vptr stk soff) (quantity_of_typ 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 -> (is_leaf_function f = true -> rs#LR = parent_ra cs) -> @@ -1688,11 +2024,15 @@ Proof. - (* leaf function *) econstructor; exists tm'. split. apply exec_straight_one. simpl. rewrite <- (sp_val _ _ _ AG). simpl. - rewrite LP'. rewrite FREE'. reflexivity. reflexivity. + rewrite LP'. rewrite FREE'. reflexivity. Simpl. split. apply agree_nextinstr. eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. + eapply parent_sp_type; eauto. + eapply Val.has_subtype; try eapply parent_sp_type; eauto. simpl; destruct Archi.ppc64; auto. split. auto. split. Simpl. split. Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype; try eapply parent_sp_type; eauto. simpl; destruct Archi.ppc64; auto. intros; Simpl. - (* regular function *) set (rs1 := nextinstr (rs#GPR0 <- (parent_ra cs))). @@ -1703,15 +2043,30 @@ Proof. simpl. unfold load1. rewrite gpr_or_zero_not_zero by congruence. unfold const_low. rewrite <- (sp_val _ _ _ AG). erewrite loadv_offset_ptr by eexact LRA'. reflexivity. - simpl. change (rs2#GPR1) with (rs#GPR1). rewrite <- (sp_val _ _ _ AG). simpl. + simpl. unfold rs2, rs1. repeat f_equal. Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype; try eapply parent_ra_type; eauto. simpl; destruct Archi.ppc64; auto. + simpl. replace (rs2#GPR1) with (rs#GPR1). rewrite <- (sp_val _ _ _ AG). simpl. rewrite LP'. rewrite FREE'. reflexivity. + unfold rs2, rs1; Simpl. + unfold rs1; Simpl. + unfold rs2; Simpl. + unfold rs3; Simpl. split. unfold rs3. apply agree_nextinstr. apply agree_change_sp with (Vptr stk soff). apply agree_nextinstr. apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto. eapply parent_sp_def; eauto. + eapply parent_sp_type; eauto. + eapply Val.has_subtype; try eapply parent_sp_type; eauto. simpl; destruct Archi.ppc64; auto. split. auto. - split. reflexivity. - split. reflexivity. + split. + unfold rs3; Simpl. unfold rs2; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype; try eapply parent_ra_type; eauto. simpl; destruct Archi.ppc64; auto. + split. + unfold rs3; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype; try eapply parent_sp_type; eauto. simpl; destruct Archi.ppc64; auto. intros. unfold rs3, rs2, rs1; Simpl. Qed. diff --git a/powerpc/Conventions1.v b/powerpc/Conventions1.v index 1de55c1af8..a6ca63200b 100644 --- a/powerpc/Conventions1.v +++ b/powerpc/Conventions1.v @@ -16,6 +16,7 @@ Require Import Coqlib. Require Import Decidableplus. Require Import AST. +Require Import Values. Require Import Events. Require Import Locations. Require Archi. @@ -144,6 +145,22 @@ Proof. destruct Archi.ptr64 eqn:?; destruct (sig_res sig) as [[]|]; destruct Archi.ppc64; simpl; auto. Qed. +Lemma loc_result_has_type: + forall res sig, + Val.has_type res (proj_sig_res sig) -> + Val.has_type_rpair res (loc_result sig) Val.loword Val.hiword mreg_type. +Proof. + intros. unfold Val.has_type_rpair, loc_result, proj_sig_res in *. +Local Transparent Archi.ptr64. + assert (P: Archi.ptr64 = false) by (unfold Archi.ptr64; auto). + destruct sig; simpl. destruct sig_res; simpl in H. + destruct t, res; simpl; auto; + simpl; try rewrite P; auto; + destruct Archi.ppc64 eqn:Q; simpl; try rewrite Q; auto. + destruct res; simpl; auto; destruct Archi.ppc64; auto. +Qed. +Local Opaque Archi.ptr64. + (** The result locations are caller-save registers *) Lemma loc_result_caller_save: @@ -206,15 +223,23 @@ Fixpoint loc_arguments_rec | (Tint | Tany32) as ty :: tys => match list_nth_z int_param_regs ir with | None => - One (S Outgoing ofs ty) :: loc_arguments_rec tys ir fr (ofs + 1) + One (S Outgoing ofs Q32) :: loc_arguments_rec tys ir fr (ofs + 1) | 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 ofs := align ofs 2 in + One (S Outgoing ofs Q32) :: loc_arguments_rec tys ir fr (ofs + 2) + | 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 - One (S Outgoing ofs ty) :: loc_arguments_rec tys ir fr (ofs + 2) + One (S Outgoing ofs Q64) :: loc_arguments_rec tys ir fr (ofs + 2) | Some freg => One (R freg) :: loc_arguments_rec tys ir (fr + 1) ofs end @@ -226,8 +251,8 @@ Fixpoint loc_arguments_rec | _, _ => let ofs := align ofs 2 in (if Archi.ptr64 - then One (S Outgoing ofs Tlong) - else Twolong (S Outgoing ofs Tint) (S Outgoing (ofs + 1) Tint)) :: + then One (S Outgoing ofs Q64) + else Twolong (S Outgoing ofs Q32) (S Outgoing (ofs + 1) Q32)) :: loc_arguments_rec tys ir fr (ofs + 2) end end. @@ -271,14 +296,14 @@ Definition size_arguments (s: signature) : Z := Definition loc_argument_acceptable (l: loc) : Prop := match l with | R r => is_callee_save r = false - | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs) + | S Outgoing ofs q => ofs >= 0 /\ (quantity_align q | ofs) | _ => False end. Definition loc_argument_charact (ofs: Z) (l: loc) : Prop := match l with | R r => In r int_param_regs \/ In r float_param_regs - | S Outgoing ofs' ty => ofs' >= ofs /\ (typealign ty | ofs') + | S Outgoing ofs' q => ofs' >= ofs /\ (quantity_align q | ofs') | _ => False end. @@ -317,11 +342,11 @@ Opaque list_nth_z. eapply IHtyl; eauto. destruct H. subst. destruct Archi.ptr64; [split|split;split]; try omega. - apply align_divides; omega. apply Z.divide_1_l. apply Z.divide_1_l. + apply Z.divide_1_l. apply Z.divide_1_l. apply Z.divide_1_l. eapply Y; eauto. omega. destruct H. subst. destruct Archi.ptr64; [split|split;split]; try omega. - apply align_divides; omega. apply Z.divide_1_l. apply Z.divide_1_l. + apply Z.divide_1_l. apply Z.divide_1_l. apply Z.divide_1_l. eapply Y; eauto. omega. - (* single *) assert (ofs <= align ofs 2) by (apply align_le; omega). @@ -398,18 +423,18 @@ Proof. 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. + forall (s: signature) (ofs: Z) (q: quantity), + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments s)) -> + ofs + typesize (typ_of_quantity q) <= 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). + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_rec tyl ir fr ofs0)) -> + ofs + typesize (typ_of_quantity q) <= size_arguments_rec tyl ir fr ofs0). { induction tyl; simpl; intros. elim H0. - destruct a. + destruct a eqn:A. - (* int *) destruct (list_nth_z int_param_regs ir); destruct H0. congruence. @@ -424,13 +449,13 @@ Proof. - (* long *) set (ir' := align ir 2) in *. assert (DFL: - In (S Outgoing ofs ty) (regs_of_rpairs + In (S Outgoing ofs q) (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)) + then One (S Outgoing (align ofs0 2) Q64) + else Twolong (S Outgoing (align ofs0 2) Q32) + (S Outgoing (align ofs0 2 + 1) Q32)) :: loc_arguments_rec tyl ir' fr (align ofs0 2 + 2))) -> - ofs + typesize ty <= size_arguments_rec tyl ir' fr (align ofs0 2 + 2)). + ofs + typesize (typ_of_quantity q) <= 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. diff --git a/powerpc/Machregs.v b/powerpc/Machregs.v index 53d99e2f3b..44c80c1fb7 100644 --- a/powerpc/Machregs.v +++ b/powerpc/Machregs.v @@ -16,6 +16,7 @@ Require Import Decidableplus. Require Import Maps. Require Import AST. Require Import Op. +Require Import Memdata. (** ** Machine registers *) @@ -93,6 +94,11 @@ Definition mreg_type (r: mreg): typ := | F24 | F25 | F26 | F27 | F28 | F29 | F30 | F31 => Tany64 end. +Lemma mreg_type_cases: forall r, mreg_type r = Tany32 \/ mreg_type r = Tany64. +Proof. + destruct r; simpl; destruct Archi.ppc64; auto. +Qed. + Open Scope positive_scope. Module IndexedMreg <: INDEXED_TYPE. @@ -100,29 +106,48 @@ Module IndexedMreg <: INDEXED_TYPE. Definition eq := mreg_eq. Definition index (r: mreg): positive := match r with - | R3 => 1 | R4 => 2 | R5 => 3 | R6 => 4 - | R7 => 5 | R8 => 6 | R9 => 7 | R10 => 8 - | R11 => 9 | R12 => 10 - | R14 => 11 | R15 => 12 | R16 => 13 - | R17 => 14 | R18 => 15 | R19 => 16 | R20 => 17 - | R21 => 18 | R22 => 19 | R23 => 20 | R24 => 21 - | R25 => 22 | R26 => 23 | R27 => 24 | R28 => 25 - | R29 => 26 | R30 => 27 | R31 => 28 - | F0 => 29 - | F1 => 30 | F2 => 31 | F3 => 32 | F4 => 33 - | F5 => 34 | F6 => 35 | F7 => 36 | F8 => 37 - | F9 => 38 | F10 => 39 | F11 => 40 | F12 => 41 - | F13 => 42 | F14 => 43 | F15 => 44 - | F16 => 45 | F17 => 46 | F18 => 47 | F19 => 48 - | F20 => 49 | F21 => 50 | F22 => 51 | F23 => 52 - | F24 => 53 | F25 => 54 | F26 => 55 | F27 => 56 - | F28 => 57 | F29 => 58 | F30 => 59 | F31 => 60 + | R3 => 2 | R4 => 4 | R5 => 6 | R6 => 8 + | R7 => 10 | R8 => 12 | R9 => 14 | R10 => 16 + | R11 => 18 | R12 => 20 + | R14 => 22 | R15 => 24 | R16 => 26 + | R17 => 28 | R18 => 30 | R19 => 32 | R20 => 34 + | R21 => 36 | R22 => 38 | R23 => 40 | R24 => 42 + | R25 => 44 | R26 => 46 | R27 => 48 | R28 => 50 + | R29 => 52 | R30 => 54 | R31 => 56 + | F0 => 58 + | F1 => 60 | F2 => 62 | F3 => 64 | F4 => 66 + | F5 => 68 | F6 => 70 | F7 => 72 | F8 => 74 + | F9 => 76 | F10 => 78 | F11 => 80 | F12 => 82 + | F13 => 84 | F14 => 86 | F15 => 88 + | F16 => 90 | F17 => 92 | F18 => 94 | F19 => 96 + | F20 => 98 | F21 => 100 | F22 => 102 | F23 => 104 + | F24 => 106 | F25 => 108 | F26 => 110 | F27 => 112 + | F28 => 114 | F29 => 116 | F30 => 118 | F31 => 120 end. Lemma index_inj: forall r1 r2, index r1 = index r2 -> r1 = r2. Proof. decide_goal. Qed. + + Open Scope Z_scope. + + Lemma scaled_index_with_size_aux: + forall r1 r2, Zpos (index r1) < Zpos (index r2) -> Zpos (index r1) + 2 <= Zpos (index r2). + Proof. + decide_goal. + Qed. + + Lemma scaled_index_with_size: + forall r1 r2, + Zpos (index r1) < Zpos (index r2) -> + Zpos (index r1) * 4 + AST.typesize (mreg_type r1) <= Zpos (index r2) * 4. + Proof. + intros. + generalize (scaled_index_with_size_aux r1 r2); intro. + assert (AST.typesize (mreg_type r1) <= 8) by (destruct (mreg_type r1); simpl; omega). + omega. + Qed. End IndexedMreg. Definition is_stack_reg (r: mreg) : bool := false. @@ -215,7 +240,7 @@ Definition destroyed_by_builtin (ef: external_function): list mreg := | _ => nil end. -Definition destroyed_by_setstack (ty: typ): list mreg := +Definition destroyed_by_setstack (q: quantity): list mreg := nil. Definition destroyed_at_function_entry: list mreg := diff --git a/powerpc/Op.v b/powerpc/Op.v index e6f942c1fd..d3bd30ac9e 100644 --- a/powerpc/Op.v +++ b/powerpc/Op.v @@ -602,6 +602,17 @@ Proof. intros; discriminate. Qed. +Lemma is_not_move_operation: + forall (F V A: Type) (genv: Genv.t F V) (sp: val) + (op: operation) (f: A -> val) (args: list A) (m: mem) (v: val), + eval_operation genv sp op (map f args) m = Some v -> + is_move_operation op args = None -> + op <> Omove. +Proof. + intros. destruct (eq_operation op Omove); auto. + subst. destruct args; try destruct args; simpl in *; congruence. +Qed. + (** [negate_condition cond] returns a condition that is logically equivalent to the negation of [cond]. *) diff --git a/powerpc/TargetPrinter.ml b/powerpc/TargetPrinter.ml index 3e11406c9a..26c5118f5b 100644 --- a/powerpc/TargetPrinter.ml +++ b/powerpc/TargetPrinter.ml @@ -626,9 +626,9 @@ module Target (System : SYSTEM):TARGET = fprintf oc " lfd %a, %a(%a)\n" freg r1 constant c ireg r2 | Plfdx(r1, r2, r3) | Plfdx_a(r1, r2, r3) -> fprintf oc " lfdx %a, %a, %a\n" freg r1 ireg r2 ireg r3 - | Plfs(r1, c, r2) -> + | Plfs(r1, c, r2) | Plfs_a(r1, c, r2) -> fprintf oc " lfs %a, %a(%a)\n" freg r1 constant c ireg r2 - | Plfsx(r1, r2, r3) -> + | Plfsx(r1, r2, r3) | Plfsx_a(r1, r2, r3) -> fprintf oc " lfsx %a, %a, %a\n" freg r1 ireg r2 ireg r3 | Plha(r1, c, r2) -> fprintf oc " lha %a, %a(%a)\n" ireg r1 constant c ireg r2 @@ -786,9 +786,9 @@ module Target (System : SYSTEM):TARGET = fprintf oc " stfdu %a, %a(%a)\n" freg r1 constant c ireg r2 | Pstfdx(r1, r2, r3) | Pstfdx_a(r1, r2, r3) -> fprintf oc " stfdx %a, %a, %a\n" freg r1 ireg r2 ireg r3 - | Pstfs(r1, c, r2) -> + | Pstfs(r1, c, r2) | Pstfs_a(r1, c, r2) -> fprintf oc " stfs %a, %a(%a)\n" freg r1 constant c ireg r2 - | Pstfsx(r1, r2, r3) -> + | Pstfsx(r1, r2, r3) | Pstfsx_a(r1, r2, r3) -> fprintf oc " stfsx %a, %a, %a\n" freg r1 ireg r2 ireg r3 | Psth(r1, c, r2) -> fprintf oc " sth %a, %a(%a)\n" ireg r1 constant c ireg r2 diff --git a/riscV/Asm.v b/riscV/Asm.v index 6d223c1d1b..20d30ae09a 100644 --- a/riscV/Asm.v +++ b/riscV/Asm.v @@ -70,7 +70,36 @@ Proof. decide equality. apply ireg_eq. Defined. Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}. Proof. decide equality. Defined. - + +Definition ireg_index (i: ireg): Z := + 2 + 2 * match i with + | X1 => 1 | X2 => 2 | X3 => 3 | X4 => 4 | X5 => 5 + | X6 => 6 | X7 => 7 | X8 => 8 | X9 => 9 | X10 => 10 + | X11 => 11 | X12 => 12 | X13 => 13 | X14 => 14 | X15 => 15 + | X16 => 16 | X17 => 17 | X18 => 18 | X19 => 19 | X20 => 20 + | X21 => 21 | X22 => 22 | X23 => 23 | X24 => 24 | X25 => 25 + | X26 => 26 | X27 => 27 | X28 => 28 | X29 => 29 | X30 => 30 + | X31 => 31 + end. + +Definition ireg0_index (i: ireg0): Z := + match i with + | X0 => 1 + | X i => ireg_index i + end. + +Definition freg_index (f: freg): Z := + 66 + 2 * match f with + | F0 => 0 | F1 => 1 | F2 => 2 | F3 => 3 + | F4 => 4 | F5 => 5 | F6 => 6 | F7 => 7 + | F8 => 8 | F9 => 9 | F10 => 10 | F11 => 11 + | F12 => 12 | F13 => 13 | F14 => 14 | F15 => 15 + | F16 => 16 | F17 => 17 | F18 => 18 | F19 => 19 + | F20 => 20 | F21 => 21 | F22 => 22 | F23 => 23 + | F24 => 24 | F25 => 25 | F26 => 26 | F27 => 27 + | F28 => 28 | F29 => 29 | F30 => 30 | F31 => 31 + end. + (** We model the following registers of the RISC-V architecture. *) Inductive preg: Type := @@ -84,12 +113,136 @@ Coercion FR: freg >-> preg. Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}. Proof. decide equality. apply ireg_eq. apply freg_eq. Defined. -Module PregEq. - Definition t := preg. - Definition eq := preg_eq. -End PregEq. +Definition preg_index (p: preg): Z := + match p with + | IR i => ireg_index i + | FR f => freg_index f + | PC => 64 + 64 + 2 + 1 + end. + +Lemma preg_index_bounds: + forall p: preg, + match p with + | IR _ => 0 < preg_index p /\ preg_index p <= 64 + | FR _ => 65 < preg_index p /\ preg_index p <= 128 + | _ => 130 < preg_index p + end. +Proof. + destruct p; unfold preg_index; try omega. + destruct i; unfold ireg_index; omega. + destruct f; unfold freg_index; omega. +Qed. + +Module Preg <: REGISTER_MODEL. + + Definition reg := preg. + Definition eq_dec := preg_eq. + + Definition type pr := + match pr with + | IR _ | PC => if Archi.ptr64 then Tany64 else Tany32 + | FR _ => Tany64 + end. + + Definition quantity_of pr := + match pr with + | IR _ | PC => if Archi.ptr64 then Q64 else Q32 + | FR _ => Q64 + end. + + Definition chunk_of pr := + match pr with + | IR _ | PC => if Archi.ptr64 then Many64 else Many32 + | FR _ => Many64 + end. + + Lemma type_cases: forall r, type r = Tany32 \/ type r = Tany64. + Proof. + destruct r; simpl; auto; destruct Archi.ptr64; auto. + Qed. + + Definition ofs (r: preg): Z := + preg_index r. + + (* A register's address: The index of its first byte. *) + Definition addr (r: preg): Z := + preg_index r * 4. + + (* The address one byte past the end of register [r]. The next nonoverlapping + register may start here. *) + Definition next_addr (r: preg): Z := addr r + AST.typesize (type r). + + Remark addr_pos: forall p, addr p > 0. + Proof. + intros. unfold addr. destruct p as [x|x|]; try destruct x; simpl; omega. + Qed. + + Lemma addr_compat: forall p, FragBlock.addr (ofs p) = addr p. + Proof. + reflexivity. + Qed. + + Lemma size_compat: + forall p, + AST.typesize (type p) = FragBlock.quantity_size (quantity_of p). + Proof. + intros. unfold type, quantity_of. + destruct p; auto; destruct Archi.ptr64; auto. + Qed. + + Lemma next_addr_compat: forall p, FragBlock.next_addr (ofs p) (quantity_of p) = next_addr p. + Proof. + unfold next_addr, addr, ofs, FragBlock.next_addr, FragBlock.addr; intros. + rewrite size_compat. auto. + Qed. + + Lemma quantity_of_compat: + forall p, + quantity_of p = quantity_of_typ (type p). + Proof. + destruct p; simpl; auto; destruct Archi.ptr64; auto. + Qed. + + Lemma chunk_of_reg_compat: + forall p, + chunk_of p = chunk_of_quantity (quantity_of p). + Proof. + destruct p; simpl; auto; destruct Archi.ptr64; auto. + Qed. + + Lemma chunk_of_reg_type: + forall p, + chunk_of p = chunk_of_type (type p). + Proof. + destruct p; simpl; auto; destruct Archi.ptr64; auto. + Qed. + + Lemma diff_outside_interval: + forall r s, r <> s -> next_addr r <= addr s \/ next_addr s <= addr r. + Proof. + intros. + unfold next_addr, addr. + assert (TS: forall p, AST.typesize (type p) = 4 \/ AST.typesize (type p) = 8). + { intro p. destruct (type_cases p) as [T | T]; rewrite T; auto. } + generalize (preg_index_bounds r), (preg_index_bounds s); intros RB SB. + destruct r eqn:R, s eqn:S; + try congruence; + try (destruct (TS r), (TS s); subst; omega); + simpl AST.typesize. + - (* two iregs *) + destruct Archi.ptr64; simpl; destruct i, i0; simpl; try omega; congruence. + - (* two fregs *) + destruct f, f0; simpl; try omega; congruence. + Qed. + +End Preg. + +Lemma pc_type: subtype Tptr (Preg.type PC) = true. +Proof. + unfold Tptr. simpl; destruct Archi.ptr64; auto. +Qed. -Module Pregmap := EMap(PregEq). +Module Pregmap := Regfile(Preg). (** Conventional names for stack pointer ([SP]) and return address ([RA]). *) @@ -260,7 +413,9 @@ Inductive instruction : Type := (* 32-bit (single-precision) floating point *) | Pfls (rd: freg) (ra: ireg) (ofs: offset) (**r load float *) + | Pfls_a (rd: freg) (ra: ireg) (ofs: offset) (**r load any32 *) | Pfss (rs: freg) (ra: ireg) (ofs: offset) (**r store float *) + | Pfss_a (rs: freg) (ra: ireg) (ofs: offset) (**r store any32 *) | Pfnegs (rd: freg) (rs: freg) (**r negation *) | Pfabss (rd: freg) (rs: freg) (**r absolute value *) @@ -435,26 +590,28 @@ Definition program := AST.program fundef unit. type [Tint] or [Tlong] (in 64 bit mode), and float registers to values of type [Tsingle] or [Tfloat]. *) -Definition regset := Pregmap.t val. +Definition regset := Pregmap.t. Definition genv := Genv.t fundef unit. Definition get0w (rs: regset) (r: ireg0) : val := match r with | X0 => Vint Int.zero - | X r => rs r + | X r => Pregmap.get r rs end. Definition get0l (rs: regset) (r: ireg0) : val := match r with | X0 => Vlong Int64.zero - | X r => rs r + | X r => Pregmap.get r rs end. -Notation "a # b" := (a b) (at level 1, only parsing) : asm. +Notation "a # b" := (Pregmap.get b a) (at level 1, only parsing) : asm. Notation "a ## b" := (get0w a b) (at level 1) : asm. Notation "a ### b" := (get0l a b) (at level 1) : asm. Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm. +Definition pregmap_read rs := fun r => Pregmap.get r rs. + Open Scope asm. (** Undefining some registers *) @@ -462,9 +619,26 @@ Open Scope asm. Fixpoint undef_regs (l: list preg) (rs: regset) : regset := match l with | nil => rs - | r :: l' => undef_regs l' (rs#r <- Vundef) + | r :: l' => (undef_regs l' rs)#r <- Vundef end. +Lemma undef_regs_other: + forall r rl rs, ~In r rl -> (undef_regs rl rs) # r = rs # r. +Proof. + induction rl; simpl; intros. auto. rewrite Pregmap.gso. apply IHrl. intuition. intuition. +Qed. + +Lemma undef_regs_same: + forall r rl rs, In r rl -> (undef_regs rl rs) # r = Vundef. +Proof. + induction rl; simpl; intros. tauto. + destruct H. + - subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + - destruct (Preg.eq_dec r a). + + subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + + rewrite Pregmap.gso; auto. +Qed. + (** Assigning a register pair *) Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := @@ -483,11 +657,11 @@ Fixpoint set_regs (rl: list preg) (vl: list val) (rs: regset) : regset := (** Assigning the result of a builtin *) -Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := +Definition set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := match res with | BR r => rs#r <- v | BR_none => rs - | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) + | BR_splitlong hi lo => rs #hi <- (Val.hiword v) #lo <- (Val.loword v) end. Section RELSEM. @@ -541,6 +715,12 @@ Axiom low_high_half: forall id ofs, Val.offset_ptr (high_half ge id ofs) (low_half ge id ofs) = Genv.symbol_address ge id ofs. +(** The high part fits into one word. *) + +Axiom high_half_type: + forall id ofs, + Val.has_type (high_half ge id ofs) Tany32. + (** The semantics is purely small-step and defined as a function from the current state (a register set + a memory state) to either [Next rs' m'] where [rs'] and [m'] are the updated register @@ -577,14 +757,14 @@ Definition eval_offset (ofs: offset) : ptrofs := Definition exec_load (chunk: memory_chunk) (rs: regset) (m: mem) (d: preg) (a: ireg) (ofs: offset) := - match Mem.loadv chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) with + match Mem.loadv chunk m (Val.offset_ptr (rs#a) (eval_offset ofs)) with | None => Stuck | Some v => Next (nextinstr (rs#d <- v)) m end. Definition exec_store (chunk: memory_chunk) (rs: regset) (m: mem) (s: preg) (a: ireg) (ofs: offset) := - match Mem.storev chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) (rs s) with + match Mem.storev chunk m (Val.offset_ptr (rs#a) (eval_offset ofs)) (rs#s) with | None => Stuck | Some m' => Next (nextinstr rs) m' end. @@ -824,8 +1004,12 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out (** 32-bit (single-precision) floating point *) | Pfls d a ofs => exec_load Mfloat32 rs m d a ofs + | Pfls_a d a ofs => + exec_load Many32 rs m d a ofs | Pfss s a ofs => exec_store Mfloat32 rs m s a ofs + | Pfss_a s a ofs => + exec_store Many32 rs m s a ofs | Pfnegs d s => Next (nextinstr (rs#d <- (Val.negfs rs#s))) m @@ -922,15 +1106,15 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pallocframe sz pos => let (m1, stk) := Mem.alloc m 0 sz in let sp := (Vptr stk Ptrofs.zero) in - match Mem.storev Mptr m1 (Val.offset_ptr sp pos) rs#SP with + match Mem.storev Mptr_any m1 (Val.offset_ptr sp pos) rs#SP with | None => Stuck - | Some m2 => Next (nextinstr (rs #X30 <- (rs SP) #SP <- sp #X31 <- Vundef)) m2 + | Some m2 => Next (nextinstr (rs #X30 <- (rs#SP) #SP <- sp #X31 <- Vundef)) m2 end | Pfreeframe sz pos => - match Mem.loadv Mptr m (Val.offset_ptr rs#SP pos) with + match Mem.loadv Mptr_any m (Val.offset_ptr rs#SP pos) with | None => Stuck | Some v => - match rs SP with + match rs#SP with | Vptr stk ofs => match Mem.free m stk 0 sz with | None => Stuck @@ -952,7 +1136,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Ploadsi rd f => Next (nextinstr (rs#X31 <- Vundef #rd <- (Vsingle f))) m | Pbtbl r tbl => - match rs r with + match rs#r with | Vint n => match list_nth_z tbl (Int.unsigned n) with | None => Stuck @@ -1015,25 +1199,54 @@ Definition preg_of (r: mreg) : preg := (** Undefine all registers except SP and callee-save registers *) -Definition undef_caller_save_regs (rs: regset) : regset := +Definition undef_caller_save_regs_spec (rs: preg -> val) : preg -> val := fun r => if preg_eq r SP || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) then rs r else Vundef. +Definition pregs_destroyed_at_call := + PC :: IR X3 :: IR X4 :: IR X31 :: IR RA + :: (map preg_of (filter (fun m => negb (is_callee_save m)) all_mregs)). + +Lemma pregs_destroyed_at_call_correct: + forall r, + preg_eq r SP + || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) + = + negb (In_dec preg_eq r pregs_destroyed_at_call). +Proof. + intros. + destruct r as [x|x|]; try destruct x; auto. +Qed. + +Definition undef_caller_save_regs (rs: regset) : regset := + undef_regs pregs_destroyed_at_call rs. + +Lemma undef_caller_save_regs_correct: + forall rs r, + (undef_caller_save_regs rs) # r = undef_caller_save_regs_spec (fun r' => rs # r) r. +Proof. + intros. unfold undef_caller_save_regs, undef_caller_save_regs_spec. + rewrite pregs_destroyed_at_call_correct. + destruct (In_dec preg_eq r pregs_destroyed_at_call) as [IN | NOTIN]. + - rewrite undef_regs_same; auto. + - rewrite undef_regs_other; auto. +Qed. + (** Extract the values of the arguments of an external call. We exploit the calling conventions from module [Conventions], except that we use RISC-V registers instead of locations. *) Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop := | extcall_arg_reg: forall r, - extcall_arg rs m (R r) (rs (preg_of r)) - | extcall_arg_stack: forall ofs ty bofs v, + extcall_arg rs m (R r) (rs#(preg_of r)) + | extcall_arg_stack: forall ofs q bofs v, bofs = Stacklayout.fe_ofs_arg + 4 * ofs -> - Mem.loadv (chunk_of_type ty) m + Mem.loadv (chunk_of_quantity q) m (Val.offset_ptr rs#SP (Ptrofs.repr bofs)) = Some v -> - extcall_arg rs m (S Outgoing ofs ty) v. + extcall_arg rs m (S Outgoing ofs q) v. Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop := | extcall_arg_one: forall l v, @@ -1059,17 +1272,17 @@ Inductive state: Type := Inductive step: state -> trace -> state -> Prop := | exec_step_internal: forall b ofs f i rs m rs' m', - rs PC = Vptr b ofs -> + rs#PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) (fn_code f) = Some i -> exec_instr f i rs m = Next rs' m' -> step (State rs m) E0 (State rs' m') | exec_step_builtin: forall b ofs f ef args res rs m vargs t vres rs' m', - rs PC = Vptr b ofs -> + rs#PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) -> - eval_builtin_args ge rs (rs SP) m args vargs -> + eval_builtin_args ge (fun r => rs#r) (rs#SP) m args vargs -> external_call ef ge vargs m t vres m' -> rs' = nextinstr (set_res res vres @@ -1078,11 +1291,11 @@ Inductive step: state -> trace -> state -> Prop := step (State rs m) t (State rs' m') | exec_step_external: forall b ef args res rs m t rs' m', - rs PC = Vptr b Ptrofs.zero -> + rs#PC = Vptr b Ptrofs.zero -> Genv.find_funct_ptr ge b = Some (External ef) -> external_call ef ge args m t res m' -> extcall_arguments rs m (ef_sig ef) args -> - rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs RA) -> + rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs#RA) -> step (State rs m) t (State rs' m'). End RELSEM. @@ -1093,7 +1306,7 @@ Inductive initial_state (p: program): state -> Prop := | initial_state_intro: forall m0, let ge := Genv.globalenv p in let rs0 := - (Pregmap.init Vundef) + Pregmap.init # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero) # SP <- Vnullptr # RA <- Vnullptr in @@ -1102,8 +1315,8 @@ Inductive initial_state (p: program): state -> Prop := Inductive final_state: state -> int -> Prop := | final_state_intro: forall rs m r, - rs PC = Vnullptr -> - rs X10 = Vint r -> + rs#PC = Vnullptr -> + rs#X10 = Vint r -> final_state (State rs m) r. Definition semantics (p: program) := diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index 0a7f8a8a54..90f43d0c9c 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -224,7 +224,7 @@ let expand_builtin_vload_common chunk base ofs res = emit (Plw (res, base, Ofsimm ofs)) | Mint64, BR(IR res) -> emit (Pld (res, base, Ofsimm ofs)) - | Mint64, BR_splitlong(BR(IR res1), BR(IR res2)) -> + | Mint64, BR_splitlong(IR res1, IR res2) -> let ofs' = Ptrofs.add ofs _4 in if base <> res2 then begin emit (Plw (res2, base, Ofsimm ofs)); @@ -414,7 +414,7 @@ let expand_builtin_inline name args res = | "__builtin_bswap64", [BA(IR a1)], BR(IR res) -> expand_bswap64 res a1 | "__builtin_bswap64", [BA_splitlong(BA(IR ah), BA(IR al))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> assert (ah = X6 && al = X5 && rh = X5 && rl = X6); expand_bswap32 X5 X5; expand_bswap32 X6 X6 @@ -437,7 +437,7 @@ let expand_builtin_inline name args res = emit (Pfmaxd(res, a1, a2)) (* 64-bit integer arithmetic for a 32-bit platform *) | "__builtin_negl", [BA_splitlong(BA(IR ah), BA(IR al))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah) rl (fun rl -> emit (Psltuw (X1, X0, X al)); @@ -446,7 +446,7 @@ let expand_builtin_inline name args res = emit (Psubw (rh, X rh, X X1))) | "__builtin_addl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = bl || rl = ah || rl = bh) rl (fun rl -> emit (Paddw (rl, X al, X bl)); @@ -455,7 +455,7 @@ let expand_builtin_inline name args res = emit (Paddw (rh, X rh, X X1))) | "__builtin_subl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = ah || rl = bh) rl (fun rl -> emit (Psltuw (X1, X al, X bl)); @@ -463,7 +463,7 @@ let expand_builtin_inline name args res = emit (Psubw (rh, X ah, X bh)); emit (Psubw (rh, X rh, X X1))) | "__builtin_mull", [BA(IR a); BA(IR b)], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> expand_int64_arith (rl = a || rl = b) rl (fun rl -> emit (Pmulw (rl, X a, X b)); diff --git a/riscV/Asmgen.v b/riscV/Asmgen.v index a704ed7495..4d285d367a 100644 --- a/riscV/Asmgen.v +++ b/riscV/Asmgen.v @@ -358,15 +358,19 @@ Definition transl_cond_op do r1 <- ireg_of a1; OK (transl_condimm_int32u c rd r1 n k) | Ccompl c, a1 :: a2 :: nil => + assertion Archi.ptr64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (transl_cond_int64s c rd r1 r2 k) | Ccomplu c, a1 :: a2 :: nil => + assertion Archi.ptr64; do r1 <- ireg_of a1; do r2 <- ireg_of a2; OK (transl_cond_int64u c rd r1 r2 k) | Ccomplimm c n, a1 :: nil => + assertion Archi.ptr64; do r1 <- ireg_of a1; OK (transl_condimm_int64s c rd r1 n k) | Ccompluimm c n, a1 :: nil => + assertion Archi.ptr64; do r1 <- ireg_of a1; OK (transl_condimm_int64u c rd r1 n k) | Ccompf c, f1 :: f2 :: nil => @@ -405,6 +409,7 @@ Definition transl_op do rd <- ireg_of res; OK (loadimm32 rd n k) | Olongconst n, nil => + assertion Archi.ptr64; do rd <- ireg_of res; OK (loadimm64 rd n k) | Ofloatconst f, nil => @@ -514,83 +519,109 @@ Definition transl_op do rd <- ireg_of res; do rs <- ireg_of a1; OK (Pcvtl2w rd rs :: k) | Ocast32signed, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; assertion (ireg_eq rd rs); OK (Pcvtw2l rd :: k) | Ocast32unsigned, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; assertion (ireg_eq rd rs); OK (Pcvtw2l rd :: Psllil rd rd (Int.repr 32) :: Psrlil rd rd (Int.repr 32) :: k) | Oaddl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Paddl rd rs1 rs2 :: k) | Oaddlimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (addimm64 rd rs n k) | Onegl, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (Psubl rd X0 rs :: k) | Osubl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Psubl rd rs1 rs2 :: k) | Omull, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pmull rd rs1 rs2 :: k) | Omullhs, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pmulhl rd rs1 rs2 :: k) | Omullhu, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pmulhul rd rs1 rs2 :: k) | Odivl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pdivl rd rs1 rs2 :: k) | Odivlu, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pdivul rd rs1 rs2 :: k) | Omodl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Preml rd rs1 rs2 :: k) | Omodlu, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Premul rd rs1 rs2 :: k) | Oandl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pandl rd rs1 rs2 :: k) | Oandlimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (andimm64 rd rs n k) | Oorl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Porl rd rs1 rs2 :: k) | Oorlimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (orimm64 rd rs n k) | Oxorl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pxorl rd rs1 rs2 :: k) | Oxorlimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (xorimm64 rd rs n k) | Oshll, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pslll rd rs1 rs2 :: k) | Oshllimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (Psllil rd rs n :: k) | Oshrl, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Psral rd rs1 rs2 :: k) | Oshrlimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (Psrail rd rs n :: k) | Oshrlu, a1 :: a2 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Psrll rd rs1 rs2 :: k) | Oshrluimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (Psrlil rd rs n :: k) | Oshrxlimm n, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- ireg_of a1; OK (if Int.eq n Int.zero then Pmv rd rs :: k else Psrail X31 rs (Int.repr 63) :: @@ -669,27 +700,35 @@ Definition transl_op OK (Pfcvtswu rd rs :: k) | Olongoffloat, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfcvtld rd rs :: k) | Olonguoffloat, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfcvtlud rd rs :: k) | Ofloatoflong, a1 :: nil => + assertion Archi.ptr64; do rd <- freg_of res; do rs <- ireg_of a1; OK (Pfcvtdl rd rs :: k) | Ofloatoflongu, a1 :: nil => + assertion Archi.ptr64; do rd <- freg_of res; do rs <- ireg_of a1; OK (Pfcvtdlu rd rs :: k) | Olongofsingle, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfcvtls rd rs :: k) | Olonguofsingle, a1 :: nil => + assertion Archi.ptr64; do rd <- ireg_of res; do rs <- freg_of a1; OK (Pfcvtlus rd rs :: k) | Osingleoflong, a1 :: nil => + assertion Archi.ptr64; do rd <- freg_of res; do rs <- ireg_of a1; OK (Pfcvtsl rd rs :: k) | Osingleoflongu, a1 :: nil => + assertion Archi.ptr64; do rd <- freg_of res; do rs <- ireg_of a1; OK (Pfcvtslu rd rs :: k) @@ -723,35 +762,29 @@ Definition indexed_memory_access Pluiw X31 hi :: Paddw X31 base X31 :: mk_instr X31 (Ofsimm (Ptrofs.of_int lo)) :: k end. -Definition loadind (base: ireg) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) := - match ty, preg_of dst with - | Tint, IR rd => OK (indexed_memory_access (Plw rd) base ofs k) - | Tlong, IR rd => OK (indexed_memory_access (Pld rd) base ofs k) - | Tsingle, FR rd => OK (indexed_memory_access (Pfls rd) base ofs k) - | Tfloat, FR rd => OK (indexed_memory_access (Pfld rd) base ofs k) - | Tany32, IR rd => OK (indexed_memory_access (Plw_a rd) base ofs k) - | Tany64, IR rd => OK (indexed_memory_access (Pld_a rd) base ofs k) - | Tany64, FR rd => OK (indexed_memory_access (Pfld_a rd) base ofs k) - | _, _ => Error (msg "Asmgen.loadind") +Definition loadind (base: ireg) (ofs: ptrofs) (q: quantity) (dst: mreg) (k: code) := + match q, preg_of dst, Archi.ptr64 with + | Q32, IR rd, _ => OK (indexed_memory_access (Plw_a rd) base ofs k) + | Q64, IR rd, true => OK (indexed_memory_access (Pld_a rd) base ofs k) + | Q32, FR rd, _ => OK (indexed_memory_access (Pfls_a rd) base ofs k) + | Q64, FR rd, _ => OK (indexed_memory_access (Pfld_a rd) base ofs k) + | _, _, _ => Error (msg "Asmgen.loadind") end. -Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) := - match ty, preg_of src with - | Tint, IR rd => OK (indexed_memory_access (Psw rd) base ofs k) - | Tlong, IR rd => OK (indexed_memory_access (Psd rd) base ofs k) - | Tsingle, FR rd => OK (indexed_memory_access (Pfss rd) base ofs k) - | Tfloat, FR rd => OK (indexed_memory_access (Pfsd rd) base ofs k) - | Tany32, IR rd => OK (indexed_memory_access (Psw_a rd) base ofs k) - | Tany64, IR rd => OK (indexed_memory_access (Psd_a rd) base ofs k) - | Tany64, FR rd => OK (indexed_memory_access (Pfsd_a rd) base ofs k) - | _, _ => Error (msg "Asmgen.storeind") +Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (q: quantity) (k: code) := + match q, preg_of src with + | Q32, IR rd => OK (indexed_memory_access (Psw_a rd) base ofs k) + | Q64, IR rd => OK (indexed_memory_access (Psd_a rd) base ofs k) + | Q32, FR rd => OK (indexed_memory_access (Pfss_a rd) base ofs k) + | Q64, FR rd => OK (indexed_memory_access (Pfsd_a rd) base ofs k) + | _, _ => Error (msg "Asmgen.storeind") end. Definition loadind_ptr (base: ireg) (ofs: ptrofs) (dst: ireg) (k: code) := - indexed_memory_access (if Archi.ptr64 then Pld dst else Plw dst) base ofs k. + indexed_memory_access (if Archi.ptr64 then Pld_a dst else Plw_a dst) base ofs k. Definition storeind_ptr (src: ireg) (base: ireg) (ofs: ptrofs) (k: code) := - indexed_memory_access (if Archi.ptr64 then Psd src else Psw src) base ofs k. + indexed_memory_access (if Archi.ptr64 then Psd_a src else Psw_a src) base ofs k. (** Translation of memory accesses: loads, and stores. *) @@ -789,6 +822,7 @@ Definition transl_load (chunk: memory_chunk) (addr: addressing) do r <- ireg_of dst; transl_memory_access (Plw r) addr args k | Mint64 => + assertion Archi.ptr64; do r <- ireg_of dst; transl_memory_access (Pld r) addr args k | Mfloat32 => diff --git a/riscV/Asmgenproof.v b/riscV/Asmgenproof.v index 5ec578861a..754ec2c5d1 100644 --- a/riscV/Asmgenproof.v +++ b/riscV/Asmgenproof.v @@ -72,7 +72,7 @@ Qed. Lemma exec_straight_exec: forall fb f c ep tf tc c' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs#PC) fb f c ep tf tc -> exec_straight tge tf tc rs m c' rs' m' -> plus step tge (State rs m) E0 (State rs' m'). Proof. @@ -84,10 +84,10 @@ Qed. Lemma exec_straight_at: forall fb f c ep tf tc c' ep' tc' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs#PC) fb f c ep tf tc -> transl_code f c' ep' = OK tc' -> exec_straight tge tf tc rs m tc' rs' m' -> - transl_code_at_pc ge (rs' PC) fb f c' ep' tf tc'. + transl_code_at_pc ge (rs'#PC) fb f c' ep' tf tc'. Proof. intros. inv H. exploit exec_straight_steps_2; eauto. @@ -310,7 +310,7 @@ Remark loadind_label: loadind base ofs ty dst k = OK c -> tail_nolabel k c. Proof. unfold loadind; intros. - destruct ty, (preg_of dst); inv H; apply indexed_memory_access_label; intros; exact I. + destruct ty, (preg_of dst), Archi.ptr64; inv H; apply indexed_memory_access_label; intros; exact I. Qed. Remark storeind_label: @@ -417,11 +417,11 @@ Lemma find_label_goto_label: forall f tf lbl rs m c' b ofs, Genv.find_funct_ptr ge b = Some (Internal f) -> transf_function f = OK tf -> - rs PC = Vptr b ofs -> + rs#PC = Vptr b ofs -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> exists tc', exists rs', goto_label tf lbl rs m = Next rs' m - /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc' + /\ transl_code_at_pc ge (rs'#PC) b f c' false tf tc' /\ forall r, r <> PC -> rs'#r = rs#r. Proof. intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2. @@ -430,7 +430,9 @@ Proof. intros [pos' [P [Q R]]]. exists tc; exists (rs#PC <- (Vptr b (Ptrofs.repr pos'))). split. unfold goto_label. rewrite P. rewrite H1. auto. - split. rewrite Pregmap.gss. constructor; auto. + split. rewrite Pregmap.gss. + rewrite Preg.chunk_of_reg_type, Val.load_result_same by (simpl; destruct Archi.ptr64; auto). + constructor; auto. rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q. auto. omega. generalize (transf_function_no_overflow _ _ H0). omega. @@ -478,7 +480,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (FIND: Genv.find_funct_ptr ge fb = Some (Internal f)) (MEXT: Mem.extends m m') - (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc) + (AT: transl_code_at_pc ge (rs#PC) fb f c ep tf tc) (AG: agree ms sp rs) (DXP: ep = true -> rs#X30 = parent_sp s), match_states (Mach.State s fb sp c ms m) @@ -488,8 +490,8 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = Vptr fb Ptrofs.zero) - (ATLR: rs RA = parent_ra s), + (ATPC: rs#PC = Vptr fb Ptrofs.zero) + (ATLR: rs#RA = parent_ra s), match_states (Mach.Callstate s fb ms m) (Asm.State rs m') | match_states_return: @@ -497,7 +499,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = parent_ra s), + (ATPC: rs#PC = parent_ra s), match_states (Mach.Returnstate s ms m) (Asm.State rs m'). @@ -506,7 +508,7 @@ Lemma exec_straight_steps: match_stack ge s -> Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> (forall k c (TR: transl_instr f i ep k = OK c), exists rs2, exec_straight tge tf c rs1 m1' k rs2 m2' @@ -529,7 +531,7 @@ Lemma exec_straight_steps_goto: Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> it1_is_parent ep i = false -> (forall k c (TR: transl_instr f i ep k = OK c), exists jmp, exists k', exists rs2, @@ -566,7 +568,7 @@ Lemma exec_straight_opt_steps_goto: Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> it1_is_parent ep i = false -> (forall k c (TR: transl_instr f i ep k = OK c), exists jmp, exists k', exists rs2, @@ -640,7 +642,7 @@ Proof. - (* Mlabel *) left; eapply exec_straight_steps; eauto; intros. - monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto. + monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. Simpl. split. apply agree_nextinstr; auto. simpl; congruence. - (* Mgetstack *) @@ -655,7 +657,7 @@ Proof. - (* Msetstack *) unfold store_stack in H. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (regmap_get src rs) (rs0#(preg_of src))). eapply preg_val; eauto. 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. @@ -684,21 +686,26 @@ Opaque loadind. simpl; intros. rewrite R; auto with asmgen. apply preg_of_not_X30; auto. (* GPR11 does not contain parent *) - rewrite chunk_of_Tptr in A. - exploit loadind_ptr_correct. eexact A. congruence. intros [rs1 [P [Q R]]]. + rewrite chunk_of_quantity_of_Tptr in A. + exploit loadind_ptr_correct. eexact A. congruence. + eapply Val.has_subtype; eauto using parent_sp_type. + simpl. unfold Tptr. destruct Archi.ptr64; auto. + intros [rs1 [P [Q R]]]. exploit loadind_correct. eexact EQ. instantiate (2 := rs1). rewrite Q. eauto. congruence. intros [rs2 [S [T U]]]. exists rs2; split. eapply exec_straight_trans; eauto. split. eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto. instantiate (1 := rs1#X30 <- (rs2#X30)). intros. rewrite Pregmap.gso; auto with asmgen. - congruence. - intros. unfold Pregmap.set. destruct (PregEq.eq r' X30). congruence. auto with asmgen. + congruence. intros. + destruct (Preg.eq_dec r' X30). subst; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto using Pregmap.get_has_type. + rewrite Pregmap.gso; auto with asmgen. simpl; intros. rewrite U; auto with asmgen. apply preg_of_not_X30; auto. - (* Mop *) - assert (eval_operation tge sp op (map rs args) m = Some v). + assert (eval_operation tge sp op (map (fun r => regmap_get r rs) args) m = Some v). rewrite <- H. apply eval_operation_preserved. exact symbols_preserved. exploit eval_operation_lessdef. eapply preg_vals; eauto. eauto. eexact H0. intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A. @@ -713,7 +720,7 @@ Local Transparent destroyed_by_op. destruct op; simpl; auto; congruence. - (* Mload *) - assert (eval_addressing tge sp addr (map rs args) = Some a). + assert (eval_addressing tge sp addr (map (fun r => regmap_get r rs) args) = Some a). rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1. intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A. @@ -726,11 +733,11 @@ Local Transparent destroyed_by_op. simpl; congruence. - (* Mstore *) - assert (eval_addressing tge sp addr (map rs args) = Some a). + assert (eval_addressing tge sp addr (map (fun r => regmap_get r rs) args) = Some a). rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1. intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (regmap_get src rs) (rs0#(preg_of src))). 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]]. @@ -745,10 +752,10 @@ Local Transparent destroyed_by_op. eapply transf_function_no_overflow; eauto. destruct ros as [rf|fid]; simpl in H; monadInv H5. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0#x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto. generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x). @@ -761,8 +768,10 @@ Local Transparent destroyed_by_op. econstructor; eauto. econstructor; eauto. eapply agree_sp_def; eauto. + eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simpl. - Simpl. rewrite <- H2. auto. + Simpl. rewrite H7. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. + Simpl. rewrite <- H2. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. + (* Direct call *) generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x). @@ -775,8 +784,10 @@ Local Transparent destroyed_by_op. econstructor; eauto. econstructor; eauto. eapply agree_sp_def; eauto. + eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simpl. - Simpl. rewrite <- H2. auto. + Simpl. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. + Simpl. rewrite <- H2. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. - (* Mtailcall *) assert (f0 = f) by congruence. subst f0. @@ -785,10 +796,10 @@ Local Transparent destroyed_by_op. eapply transf_function_no_overflow; eauto. exploit Mem.loadv_extends. eauto. eexact H1. auto. simpl. intros [parent' [A B]]. destruct ros as [rf|fid]; simpl in H; monadInv H7. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0#x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H7; intros LD; inv LD; auto. exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z). exploit exec_straight_steps_2; eauto using functions_transl. @@ -802,7 +813,10 @@ Local Transparent destroyed_by_op. (* match states *) econstructor; eauto. apply agree_set_other; auto with asmgen. - Simpl. rewrite Z by (rewrite <- (ireg_of_eq _ _ EQ1); eauto with asmgen). assumption. + Simpl. + rewrite Z by (rewrite <- (ireg_of_eq _ _ EQ1); eauto with asmgen). rewrite H9. + simpl. unfold Val.load_result. destruct Archi.ptr64; auto. + rewrite Z by (rewrite <- (ireg_of_eq _ _ EQ1); eauto with asmgen). Simpl. + (* Direct call *) exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z). exploit exec_straight_steps_2; eauto using functions_transl. @@ -814,14 +828,15 @@ Local Transparent destroyed_by_op. simpl. reflexivity. traceEq. (* match states *) + unfold Genv.symbol_address. rewrite symbols_preserved, H. econstructor; eauto. apply agree_set_other; auto with asmgen. - Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. auto. + Simpl. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. Simpl. - (* Mbuiltin *) - inv AT. monadInv H4. + inv AT. monadInv H5. exploit functions_transl; eauto. intro FN. - generalize (transf_function_no_overflow _ _ H3); intro NOOV. + generalize (transf_function_no_overflow _ _ H4); intro NOOV. exploit builtin_args_match; eauto. intros [vargs' [P Q]]. exploit external_call_mem_extends; eauto. intros [vres' [m2' [A [B [C D]]]]]. @@ -836,12 +851,20 @@ Local Transparent destroyed_by_op. instantiate (2 := tf); instantiate (1 := x). unfold nextinstr. rewrite Pregmap.gss. rewrite set_res_other. rewrite undef_regs_other_2. rewrite Pregmap.gso by congruence. - rewrite <- H1. simpl. econstructor; eauto. + rewrite <- H2. simpl. unfold Val.load_result. destruct Archi.ptr64; econstructor; eauto. + eapply code_tail_next_int; eauto. eapply code_tail_next_int; eauto. rewrite preg_notin_charact. intros. auto with asmgen. auto with asmgen. apply agree_nextinstr. eapply agree_set_res; auto. eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto. apply Pregmap.gso; auto with asmgen. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + destruct res; simpl in H1; InvBooleans; simpl; auto. + eapply Val.has_subtype; eauto. + split. + destruct vres'; simpl; auto; destruct (mreg_type lo); auto; congruence. + destruct vres'; simpl; auto; destruct (mreg_type hi); auto; congruence. congruence. - (* Mgoto *) @@ -914,6 +937,11 @@ Local Transparent destroyed_by_op. (* match states *) econstructor; eauto. apply agree_set_other; auto with asmgen. + Simpl. rewrite X. + exploit parent_ra_type; eauto; intro. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tptr); auto. + simpl; unfold Tptr; destruct Archi.ptr64; auto. - (* internal function *) exploit functions_translated; eauto. intros [tf [A B]]. monadInv B. @@ -934,8 +962,12 @@ Local Transparent destroyed_by_op. set (tf := {| fn_sig := Mach.fn_sig f; fn_code := tfbody |}) in *. set (rs2 := nextinstr (rs0#X30 <- (parent_sp s) #SP <- sp #X31 <- Vundef)). exploit (storeind_ptr_correct tge tf SP (fn_retaddr_ofs f) RA x0 rs2 m2'). - rewrite chunk_of_Tptr in P. change (rs2 X1) with (rs0 X1). rewrite ATLR. - change (rs2 X2) with sp. eexact P. + rewrite chunk_of_quantity_of_Tptr in P. + replace (rs2#X1) with (rs0#X1). rewrite ATLR. + replace (rs2#X2) with sp. eexact P. + unfold rs2; Simpl. + unfold sp. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. + unfold rs2; Simpl. congruence. congruence. intros (rs3 & U & V). assert (EXEC_PROLOGUE: @@ -945,8 +977,8 @@ Local Transparent destroyed_by_op. { change (fn_code tf) with tfbody; unfold tfbody. apply exec_straight_step with rs2 m2'. unfold exec_instr. rewrite C. fold sp. - rewrite <- (sp_val _ _ _ AG). rewrite chunk_of_Tptr in F. rewrite F. reflexivity. - reflexivity. + rewrite <- (sp_val _ _ _ AG). rewrite chunk_of_quantity_of_Tptr in F. rewrite F. reflexivity. + unfold rs2. Simpl. eexact U. } exploit exec_straight_steps_2; eauto using functions_transl. omega. constructor. intros (ofs' & X & Y). @@ -962,7 +994,13 @@ Local Transparent destroyed_by_op. Local Transparent destroyed_at_function_entry. simpl; intros; Simpl. unfold sp; congruence. - intros. rewrite V by auto with asmgen. reflexivity. + unfold sp, Tptr. simpl. destruct Archi.ptr64; auto. + unfold sp; simpl. destruct Archi.ptr64; auto. + intros. rewrite V by auto with asmgen. unfold rs2; Simpl. + exploit agree_sp_type; eauto; intro. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tptr); auto. + simpl. unfold Tptr. destruct Archi.ptr64; auto. - (* external function *) exploit functions_translated; eauto. @@ -977,6 +1015,20 @@ Local Transparent destroyed_at_function_entry. econstructor; eauto. unfold loc_external_result. apply agree_set_other; auto. apply agree_set_pair; auto. apply agree_undef_caller_save_regs; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + unfold Val.has_type_rpair in H3. + unfold loc_result, proj_sig_res in *. + destruct Archi.ptr64, (sig_res (ef_sig ef)). + destruct t0; simpl in *; auto. + simpl in *; auto. + destruct t0; simpl in *; auto. + simpl in *; auto. + rewrite Pregmap.gss. rewrite ATLR. + exploit parent_ra_type; eauto; intro. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tptr); auto. + simpl. unfold Tptr. destruct Archi.ptr64; auto. - (* return *) inv STACKS. simpl in *. @@ -998,8 +1050,12 @@ Proof. econstructor; eauto. constructor. apply Mem.extends_refl. - split. auto. simpl. unfold Vnullptr; destruct Archi.ptr64; congruence. - intros. rewrite Regmap.gi. auto. + split. Simpl. simpl. unfold Val.load_result, Vnullptr; destruct Archi.ptr64; auto. + simpl. unfold Vnullptr; destruct Archi.ptr64; congruence. + simpl. unfold Vnullptr, Tptr; destruct Archi.ptr64; simpl; auto. + intros. unfold regmap_get, Regmap.get. rewrite FragBlock.gi. auto. + Simpl. simpl. unfold Val.load_result. destruct Archi.ptr64; auto. + Simpl. simpl. unfold Val.load_result, Vnullptr. destruct Archi.ptr64; auto. unfold Genv.symbol_address. rewrite (match_program_main TRANSF). rewrite symbols_preserved. diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v index 7f070c120e..86e69db790 100644 --- a/riscV/Asmgenproof1.v +++ b/riscV/Asmgenproof1.v @@ -90,6 +90,47 @@ Qed. Hint Resolve ireg_of_not_X31 ireg_of_not_X31': asmgen. +Lemma data_preg_nextinstr_PC: + forall rs r v, + r <> PC -> + Pregmap.get PC (nextinstr rs # r <- v) = + Val.offset_ptr (Pregmap.get PC rs) Ptrofs.one. +Proof. + intros. + unfold nextinstr. rewrite Pregmap.gss. + rewrite Pregmap.gso by auto. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype; eauto using pc_type, Val.offset_ptr_type. +Qed. + +Lemma data_preg_nextinstr_set2_PC: + forall rs r1 r2 v1 v2, + r1 <> PC -> + r2 <> PC -> + Pregmap.get PC (nextinstr rs # r1 <- v1 # r2 <- v2) = + Val.offset_ptr (Pregmap.get PC rs) Ptrofs.one. +Proof. + intros. + unfold nextinstr. rewrite Pregmap.gss. + rewrite !Pregmap.gso by auto. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype; eauto using pc_type, Val.offset_ptr_type. +Qed. + +Lemma load_result_symbol_address: + forall (ge: genv) id ofs (r: ireg), + Genv.symbol_address ge id ofs = + Val.load_result (Preg.chunk_of r) (Genv.symbol_address ge id ofs). +Proof. + intros. + replace (Preg.chunk_of r) with (chunk_of_type (if Archi.ptr64 then Tany64 else Tany32)). + rewrite Val.load_result_same; auto. unfold Genv.symbol_address. + destruct (Genv.find_symbol ge id); simpl; destruct Archi.ptr64; auto. + simpl; destruct Archi.ptr64; auto. +Qed. + +Hint Resolve data_preg_nextinstr_PC data_preg_nextinstr_set2_PC load_result_symbol_address: asmgen. + (** Useful simplification tactic *) Ltac Simplif := @@ -101,6 +142,43 @@ Ltac Simplif := Ltac Simpl := repeat Simplif. +Ltac invert_load_result := + simpl get0w; + try apply Val.lessdef_same; + match goal with + | [ |- ?v = Val.load_result ?c ?v ] => + symmetry + | _ => + idtac + end; + match goal with + | [ |- Val.load_result ?c ?v = ?v ] => + simpl c; + match goal with + | [ H: Archi.ptr64 = true |- _ ] => + rewrite !H + | _ => + idtac + end + end; + match goal with + | [ |- Val.load_result (if Archi.ptr64 then _ else _) ?v = ?v ] => + destruct Archi.ptr64; invert_load_result + | [ |- Val.load_result ?c (?f ?arg1 ?arg2 ?arg3) = ?f ?arg1 ?arg2 ?arg3 ] => + destruct arg1, arg2, arg3; simpl; auto + | [ |- Val.load_result ?c (?f ?arg1 ?arg2) = ?f ?arg1 ?arg2 ] => + destruct arg1, arg2; simpl; auto; + match goal with + | [ |- match (if ?cond then _ else _) with _ => _ end = if ?cond then _ else _ ] => + destruct cond; auto + | _ => + idtac + end + | [ |- Val.load_result ?c (?f ?arg1) = ?f ?arg1 ] => + destruct arg1; simpl; auto + end; + try congruence. + (** * Correctness of RISC-V constructor functions *) Section CONSTRUCTORS. @@ -120,12 +198,13 @@ Proof. unfold load_hilo32; intros. predSpec Int.eq Int.eq_spec lo Int.zero. - subst lo. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int.add_zero. Simpl. + apply exec_straight_one. simpl; eauto. auto with asmgen. + split. rewrite Int.add_zero. Simpl. invert_load_result. intros; Simpl. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split. Simpl. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. + split. Simpl. + rewrite Preg.chunk_of_reg_type. destruct (Preg.type_cases rd) as [T|T]; rewrite T; auto. intros; Simpl. Qed. @@ -139,8 +218,8 @@ Proof. unfold loadimm32; intros. generalize (make_immed32_sound n); intros E. destruct (make_immed32 n). - subst imm. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int.add_zero_l; Simpl. + apply exec_straight_one. simpl; eauto. auto with asmgen. + split. rewrite Int.add_zero_l; Simpl. invert_load_result. intros; Simpl. - rewrite E. apply load_hilo32_correct. Qed. @@ -157,19 +236,19 @@ Lemma opimm32_correct: r1 <> X31 -> exists rs', exec_straight ge fn (opimm32 op opi rd r1 n k) rs m k rs' m - /\ rs'#rd = sem rs##r1 (Vint n) + /\ rs'#rd = Val.load_result (Preg.chunk_of rd) (sem rs##r1 (Vint n)) /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. Proof. intros. unfold opimm32. generalize (make_immed32_sound n); intros E. destruct (make_immed32 n). - subst imm. econstructor; split. - apply exec_straight_one. rewrite H0. simpl; eauto. auto. + apply exec_straight_one. rewrite H0. simpl; eauto. auto with asmgen. split. Simpl. intros; Simpl. - destruct (load_hilo32_correct X31 hi lo (op rd r1 X31 :: k) rs m) as (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. - rewrite H; eauto. auto. + rewrite H; eauto. auto with asmgen. split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence. intros; Simpl. Qed. @@ -178,6 +257,7 @@ Qed. Lemma load_hilo64_correct: forall rd hi lo k rs m, + Archi.ptr64 = true -> exists rs', exec_straight ge fn (load_hilo64 rd hi lo k) rs m k rs' m /\ rs'#rd = Vlong (Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo) @@ -186,17 +266,18 @@ Proof. unfold load_hilo64; intros. predSpec Int64.eq Int64.eq_spec lo Int64.zero. - subst lo. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int64.add_zero. Simpl. + apply exec_straight_one. simpl; eauto. auto with asmgen. + split. rewrite Int64.add_zero. Simpl. simpl Preg.chunk_of; rewrite H; auto. intros; Simpl. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split. Simpl. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. + split. Simpl. simpl Preg.chunk_of; rewrite H; auto. intros; Simpl. Qed. Lemma loadimm64_correct: forall rd n k rs m, + Archi.ptr64 = true -> exists rs', exec_straight ge fn (loadimm64 rd n k) rs m k rs' m /\ rs'#rd = Vlong n @@ -205,14 +286,14 @@ Proof. unfold loadimm64; intros. generalize (make_immed64_sound n); intros E. destruct (make_immed64 n). - subst imm. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int64.add_zero_l; Simpl. + apply exec_straight_one. simpl; eauto. auto with asmgen. + split. rewrite Int64.add_zero_l; Simpl. simpl Preg.chunk_of; rewrite H; auto. intros; Simpl. - exploit load_hilo64_correct; eauto. intros (rs' & A & B & C). rewrite E. exists rs'; eauto. - subst imm. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. Simpl. + apply exec_straight_one. simpl; eauto. rewrite nextinstr_pc, !Pregmap.gso; auto with asmgen. + split. Simpl. simpl Preg.chunk_of; rewrite H; auto. intros; Simpl. Qed. @@ -225,6 +306,7 @@ Lemma opimm64_correct: (forall d s n rs, exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs###s (Vlong n)))) m) -> forall rd r1 n k rs, + Archi.ptr64 = true -> r1 <> X31 -> exists rs', exec_straight ge fn (opimm64 op opi rd r1 n k) rs m k rs' m @@ -234,18 +316,20 @@ Proof. intros. unfold opimm64. generalize (make_immed64_sound n); intros E. destruct (make_immed64 n). - subst imm. econstructor; split. - apply exec_straight_one. rewrite H0. simpl; eauto. auto. - split. Simpl. intros; Simpl. + apply exec_straight_one. rewrite H0. simpl; eauto. auto with asmgen. + split. Simpl. simpl Preg.chunk_of; rewrite H1; auto. intros; Simpl. - destruct (load_hilo64_correct X31 hi lo (op rd r1 X31 :: k) rs m) as (rs' & A & B & C). + auto. econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. - rewrite H; eauto. auto. - split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence. + rewrite H; eauto. auto with asmgen. + split. Simpl. simpl. rewrite B, C, E, H1. auto. congruence. congruence. intros; Simpl. - subst imm. econstructor; split. - eapply exec_straight_two. simpl; eauto. rewrite H. simpl; eauto. auto. auto. - split. Simpl. intros; Simpl. + eapply exec_straight_two. simpl; eauto. rewrite H. simpl; eauto. + rewrite nextinstr_pc, !Pregmap.gso; auto with asmgen. auto with asmgen. + split. Simpl. simpl Preg.chunk_of; rewrite H1; auto. intros; Simpl. Qed. (** Add offset to pointer *) @@ -261,20 +345,24 @@ Proof. unfold addptrofs; intros. destruct (Ptrofs.eq_dec n Ptrofs.zero). - subst n. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. Simpl. destruct (rs r1); simpl; auto. rewrite Ptrofs.add_zero; auto. + apply exec_straight_one. simpl; eauto. auto with asmgen. + split. Simpl. destruct (rs#r1); unfold Val.offset_ptr; auto. rewrite Ptrofs.add_zero. + rewrite Preg.chunk_of_reg_type, Val.load_result_same. auto. + simpl. destruct Archi.ptr64; auto. intros; Simpl. - destruct Archi.ptr64 eqn:SF. + unfold addimm64. exploit (opimm64_correct Paddl Paddil Val.addl); eauto. intros (rs' & A & B & C). exists rs'; split. eexact A. split; auto. - rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF. + rewrite B. simpl. destruct (rs#r1); simpl; auto. rewrite SF. rewrite Ptrofs.of_int64_to_int64 by auto. auto. + unfold addimm32. exploit (opimm32_correct Paddw Paddiw Val.add); eauto. intros (rs' & A & B & C). exists rs'; split. eexact A. split; auto. - rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF. - rewrite Ptrofs.of_int_to_int by auto. auto. + rewrite B. simpl. destruct (rs#r1); simpl; auto. rewrite SF. + rewrite Ptrofs.of_int_to_int by auto. + change Many32 with (chunk_of_type Tany32). + rewrite Val.load_result_same; auto. Qed. Lemma addptrofs_correct_2: @@ -356,9 +444,9 @@ Lemma transl_cond_float_correct: Proof. intros. destruct cmp; simpl in H; inv H; auto. - rewrite Val.negate_cmpf_eq. auto. -- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpf, Val.cmpf_bool. +- simpl. f_equal. f_equal. f_equal. destruct (rs#r2), (rs#r1); auto. unfold Val.cmpf, Val.cmpf_bool. rewrite <- Float.cmp_swap. auto. -- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpf, Val.cmpf_bool. +- simpl. f_equal. f_equal. f_equal. destruct (rs#r2), (rs#r1); auto. unfold Val.cmpf, Val.cmpf_bool. rewrite <- Float.cmp_swap. auto. Qed. @@ -369,11 +457,11 @@ Lemma transl_cond_single_correct: exec_instr ge fn insn rs m = Next (nextinstr (rs#rd <- v)) m. Proof. intros. destruct cmp; simpl in H; inv H; auto. -- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool. +- simpl. f_equal. f_equal. f_equal. destruct (rs#r2), (rs#r1); auto. unfold Val.cmpfs, Val.cmpfs_bool. rewrite Float32.cmp_ne_eq. destruct (Float32.cmp Ceq f0 f); auto. -- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool. +- simpl. f_equal. f_equal. f_equal. destruct (rs#r2), (rs#r1); auto. unfold Val.cmpfs, Val.cmpfs_bool. rewrite <- Float32.cmp_swap. auto. -- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool. +- simpl. f_equal. f_equal. f_equal. destruct (rs#r2), (rs#r1); auto. unfold Val.cmpfs, Val.cmpfs_bool. rewrite <- Float32.cmp_swap. auto. Qed. @@ -419,7 +507,7 @@ Qed. Lemma transl_cbranch_correct_1: forall cond args lbl k c m ms b sp rs m', transl_cbranch cond args lbl k = OK c -> - eval_condition cond (List.map ms args) m = Some b -> + eval_condition cond (List.map (fun r => Regmap.get r ms) args) m = Some b -> agree ms sp rs -> Mem.extends m m' -> exists rs', exists insn, @@ -428,9 +516,9 @@ Lemma transl_cbranch_correct_1: /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. Proof. intros until m'; intros TRANSL EVAL AG MEXT. - set (vl' := map rs (map preg_of args)). + set (vl' := map (fun r => rs#r) (map preg_of args)). assert (EVAL': eval_condition cond vl' m' = Some b). - { apply eval_condition_lessdef with (map ms args) m; auto. eapply preg_vals; eauto. } + { apply eval_condition_lessdef with (map (fun r => Regmap.get r ms) args) m; auto. eapply preg_vals; eauto. } clear EVAL MEXT AG. destruct cond; simpl in TRANSL; ArgsInv. - exists rs, (transl_cbranch_int32s c0 x x0 lbl). @@ -460,7 +548,14 @@ Proof. - predSpec Int64.eq Int64.eq_spec n Int64.zero. + subst n. exists rs, (transl_cbranch_int64s c0 x X0 lbl). intuition auto. constructor. apply transl_cbranch_int64s_correct; auto. -+ exploit (loadimm64_correct X31 n); eauto. intros (rs' & A & B & C). ++ exploit (loadimm64_correct X31 n); eauto. + assert (Archi.ptr64 = true). + { destruct (rs#x) eqn:V; simpl in EVAL'; try congruence. + assert (Val.has_type (Vlong i) (Preg.type x)) by (rewrite <- V; apply Pregmap.get_has_type). + destruct (Preg.type_cases x) as [T|T]; rewrite T in H0; auto. + simpl in T. destruct Archi.ptr64; congruence. } + auto. + intros (rs' & A & B & C). exists rs', (transl_cbranch_int64s c0 x X31 lbl). split. constructor; eexact A. split; auto. apply transl_cbranch_int64s_correct; auto. @@ -468,7 +563,15 @@ Proof. - predSpec Int64.eq Int64.eq_spec n Int64.zero. + subst n. exists rs, (transl_cbranch_int64u c0 x X0 lbl). intuition auto. constructor. apply transl_cbranch_int64u_correct; auto. -+ exploit (loadimm64_correct X31 n); eauto. intros (rs' & A & B & C). ++ exploit (loadimm64_correct X31 n); eauto. + assert (Archi.ptr64 = true). + { destruct (rs#x) eqn:V; simpl in EVAL'; try congruence. + assert (Val.has_type (Vlong i) (Preg.type x)) by (rewrite <- V; apply Pregmap.get_has_type). + destruct (Preg.type_cases x) as [T|T]; rewrite T in H0; auto. + simpl in T. destruct Archi.ptr64; congruence. + destruct Archi.ptr64; simpl in EVAL'; congruence. } + auto. + intros (rs' & A & B & C). exists rs', (transl_cbranch_int64u c0 x X31 lbl). split. constructor; eexact A. split; auto. apply transl_cbranch_int64u_correct; auto. @@ -478,43 +581,59 @@ Proof. assert (V: v = Val.of_bool (eqb normal b)). { unfold v, Val.cmpf. rewrite EVAL'. destruct normal, b; reflexivity. } econstructor; econstructor. - split. constructor. apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto. - split. rewrite V; destruct normal, b; reflexivity. + split. constructor. apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto with asmgen. + split. rewrite V. destruct normal; unfold exec_instr, eval_branch. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. intros; Simpl. - destruct (transl_cond_float c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2. - assert (EVAL'': Val.cmpf_bool c0 (rs x) (rs x0) = Some (negb b)). - { destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; inv EVAL'; auto. } + assert (EVAL'': Val.cmpf_bool c0 (rs#x) (rs#x0) = Some (negb b)). + { destruct (Val.cmpf_bool c0 (rs#x) (rs#x0)) as [[]|]; inv EVAL'; auto. } set (v := if normal then Val.cmpf c0 rs#x rs#x0 else Val.notbool (Val.cmpf c0 rs#x rs#x0)). assert (V: v = Val.of_bool (xorb normal b)). { unfold v, Val.cmpf. rewrite EVAL''. destruct normal, b; reflexivity. } econstructor; econstructor. - split. constructor. apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto. - split. rewrite V; destruct normal, b; reflexivity. + split. constructor. apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto with asmgen. + split. rewrite V. destruct normal; unfold exec_instr, eval_branch. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. intros; Simpl. - destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2. set (v := if normal then Val.cmpfs c0 rs#x rs#x0 else Val.notbool (Val.cmpfs c0 rs#x rs#x0)). assert (V: v = Val.of_bool (eqb normal b)). { unfold v, Val.cmpfs. rewrite EVAL'. destruct normal, b; reflexivity. } econstructor; econstructor. - split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto. - split. rewrite V; destruct normal, b; reflexivity. + split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto with asmgen. + split. rewrite V. destruct normal; unfold exec_instr, eval_branch. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. intros; Simpl. - destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2. - assert (EVAL'': Val.cmpfs_bool c0 (rs x) (rs x0) = Some (negb b)). - { destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; inv EVAL'; auto. } + assert (EVAL'': Val.cmpfs_bool c0 (rs#x) (rs#x0) = Some (negb b)). + { destruct (Val.cmpfs_bool c0 (rs#x) (rs#x0)) as [[]|]; inv EVAL'; auto. } set (v := if normal then Val.cmpfs c0 rs#x rs#x0 else Val.notbool (Val.cmpfs c0 rs#x rs#x0)). assert (V: v = Val.of_bool (xorb normal b)). { unfold v, Val.cmpfs. rewrite EVAL''. destruct normal, b; reflexivity. } econstructor; econstructor. - split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto. - split. rewrite V; destruct normal, b; reflexivity. + split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto with asmgen. + split. rewrite V. destruct normal; unfold exec_instr, eval_branch. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. + unfold get0w. rewrite nextinstr_inv, Pregmap.gss by congruence. + simpl; destruct b, Archi.ptr64; reflexivity. intros; Simpl. Qed. Lemma transl_cbranch_correct_true: forall cond args lbl k c m ms sp rs m', transl_cbranch cond args lbl k = OK c -> - eval_condition cond (List.map ms args) m = Some true -> + eval_condition cond (List.map (fun r => Regmap.get r ms) args) m = Some true -> agree ms sp rs -> Mem.extends m m' -> exists rs', exists insn, @@ -528,7 +647,7 @@ Qed. Lemma transl_cbranch_correct_false: forall cond args lbl k c m ms sp rs m', transl_cbranch cond args lbl k = OK c -> - eval_condition cond (List.map ms args) m = Some false -> + eval_condition cond (List.map (fun r => Regmap.get r ms) args) m = Some false -> agree ms sp rs -> Mem.extends m m' -> exists rs', @@ -538,7 +657,7 @@ Proof. intros. exploit transl_cbranch_correct_1; eauto. simpl. intros (rs' & insn & A & B & C). exists (nextinstr rs'). - split. eapply exec_straight_opt_right; eauto. apply exec_straight_one; auto. + split. eapply exec_straight_opt_right; eauto. apply exec_straight_one; auto. Simpl. intros; Simpl. Qed. @@ -552,23 +671,31 @@ Lemma transl_cond_int32s_correct: /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. destruct (rs##r1); auto. destruct (rs##r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. destruct (rs##r1); auto. destruct (rs##r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. + unfold Val.cmp. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool. simpl. rewrite (Val.negate_cmp_bool Clt). - destruct (Val.cmp_bool Clt rs##r2 rs##r1) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool. auto. + destruct (Val.cmp_bool Clt rs##r2 rs##r1) as [[]|], Archi.ptr64; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. + split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool. + simpl Preg.chunk_of. destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmp. rewrite (Val.negate_cmp_bool Clt). - destruct (Val.cmp_bool Clt rs##r1 rs##r2) as [[]|]; auto. + simpl Preg.chunk_of. + destruct (Val.cmp_bool Clt rs##r1 rs##r2) as [[]|], Archi.ptr64; auto. Qed. Lemma transl_cond_int32u_correct: @@ -579,76 +706,97 @@ Lemma transl_cond_int32u_correct: /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool. simpl. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Cle). - destruct (Val.cmpu_bool (Mem.valid_pointer m) Cle rs##r1 rs##r2) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool. auto. + destruct (Val.cmpu_bool (Mem.valid_pointer m) Cle rs##r1 rs##r2) as [[]|], Archi.ptr64; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. + split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmpu. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Clt). - destruct (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##r1 rs##r2) as [[]|]; auto. + simpl Preg.chunk_of. + destruct (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##r1 rs##r2) as [[]|], Archi.ptr64; auto. Qed. Lemma transl_cond_int64s_correct: forall cmp rd r1 r2 k rs m, + Archi.ptr64 = true -> exists rs', exec_straight ge fn (transl_cond_int64s cmp rd r1 r2 k) rs m k rs' m /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs###r1 rs###r2)) rs'#rd /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. destruct (rs###r1); auto. destruct (rs###r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + simpl; rewrite H; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. destruct (rs###r1); auto. destruct (rs###r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + simpl; rewrite H; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. + simpl; rewrite H; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool. - simpl. rewrite (Val.negate_cmpl_bool Clt). + simpl. rewrite (Val.negate_cmpl_bool Clt), H. destruct (Val.cmpl_bool Clt rs###r2 rs###r1) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool. auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. + split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool. + simpl; rewrite H; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmpl. rewrite (Val.negate_cmpl_bool Clt). + simpl; rewrite H. destruct (Val.cmpl_bool Clt rs###r1 rs###r2) as [[]|]; auto. Qed. Lemma transl_cond_int64u_correct: forall cmp rd r1 r2 k rs m, + Archi.ptr64 = true -> exists rs', exec_straight ge fn (transl_cond_int64u cmp rd r1 r2 k) rs m k rs' m /\ rs'#rd = Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs###r1 rs###r2) /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + simpl; rewrite H; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. + simpl; rewrite H; auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. split; intros; Simpl. + simpl; rewrite H; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool. - simpl. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Cle). + simpl. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Cle), H. destruct (Val.cmplu_bool (Mem.valid_pointer m) Cle rs###r1 rs###r2) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool. auto. +- econstructor; split. apply exec_straight_one; [simpl; eauto|auto with asmgen]. + split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool. + simpl; rewrite H; auto. - econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. unfold Val.cmplu. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Clt). + simpl; rewrite H. destruct (Val.cmplu_bool (Mem.valid_pointer m) Clt rs###r1 rs###r2) as [[]|]; auto. Qed. @@ -674,26 +822,36 @@ Proof. exists rs2; split. eapply exec_straight_trans. eexact A1. eexact A2. split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. auto. - intros; transitivity (rs1 r); auto. } + intros; transitivity (rs1#r); auto. } destruct cmp. + unfold xorimm32. exploit (opimm32_correct Pxorw Pxoriw Val.xor); eauto. intros (rs1 & A1 & B1 & C1). exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2). exists rs2; split. eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. + split. simpl in B2; rewrite B1 in B2; simpl in B2. + destruct Archi.ptr64. + simpl in B2. destruct (rs#r1); auto. + unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2. + simpl in B2. destruct (rs#r1); auto. unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. + intros; transitivity (rs1#r); auto. + unfold xorimm32. exploit (opimm32_correct Pxorw Pxoriw Val.xor); eauto. intros (rs1 & A1 & B1 & C1). exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2). exists rs2; split. eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. + split. simpl in B2; rewrite B1 in B2; simpl in B2. + destruct Archi.ptr64. + simpl in B2. destruct (rs#r1); auto. + unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2. + simpl in B2. destruct (rs#r1); auto. unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. + intros; transitivity (rs1#r); auto. + exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. rewrite B1; auto. + exists rs1; split. eexact A1. split; auto. rewrite B1. + unfold Val.cmp. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. + predSpec Int.eq Int.eq_spec n (Int.repr Int.max_signed). * subst n. exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. @@ -703,7 +861,10 @@ Proof. generalize (Int.signed_range i); omega. * exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. - rewrite B1. unfold Val.cmp; simpl; destruct (rs#r1); simpl; auto. + rewrite B1. + unfold Val.cmp. simpl Preg.chunk_of. + rewrite Val.load_result_of_optbool by (destruct Archi.ptr64; auto). + simpl; destruct (rs#r1); simpl; auto. unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1). destruct (zlt (Int.signed n) (Int.signed i)). rewrite zlt_false by omega. auto. @@ -738,13 +899,15 @@ Proof. exists rs2; split. eapply exec_straight_trans. eexact A1. eexact A2. split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. rewrite B2; auto. - intros; transitivity (rs1 r); auto. } + intros; transitivity (rs1#r); auto. } destruct cmp. + apply DFL. + apply DFL. + exploit (opimm32_correct Psltuw Psltiuw (Val.cmpu (Mem.valid_pointer m) Clt) m); eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. rewrite B1; auto. + exists rs1; split. eexact A1. split; auto. rewrite B1. + unfold Val.cmpu. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. + apply DFL. + apply DFL. + apply DFL. @@ -752,6 +915,7 @@ Qed. Lemma transl_condimm_int64s_correct: forall cmp rd r1 n k rs m, + Archi.ptr64 = true -> r1 <> X31 -> exists rs', exec_straight ge fn (transl_condimm_int64s cmp rd r1 n k) rs m k rs' m @@ -760,7 +924,7 @@ Lemma transl_condimm_int64s_correct: Proof. intros. unfold transl_condimm_int64s. predSpec Int64.eq Int64.eq_spec n Int64.zero. -- subst n. exploit transl_cond_int64s_correct. intros (rs' & A & B & C). +- subst n. exploit transl_cond_int64s_correct; auto. intros (rs' & A & B & C). exists rs'; eauto. - assert (DFL: exists rs', @@ -772,7 +936,7 @@ Proof. exists rs2; split. eapply exec_straight_trans. eexact A1. eexact A2. split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. auto. - intros; transitivity (rs1 r); auto. } + intros; transitivity (rs1#r); auto. } destruct cmp. + unfold xorimm64. exploit (opimm64_correct Pxorl Pxoril Val.xorl); eauto. intros (rs1 & A1 & B1 & C1). @@ -781,7 +945,7 @@ Proof. eapply exec_straight_trans. eexact A1. eexact A2. split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. unfold Val.cmpl in B2; simpl in B2; rewrite Int64.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. + intros; transitivity (rs1#r); auto. + unfold xorimm64. exploit (opimm64_correct Pxorl Pxoril Val.xorl); eauto. intros (rs1 & A1 & B1 & C1). exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2). @@ -789,7 +953,7 @@ Proof. eapply exec_straight_trans. eexact A1. eexact A2. split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. unfold Val.cmpl in B2; simpl in B2; rewrite Int64.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. + intros; transitivity (rs1#r); auto. + exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. rewrite B1; auto. + predSpec Int64.eq Int64.eq_spec n (Int64.repr Int64.max_signed). @@ -808,7 +972,7 @@ Proof. rewrite zlt_true by omega. auto. rewrite Int64.add_signed. symmetry; apply Int64.signed_repr. assert (Int64.signed n <> Int64.max_signed). - { red; intros E. elim H1. rewrite <- (Int64.repr_signed n). rewrite E. auto. } + { red; intros E. elim H2. rewrite <- (Int64.repr_signed n). rewrite E. auto. } generalize (Int64.signed_range n); omega. + apply DFL. + apply DFL. @@ -816,6 +980,7 @@ Qed. Lemma transl_condimm_int64u_correct: forall cmp rd r1 n k rs m, + Archi.ptr64 = true -> r1 <> X31 -> exists rs', exec_straight ge fn (transl_condimm_int64u cmp rd r1 n k) rs m k rs' m @@ -824,7 +989,7 @@ Lemma transl_condimm_int64u_correct: Proof. intros. unfold transl_condimm_int64u. predSpec Int64.eq Int64.eq_spec n Int64.zero. -- subst n. exploit transl_cond_int64u_correct. intros (rs' & A & B & C). +- subst n. exploit transl_cond_int64u_correct; auto. intros (rs' & A & B & C). exists rs'; split. eexact A. split; auto. rewrite B; auto. - assert (DFL: exists rs', @@ -836,7 +1001,7 @@ Proof. exists rs2; split. eapply exec_straight_trans. eexact A1. eexact A2. split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. rewrite B2; auto. - intros; transitivity (rs1 r); auto. } + intros; transitivity (rs1#r); auto. } destruct cmp. + apply DFL. + apply DFL. @@ -853,7 +1018,7 @@ Lemma transl_cond_op_correct: transl_cond_op cond rd args k = OK c -> exists rs', exec_straight ge fn c rs m k rs' m - /\ Val.lessdef (Val.of_optbool (eval_condition cond (map rs (map preg_of args)) m)) rs'#rd + /\ Val.lessdef (Val.of_optbool (eval_condition cond (map (fun r => rs#r) (map preg_of args)) m)) rs'#rd /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. Proof. assert (MKTOT: forall ob, Val.of_optbool ob = Val.maketotal (option_map Val.of_bool ob)). @@ -885,60 +1050,72 @@ Proof. exists rs'; repeat split; eauto. rewrite MKTOT; eauto. + (* cmpf *) destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR. - fold (Val.cmpf c0 (rs x) (rs x0)). - set (v := Val.cmpf c0 (rs x) (rs x0)). + fold (Val.cmpf c0 (rs#x) (rs#x0)). + set (v := Val.cmpf c0 (rs#x) (rs#x0)). destruct normal; inv EQ2. * econstructor; split. - apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto. + apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto with asmgen. split; intros; Simpl. + unfold v, Val.cmpf. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. * econstructor; split. eapply exec_straight_two. eapply transl_cond_float_correct with (v := Val.notbool v); eauto. simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpf. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; auto. + auto with asmgen. auto with asmgen. + split; intros; Simpl. unfold v, Val.cmpf. simpl Preg.chunk_of. + destruct (Val.cmpf_bool c0 (rs#x) (rs#x0)) as [[]|], Archi.ptr64; auto. + (* notcmpf *) destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR. - rewrite Val.notbool_negb_3. fold (Val.cmpf c0 (rs x) (rs x0)). - set (v := Val.cmpf c0 (rs x) (rs x0)). + rewrite Val.notbool_negb_3. fold (Val.cmpf c0 (rs#x) (rs#x0)). + set (v := Val.cmpf c0 (rs#x) (rs#x0)). destruct normal; inv EQ2. * econstructor; split. eapply exec_straight_two. eapply transl_cond_float_correct with (v := v); eauto. simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpf. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; auto. + auto with asmgen. auto with asmgen. + split; intros; Simpl. unfold v, Val.cmpf. simpl Preg.chunk_of. + destruct (Val.cmpf_bool c0 (rs#x) (rs#x0)) as [[]|], Archi.ptr64; auto. * econstructor; split. - apply exec_straight_one. eapply transl_cond_float_correct with (v := Val.notbool v); eauto. auto. + apply exec_straight_one. eapply transl_cond_float_correct with (v := Val.notbool v); eauto. auto with asmgen. split; intros; Simpl. + simpl. unfold Val.notbool. destruct Archi.ptr64, v; auto. + destruct (Int.eq i Int.zero); auto. + (* cmpfs *) destruct (transl_cond_single c0 rd x x0) as [insn normal] eqn:TR. - fold (Val.cmpfs c0 (rs x) (rs x0)). - set (v := Val.cmpfs c0 (rs x) (rs x0)). + fold (Val.cmpfs c0 (rs#x) (rs#x0)). + set (v := Val.cmpfs c0 (rs#x) (rs#x0)). destruct normal; inv EQ2. * econstructor; split. - apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto. + apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto with asmgen. split; intros; Simpl. + unfold v, Val.cmpfs. simpl Preg.chunk_of. + destruct Archi.ptr64; try rewrite Val.load_result_of_optbool; auto. * econstructor; split. eapply exec_straight_two. eapply transl_cond_single_correct with (v := Val.notbool v); eauto. simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpfs. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; auto. + auto with asmgen. auto with asmgen. + split; intros; Simpl. unfold v, Val.cmpfs. simpl Preg.chunk_of. + destruct (Val.cmpfs_bool c0 (rs#x) (rs#x0)) as [[]|], Archi.ptr64; auto. + (* notcmpfs *) destruct (transl_cond_single c0 rd x x0) as [insn normal] eqn:TR. - rewrite Val.notbool_negb_3. fold (Val.cmpfs c0 (rs x) (rs x0)). - set (v := Val.cmpfs c0 (rs x) (rs x0)). + rewrite Val.notbool_negb_3. fold (Val.cmpfs c0 (rs#x) (rs#x0)). + set (v := Val.cmpfs c0 (rs#x) (rs#x0)). destruct normal; inv EQ2. * econstructor; split. eapply exec_straight_two. eapply transl_cond_single_correct with (v := v); eauto. simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpfs. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; auto. + auto with asmgen. auto with asmgen. + split; intros; Simpl. unfold v, Val.cmpfs. simpl Preg.chunk_of. + destruct (Val.cmpfs_bool c0 (rs#x) (rs#x0)) as [[]|], Archi.ptr64; auto. * econstructor; split. - apply exec_straight_one. eapply transl_cond_single_correct with (v := Val.notbool v); eauto. auto. + apply exec_straight_one. eapply transl_cond_single_correct with (v := Val.notbool v); eauto. auto with asmgen. split; intros; Simpl. + simpl. unfold Val.notbool. destruct Archi.ptr64, v; auto. + destruct (Int.eq i Int.zero); auto. Qed. (** Some arithmetic properties. *) @@ -964,24 +1141,38 @@ end. Ltac TranslOpSimpl := econstructor; split; - [ apply exec_straight_one; [simpl; eauto | reflexivity] - | split; [ apply Val.lessdef_same; Simpl; fail | intros; Simpl; fail ] ]. + [ apply exec_straight_one; [simpl; eauto | auto with asmgen] + | split; [ apply Val.lessdef_same; Simpl; invert_load_result; fail | intros; Simpl; invert_load_result; fail ] ]. + +Ltac ExecOne := + econstructor; split; + [ apply exec_straight_one; [simpl; eauto | auto with asmgen] + | split; intros; Simpl; try invert_load_result ]. Lemma transl_op_correct: forall op args res k (rs: regset) m v c, transl_op op args res k = OK c -> - eval_operation ge (rs#SP) op (map rs (map preg_of args)) m = Some v -> + eval_operation ge (rs#SP) op (map (fun r => rs#r) (map preg_of args)) m = Some v -> exists rs', exec_straight ge fn c rs m k rs' m /\ Val.lessdef v rs'#(preg_of res) - /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r. + /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs'#r = rs#r. Proof. assert (SAME: forall v1 v2, v1 = v2 -> Val.lessdef v2 v1). { intros; subst; auto. } Opaque Int.eq. intros until c; intros TR EV. unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; SimplEval EV; try TranslOpSimpl. - (* move *) - destruct (preg_of res), (preg_of m0); inv TR; TranslOpSimpl. + destruct (preg_of res), (preg_of m0); inv TR. + + (* ireg *) + econstructor; split. apply exec_straight_one; simpl; eauto. auto with asmgen. + intuition Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + change (Preg.type i) with (Preg.type i0). + apply Pregmap.get_has_type. + + (* freg *) + econstructor; split. apply exec_straight_one; simpl; eauto. auto with asmgen. + intuition Simpl. - (* intconst *) exploit loadimm32_correct; eauto. intros (rs' & A & B & C). exists rs'; split; eauto. rewrite B; auto with asmgen. @@ -991,18 +1182,18 @@ Opaque Int.eq. - (* floatconst *) destruct (Float.eq_dec n Float.zero). + subst n. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. + apply exec_straight_one. simpl; eauto. auto with asmgen. split; intros; Simpl. + econstructor; split. - apply exec_straight_one. simpl; eauto. auto. + apply exec_straight_one. simpl; eauto. auto with asmgen. split; intros; Simpl. - (* singleconst *) destruct (Float32.eq_dec n Float32.zero). + subst n. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. + apply exec_straight_one. simpl; eauto. auto with asmgen. split; intros; Simpl. + econstructor; split. - apply exec_straight_one. simpl; eauto. auto. + apply exec_straight_one. simpl; eauto. auto with asmgen. split; intros; Simpl. - (* addrsymbol *) destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)). @@ -1010,11 +1201,12 @@ Opaque Int.eq. exploit (addptrofs_correct x x ofs k rs1 m); eauto with asmgen. intros (rs2 & A & B & C). exists rs2; split. - apply exec_straight_step with rs1 m; auto. + apply exec_straight_step with rs1 m; auto. unfold rs1; auto with asmgen. split. replace ofs with (Ptrofs.add Ptrofs.zero ofs) by (apply Ptrofs.add_zero_l). rewrite Genv.shift_symbol_address. - replace (rs1 x) with (Genv.symbol_address ge id Ptrofs.zero) in B by (unfold rs1; Simpl). + replace (rs1#x) with (Genv.symbol_address ge id Ptrofs.zero) in B. exact B. + unfold rs1. Simpl. intros. rewrite C by eauto with asmgen. unfold rs1; Simpl. + TranslOpSimpl. - (* stackoffset *) @@ -1022,50 +1214,86 @@ Opaque Int.eq. exists rs'; split; eauto. auto with asmgen. - (* cast8signed *) econstructor; split. - eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. + eapply exec_straight_two. simpl;eauto. simpl;eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto. - destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. + destruct (rs#x0); auto; simpl. rewrite A; simpl. destruct Archi.ptr64; simpl; rewrite A. + apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. - (* cast16signed *) econstructor; split. - eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. + eapply exec_straight_two. simpl;eauto. simpl;eauto. auto with asmgen. auto with asmgen. split; intros; Simpl. assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto. - destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. + destruct (rs#x0); auto; simpl. rewrite A; simpl. destruct Archi.ptr64; simpl; rewrite A. + apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. - (* addimm *) exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). exists rs'; split; eauto. rewrite B; auto with asmgen. + split. invert_load_result. auto with asmgen. +- (* neg *) + ExecOne. + destruct (rs#x0); simpl; auto. destruct Archi.ptr64; simpl; auto. +- (* sub *) + ExecOne. destruct (eq_block _ _); auto. +- (* div *) + ExecOne. + subst v. apply Val.lessdef_same; Simpl. + destruct (rs#x0), (rs#x1); simpl; destruct Archi.ptr64; auto. + destruct (Int.eq _ _ || _); auto. +- (* divu *) + ExecOne. + subst v. + apply Val.lessdef_same; Simpl. + destruct (rs#x0), (rs#x1); simpl; destruct Archi.ptr64; auto. + destruct (Int.eq _ _); auto. +- (* mod *) + ExecOne. + subst v. apply Val.lessdef_same; Simpl. + destruct (rs#x0), (rs#x1); simpl; destruct Archi.ptr64; auto. + destruct (Int.eq _ _ || _); auto. +- (* modu *) + ExecOne. + subst v. apply Val.lessdef_same; Simpl. + destruct (rs#x0), (rs#x1); simpl; destruct Archi.ptr64; auto. + destruct (Int.eq _ _); auto. - (* andimm *) exploit (opimm32_correct Pandw Pandiw Val.and); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). exists rs'; split; eauto. rewrite B; auto with asmgen. + split. invert_load_result. auto with asmgen. - (* orimm *) exploit (opimm32_correct Porw Poriw Val.or); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). exists rs'; split; eauto. rewrite B; auto with asmgen. + split. invert_load_result. auto with asmgen. - (* xorimm *) exploit (opimm32_correct Pxorw Pxoriw Val.xor); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). exists rs'; split; eauto. rewrite B; auto with asmgen. + split. invert_load_result. auto with asmgen. - (* shrximm *) clear H. exploit Val.shrx_shr_2; eauto. intros E; subst v; clear EV. destruct (Int.eq n Int.zero). -+ econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. ++ econstructor; split. apply exec_straight_one. simpl; eauto. auto with asmgen. + split; intros; Simpl. + change (Preg.chunk_of x) with (Preg.chunk_of x0). rewrite Pregmap.get_load_result; auto. + change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n). econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto with asmgen. + eapply exec_straight_step. simpl; reflexivity. auto with asmgen. + eapply exec_straight_step. simpl; reflexivity. auto with asmgen. + apply exec_straight_one. simpl; reflexivity. auto with asmgen. split; intros; Simpl. + simpl. destruct Archi.ptr64; auto. + rewrite Val.load_result_add, !Val.load_result_shr, Val.load_result_shru; auto. - (* longofintu *) econstructor; split. - eapply exec_straight_three. simpl; eauto. simpl; eauto. simpl; eauto. auto. auto. auto. - split; intros; Simpl. destruct (rs x0); auto. simpl. + eapply exec_straight_three. simpl; eauto. simpl; eauto. simpl; eauto. + auto with asmgen. auto with asmgen. auto with asmgen. + split; intros; Simpl. destruct (rs#x0); auto. simpl. rewrite Heqb; simpl. assert (A: Int.ltu (Int.repr 32) Int64.iwordsize' = true) by auto. rewrite A; simpl. rewrite A. apply Val.lessdef_same. f_equal. rewrite cast32unsigned_from_cast32signed. apply Int64.zero_ext_shru_shl. compute; auto. @@ -1073,6 +1301,16 @@ Opaque Int.eq. exploit (opimm64_correct Paddl Paddil Val.addl); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). exists rs'; split; eauto. rewrite B; auto with asmgen. +- (* negl *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* divl *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* divlu *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* modl *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* modlu *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. - (* andimm *) exploit (opimm64_correct Pandl Pandil Val.andl); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). @@ -1088,15 +1326,43 @@ Opaque Int.eq. - (* shrxlimm *) clear H. exploit Val.shrxl_shrl_2; eauto. intros E; subst v; clear EV. destruct (Int.eq n Int.zero). -+ econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. ++ econstructor; split. apply exec_straight_one. simpl; eauto. auto with asmgen. + split; intros; Simpl. simpl; rewrite Heqb; auto. + change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n). econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. + eapply exec_straight_step. simpl; reflexivity. auto with asmgen. + eapply exec_straight_step. simpl; reflexivity. auto with asmgen. + eapply exec_straight_step. simpl; reflexivity. auto with asmgen. + apply exec_straight_one. simpl; reflexivity. auto with asmgen. + split; intros; Simpl. simpl; rewrite Heqb; auto. +- (* intoffloat *) + ExecOne. + subst v. apply Val.lessdef_same. + simpl Preg.chunk_of. destruct (rs#x0), Archi.ptr64; auto. destruct f; auto. + simpl. destruct (Float.to_int _); auto. +- (* intuoffloat *) + ExecOne. + subst v. apply Val.lessdef_same. + simpl Preg.chunk_of. destruct (rs#x0), Archi.ptr64; auto. destruct f; auto. + simpl. destruct (Float.to_intu _); auto. +- (* intofsingle *) + ExecOne. + subst v. apply Val.lessdef_same. + simpl Preg.chunk_of. destruct (rs#x0), Archi.ptr64; auto. destruct f; auto. + simpl. destruct (Float32.to_int _); auto. +- (* intuofsingle *) + ExecOne. + subst v. apply Val.lessdef_same. + simpl Preg.chunk_of. destruct (rs#x0), Archi.ptr64; auto. destruct f; auto. + simpl. destruct (Float32.to_intu _); auto. +- (* longoffloat *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* longuoffloat *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* longofsingle *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. +- (* longuofsingle *) + ExecOne. simpl Preg.chunk_of; rewrite Heqb; auto. - (* cond *) exploit transl_cond_op_correct; eauto. intros (rs' & A & B & C). exists rs'; split. eexact A. eauto with asmgen. @@ -1122,15 +1388,15 @@ Proof. split; auto. simpl. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto. + econstructor; econstructor; econstructor; split. constructor. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl. + simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. + split; intros; Simpl. destruct (rs#base); auto; simpl; rewrite SF; auto. simpl. rewrite Ptrofs.add_assoc. f_equal. f_equal. rewrite <- (Ptrofs.of_int64_to_int64 SF ofs). rewrite EQ. symmetry; auto with ptrofs. + econstructor; econstructor; econstructor; split. constructor. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold eval_offset. destruct (rs base); auto; simpl. rewrite SF. simpl. + simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. + split; intros; Simpl. unfold eval_offset. destruct (rs#base); auto; simpl; rewrite SF; auto. simpl. rewrite Ptrofs.add_zero. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto. - generalize (make_immed32_sound (Ptrofs.to_int ofs)); intros EQ. destruct (make_immed32 (Ptrofs.to_int ofs)). @@ -1139,8 +1405,8 @@ Proof. split; auto. simpl. subst imm. rewrite Ptrofs.of_int_to_int by auto. auto. + econstructor; econstructor; econstructor; split. constructor. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl. + simpl; eauto. simpl; eauto. auto with asmgen. auto with asmgen. + split; intros; Simpl. destruct (rs#base); auto; simpl; rewrite SF; auto. simpl. rewrite SF. simpl. rewrite Ptrofs.add_assoc. f_equal. f_equal. rewrite <- (Ptrofs.of_int_to_int SF ofs). rewrite EQ. symmetry; auto with ptrofs. @@ -1153,6 +1419,7 @@ Lemma indexed_load_access_correct: forall (base: ireg) ofs k (rs: regset) v, Mem.loadv chunk m (Val.offset_ptr rs#base ofs) = Some v -> base <> X31 -> rd <> PC -> + Val.has_type v (Preg.type rd) -> exists rs', exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m /\ rs'#rd = v @@ -1165,6 +1432,7 @@ Proof. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC. unfold exec_load. rewrite B, LOAD. eauto. Simpl. split; intros; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. Qed. Lemma indexed_store_access_correct: @@ -1183,14 +1451,14 @@ Proof. intros (base' & ofs' & rs' & A & B & C). econstructor; split. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC. - unfold exec_store. rewrite B, C, STORE by auto. eauto. auto. + unfold exec_store. rewrite B, C, STORE by auto. eauto. Simpl. intros; Simpl. Qed. Lemma loadind_correct: - forall (base: ireg) ofs ty dst k c (rs: regset) m v, - loadind base ofs ty dst k = OK c -> - Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v -> + forall (base: ireg) ofs q dst k c (rs: regset) m v, + loadind base ofs q dst k = OK c -> + Mem.loadv (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) = Some v -> base <> X31 -> exists rs', exec_straight ge fn c rs m k rs' m @@ -1202,16 +1470,20 @@ Proof. c = indexed_memory_access mk_instr base ofs k /\ forall base' ofs' rs', exec_instr ge fn (mk_instr base' ofs') rs' m = - exec_load ge (chunk_of_type ty) rs' m (preg_of dst) base' ofs'). - { unfold loadind in TR. destruct ty, (preg_of dst); inv TR; econstructor; split; eauto. } + exec_load ge (chunk_of_quantity q) rs' m (preg_of dst) base' ofs'). + { unfold loadind in TR. destruct q, (preg_of dst), Archi.ptr64; inv TR; econstructor; split; eauto. } destruct A as (mk_instr & B & C). subst c. eapply indexed_load_access_correct; eauto with asmgen. + unfold loadind in TR. + destruct q eqn:Q, (preg_of dst) eqn:P, Archi.ptr64 eqn:PTR64; try inv TR; + simpl; try rewrite PTR64; auto using Val.has_type_Tany64. + apply Mem.loadv_type in LOAD; apply LOAD. Qed. Lemma storeind_correct: - forall (base: ireg) ofs ty src k c (rs: regset) m m', - storeind src base ofs ty k = OK c -> - Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) rs#(preg_of src) = Some m' -> + forall (base: ireg) ofs q src k c (rs: regset) m m', + storeind src base ofs q k = OK c -> + Mem.storev (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) rs#(preg_of src) = Some m' -> base <> X31 -> exists rs', exec_straight ge fn c rs m k rs' m' @@ -1222,41 +1494,42 @@ Proof. c = indexed_memory_access mk_instr base ofs k /\ forall base' ofs' rs', exec_instr ge fn (mk_instr base' ofs') rs' m = - exec_store ge (chunk_of_type ty) rs' m (preg_of src) base' ofs'). - { unfold storeind in TR. destruct ty, (preg_of src); inv TR; econstructor; split; eauto. } + exec_store ge (chunk_of_quantity q) rs' m (preg_of src) base' ofs'). + { unfold storeind in TR. destruct q, (preg_of src); inv TR; econstructor; split; eauto. } destruct A as (mk_instr & B & C). subst c. eapply indexed_store_access_correct; eauto with asmgen. Qed. Lemma loadind_ptr_correct: forall (base: ireg) ofs (dst: ireg) k (rs: regset) m v, - Mem.loadv Mptr m (Val.offset_ptr rs#base ofs) = Some v -> + Mem.loadv Mptr_any m (Val.offset_ptr rs#base ofs) = Some v -> base <> X31 -> + Val.has_type v (Preg.type dst) -> exists rs', exec_straight ge fn (loadind_ptr base ofs dst k) rs m k rs' m /\ rs'#dst = v /\ forall r, r <> PC -> r <> X31 -> r <> dst -> rs'#r = rs#r. Proof. intros. eapply indexed_load_access_correct; eauto with asmgen. - intros. unfold Mptr. destruct Archi.ptr64; auto. + intros. unfold Mptr_any. destruct Archi.ptr64; auto. Qed. Lemma storeind_ptr_correct: forall (base: ireg) ofs (src: ireg) k (rs: regset) m m', - Mem.storev Mptr m (Val.offset_ptr rs#base ofs) rs#src = Some m' -> + Mem.storev Mptr_any m (Val.offset_ptr rs#base ofs) rs#src = Some m' -> base <> X31 -> src <> X31 -> exists rs', exec_straight ge fn (storeind_ptr src base ofs k) rs m k rs' m' /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. Proof. intros. eapply indexed_store_access_correct with (r1 := src); eauto with asmgen. - intros. unfold Mptr. destruct Archi.ptr64; auto. + intros. unfold Mptr_any. destruct Archi.ptr64; auto. Qed. Lemma transl_memory_access_correct: forall mk_instr addr args k c (rs: regset) m v, transl_memory_access mk_instr addr args k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> + eval_addressing ge rs#SP addr (map (fun r => rs#r) (map preg_of args)) = Some v -> exists base ofs rs', exec_straight_opt c rs m (mk_instr base ofs :: k) rs' m /\ Val.offset_ptr rs'#base (eval_offset ge ofs) = v @@ -1268,8 +1541,11 @@ Proof. inv EV. apply indexed_memory_access_correct; eauto with asmgen. - (* global *) simpl in EV. inv EV. inv TR. econstructor; econstructor; econstructor; split. - constructor. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. unfold eval_offset. apply low_high_half. + constructor. apply exec_straight_one. simpl; eauto. auto with asmgen. + split; intros; Simpl. unfold eval_offset. + rewrite Preg.chunk_of_reg_type, Val.load_result_same. + apply low_high_half. + simpl. destruct Archi.ptr64; auto using high_half_type, Val.has_type_Tany64. - (* stack *) inv TR. inv EV. apply indexed_memory_access_correct; eauto with asmgen. Qed. @@ -1279,9 +1555,10 @@ Lemma transl_load_access_correct: (forall base ofs rs, exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) -> transl_memory_access mk_instr addr args k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> + eval_addressing ge rs#SP addr (map (fun r => rs#r) (map preg_of args)) = Some v -> Mem.loadv chunk m v = Some v' -> rd <> PC -> + Val.has_type v' (Preg.type rd) -> exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#rd = v' @@ -1294,6 +1571,7 @@ Proof. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite INSTR. unfold exec_load. rewrite B, LOAD. reflexivity. Simpl. split; intros; Simpl. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. Qed. Lemma transl_store_access_correct: @@ -1301,7 +1579,7 @@ Lemma transl_store_access_correct: (forall base ofs rs, exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) -> transl_memory_access mk_instr addr args k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> + eval_addressing ge rs#SP addr (map (fun r => rs#r) (map preg_of args)) = Some v -> Mem.storev chunk m v rs#r1 = Some m' -> r1 <> PC -> r1 <> X31 -> exists rs', @@ -1313,14 +1591,14 @@ Proof. intros (base & ofs & rs' & A & B & C). econstructor; split. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. - rewrite INSTR. unfold exec_store. rewrite B, C, STORE by auto. reflexivity. auto. + rewrite INSTR. unfold exec_store. rewrite B, C, STORE by auto. reflexivity. Simpl. intros; Simpl. Qed. 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 -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge rs#SP addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> exists rs', exec_straight ge fn c rs m k rs' m @@ -1335,12 +1613,19 @@ Proof. { unfold transl_load in TR; destruct chunk; ArgsInv; econstructor; (split; [eassumption|auto]). } destruct A as (mk_instr & B & C). eapply transl_load_access_correct; eauto with asmgen. + eapply Val.has_subtype; eauto using Mem.loadv_type. + destruct Archi.ptr64 eqn:PTR64. + - unfold Preg.type. rewrite PTR64. destruct (preg_of dst), chunk; auto. + - destruct chunk; simpl in TR; try inversion TR; simpl; + destruct (Preg.type_cases (preg_of dst)) as [T|T]; rewrite T; auto. + unfold assertion_failed in H0. rewrite PTR64 in H0; congruence. + destruct dst; simpl in TR; simpl in T; congruence. Qed. Lemma transl_store_correct: forall chunk addr args src k c (rs: regset) m a m', transl_store chunk addr args src k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge rs#SP addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.storev chunk m a rs#(preg_of src) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' @@ -1366,8 +1651,8 @@ Qed. Lemma make_epilogue_correct: forall ge0 f m stk soff cs m' ms rs k tm, - load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) -> - load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs) -> + load_stack m (Vptr stk soff) (quantity_of_typ Tptr) f.(fn_link_ofs) = Some (parent_sp cs) -> + load_stack m (Vptr stk soff) (quantity_of_typ 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 -> @@ -1387,21 +1672,31 @@ Proof. exploit lessdef_parent_ra; eauto. intros EQ; subst ra'; clear LDRA'. exploit Mem.free_parallel_extends; eauto. intros (tm' & FREE' & MEXT'). unfold make_epilogue. - rewrite chunk_of_Tptr in *. + replace (chunk_of_quantity (quantity_of_typ Tptr)) with Mptr_any in * by + (unfold Tptr, Mptr_any; destruct Archi.ptr64; reflexivity). exploit (loadind_ptr_correct SP (fn_retaddr_ofs f) RA (Pfreeframe (fn_stacksize f) (fn_link_ofs f) :: k) rs tm). rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. congruence. + exploit Val.has_subtype; eauto using parent_ra_type. + simpl. unfold Tptr. destruct Archi.ptr64; auto. intros (rs1 & A1 & B1 & C1). econstructor; econstructor; split. eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl. rewrite (C1 X2) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. - rewrite FREE'. eauto. auto. + rewrite FREE'. eauto. auto with asmgen. split. apply agree_nextinstr. apply agree_set_other; auto with asmgen. apply agree_change_sp with (Vptr stk soff). apply agree_exten with rs; auto. intros; apply C1; auto with asmgen. eapply parent_sp_def; eauto. + eapply parent_sp_type; eauto. + exploit Val.has_subtype; eauto using parent_sp_type. + simpl. unfold Tptr. destruct Archi.ptr64; auto. split. auto. split. Simpl. split. Simpl. + rewrite Preg.chunk_of_reg_type. + rewrite Val.load_result_same; auto. + exploit Val.has_subtype; eauto using parent_sp_type. + simpl. unfold Tptr. destruct Archi.ptr64; auto. intros. Simpl. Qed. diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index df7ddfd20d..ce6d4cde79 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -19,7 +19,7 @@ machine registers and stack slots. *) Require Import Coqlib Decidableplus. -Require Import AST Machregs Locations. +Require Import AST Values Machregs Locations. (** * Classification of machine registers *) @@ -132,6 +132,17 @@ Proof. destruct (sig_res sig) as [[]|]; auto; destruct Archi.ptr64; auto. Qed. +Lemma loc_result_has_type: + forall res sig, + Val.has_type res (proj_sig_res sig) -> + Val.has_type_rpair res (loc_result sig) Val.loword Val.hiword mreg_type. +Proof. + intros. unfold Val.has_type_rpair, loc_result, proj_sig_res in *. + destruct sig; simpl. destruct sig_res; simpl in H. + destruct t, res; simpl; auto; destruct Archi.ptr64 eqn:P; simpl; try rewrite P; auto. + destruct res; simpl; auto; destruct Archi.ptr64; auto. +Qed. + (** The result locations are caller-save registers *) Lemma loc_result_caller_save: @@ -216,7 +227,7 @@ Definition one_arg (regs: list mreg) (rn: Z) (ofs: Z) (ty: typ) One(R r) :: rec (rn + 1) 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 (quantity_of_typ ty)) :: rec rn (ofs + (if Archi.ptr64 then 2 else typesize ty)) end. Definition two_args (regs: list mreg) (rn: Z) (ofs: Z) @@ -227,7 +238,7 @@ Definition two_args (regs: list mreg) (rn: Z) (ofs: Z) 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) :: + Twolong (S Outgoing (ofs + 1) Q32) (S Outgoing ofs Q32) :: rec rn (ofs + 2) end. @@ -239,7 +250,7 @@ Definition hybrid_arg (regs: list mreg) (rn: Z) (ofs: Z) (ty: typ) One (R r) :: rec (rn + 2) ofs | None => let ofs := align ofs 2 in - One (S Outgoing ofs ty) :: rec rn (ofs + 2) + One (S Outgoing ofs (quantity_of_typ ty)) :: rec rn (ofs + 2) end. Fixpoint loc_arguments_rec (va: bool) @@ -271,7 +282,7 @@ Definition loc_arguments (s: signature) : list (rpair loc) := Definition max_outgoing_1 (accu: Z) (l: loc) : Z := match l with - | S Outgoing ofs ty => Z.max accu (ofs + typesize ty) + | S Outgoing ofs q => Z.max accu (ofs + typesize (typ_of_quantity q)) | _ => accu end. @@ -290,7 +301,7 @@ Definition size_arguments (s: signature) : Z := Definition loc_argument_acceptable (l: loc) : Prop := match l with | R r => is_callee_save r = false - | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs) + | S Outgoing ofs q => ofs >= 0 /\ (quantity_align q | ofs) | _ => False end. @@ -318,7 +329,8 @@ Proof. 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. auto using align_divides, typealign_pos. + - subst p; simpl. split; auto. + destruct ty; auto using align_divides, typealign_pos. apply Z.divide_1_l. - eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. } @@ -328,7 +340,7 @@ Proof. set (rn' := align rn 2). 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) + assert (DFL: OK (Twolong (S Outgoing (ofs' + 1) Q32) (S Outgoing ofs' Q32) :: f rn' (ofs' + 2))). { red; simpl; intros. destruct H. - subst p; simpl. @@ -349,7 +361,7 @@ Proof. 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. + - subst p; simpl. split. apply (AL ofs Tlong OO). apply Z.divide_1_l. - eapply OF; [idtac|eauto]. generalize (AL ofs Tlong OO); simpl; omega. } assert (D: OKREGS int_param_regs). @@ -408,23 +420,23 @@ Proof. 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. + forall (s: signature) (ofs: Z) (q: quantity), + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments s)) -> + ofs + typesize (typ_of_quantity q) <= size_arguments s. Proof. - intros until ty. + intros until q. 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). + In (S Outgoing ofs q) (regs_of_rpair p) -> + ofs + typesize (typ_of_quantity q) <= 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). + In (S Outgoing ofs q) (regs_of_rpairs l) -> + ofs + typesize (typ_of_quantity q) <= fold_left max_outgoing_2 l n). { induction l; simpl; intros. - contradiction. - rewrite in_app_iff in H. destruct H. diff --git a/riscV/Machregs.v b/riscV/Machregs.v index d8bb4a4ba3..050406ed2a 100644 --- a/riscV/Machregs.v +++ b/riscV/Machregs.v @@ -22,6 +22,7 @@ Require Import Maps. Require Import AST. Require Import Integers. Require Import Op. +Require Import Memdata. (** ** Machine registers *) @@ -98,6 +99,11 @@ Definition mreg_type (r: mreg): typ := | F24 | F25 | F26 | F27 | F28 | F29 | F30 | F31 => Tany64 end. +Lemma mreg_type_cases: forall r, mreg_type r = Tany32 \/ mreg_type r = Tany64. +Proof. + destruct r; simpl; auto; destruct Archi.ptr64; auto. +Qed. + Open Scope positive_scope. Module IndexedMreg <: INDEXED_TYPE. @@ -105,28 +111,47 @@ Module IndexedMreg <: INDEXED_TYPE. Definition eq := mreg_eq. Definition index (r: mreg): positive := match r with - | R5 => 1 | R6 => 2 | R7 => 3 - | R8 => 4 | R9 => 5 | R10 => 6 | R11 => 7 - | R12 => 8 | R13 => 9 | R14 => 10 | R15 => 11 - | R16 => 12 | R17 => 13 | R18 => 14 | R19 => 15 - | R20 => 16 | R21 => 17 | R22 => 18 | R23 => 19 - | R24 => 20 | R25 => 21 | R26 => 22 | R27 => 23 - | R28 => 24 | R29 => 25 | R30 => 26 - - | F0 => 28 | F1 => 29 | F2 => 30 | F3 => 31 - | F4 => 32 | F5 => 33 | F6 => 34 | F7 => 35 - | F8 => 36 | F9 => 37 | F10 => 38 | F11 => 39 - | F12 => 40 | F13 => 41 | F14 => 42 | F15 => 43 - | F16 => 44 | F17 => 45 | F18 => 46 | F19 => 47 - | F20 => 48 | F21 => 49 | F22 => 50 | F23 => 51 - | F24 => 52 | F25 => 53 | F26 => 54 | F27 => 55 - | F28 => 56 | F29 => 57 | F30 => 58 | F31 => 59 + | R5 => 2 | R6 => 4 | R7 => 6 + | R8 => 8 | R9 => 10 | R10 => 12 | R11 => 14 + | R12 => 16 | R13 => 18 | R14 => 20 | R15 => 22 + | R16 => 24 | R17 => 26 | R18 => 28 | R19 => 30 + | R20 => 32 | R21 => 34 | R22 => 36 | R23 => 38 + | R24 => 40 | R25 => 42 | R26 => 44 | R27 => 46 + | R28 => 48 | R29 => 50 | R30 => 52 + + | F0 => 56 | F1 => 58 | F2 => 60 | F3 => 62 + | F4 => 64 | F5 => 66 | F6 => 68 | F7 => 70 + | F8 => 72 | F9 => 74 | F10 => 76 | F11 => 78 + | F12 => 80 | F13 => 82 | F14 => 84 | F15 => 86 + | F16 => 88 | F17 => 90 | F18 => 92 | F19 => 94 + | F20 => 96 | F21 => 98 | F22 => 100 | F23 => 102 + | F24 => 104 | F25 => 106 | F26 => 108 | F27 => 110 + | F28 => 112 | F29 => 114 | F30 => 116 | F31 => 118 end. Lemma index_inj: forall r1 r2, index r1 = index r2 -> r1 = r2. Proof. decide_goal. Qed. + + Open Scope Z_scope. + + Lemma scaled_index_with_size_aux: + forall r1 r2, Zpos (index r1) < Zpos (index r2) -> Zpos (index r1) + 2 <= Zpos (index r2). + Proof. + decide_goal. + Qed. + + Lemma scaled_index_with_size: + forall r1 r2, + Zpos (index r1) < Zpos (index r2) -> + Zpos (index r1) * 4 + AST.typesize (mreg_type r1) <= Zpos (index r2) * 4. + Proof. + intros. + generalize (scaled_index_with_size_aux r1 r2); intro. + assert (AST.typesize (mreg_type r1) <= 8) by (destruct (mreg_type r1); simpl; omega). + omega. + Qed. End IndexedMreg. Definition is_stack_reg (r: mreg) : bool := false. @@ -197,7 +222,7 @@ Definition destroyed_by_builtin (ef: external_function): list mreg := | _ => nil end. -Definition destroyed_by_setstack (ty: typ): list mreg := nil. +Definition destroyed_by_setstack (q: quantity): list mreg := nil. Definition destroyed_at_function_entry: list mreg := R30 :: nil. diff --git a/riscV/Op.v b/riscV/Op.v index bb04f78641..7c401cec01 100644 --- a/riscV/Op.v +++ b/riscV/Op.v @@ -691,6 +691,17 @@ Proof. intros; discriminate. Qed. +Lemma is_not_move_operation: + forall (F V A: Type) (genv: Genv.t F V) (sp: val) + (op: operation) (f: A -> val) (args: list A) (m: mem) (v: val), + eval_operation genv sp op (map f args) m = Some v -> + is_move_operation op args = None -> + op <> Omove. +Proof. + intros. destruct (eq_operation op Omove); auto. + subst. destruct args; try destruct args; simpl in *; congruence. +Qed. + (** [negate_condition cond] returns a condition that is logically equivalent to the negation of [cond]. *) diff --git a/riscV/Stacklayout.v b/riscV/Stacklayout.v index d0c6a526cc..f5ce99674e 100644 --- a/riscV/Stacklayout.v +++ b/riscV/Stacklayout.v @@ -54,8 +54,8 @@ Lemma frame_env_separated: m |= range sp 0 (fe_stack_data fe) ** range sp (fe_stack_data fe + bound_stack_data b) (fe_size fe) ** P -> m |= range sp (fe_ofs_local fe) (fe_ofs_local fe + 4 * bound_local b) ** range sp fe_ofs_arg (fe_ofs_arg + 4 * bound_outgoing b) - ** range sp (fe_ofs_link fe) (fe_ofs_link fe + size_chunk Mptr) - ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + size_chunk Mptr) + ** range sp (fe_ofs_link fe) (fe_ofs_link fe + size_chunk Mptr_any) + ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + size_chunk Mptr_any) ** range sp (fe_ofs_callee_save fe) (size_callee_save_area b (fe_ofs_callee_save fe)) ** P. Proof. @@ -89,6 +89,7 @@ Local Opaque Z.add Z.mul sepconj range. rewrite sep_swap45. (* Apply range_split and range_split2 repeatedly *) unfold fe_ofs_arg. + replace (size_chunk Mptr_any) with w by (unfold Mptr_any; destruct Archi.ptr64; auto). apply range_split_2. fold olink; omega. omega. apply range_split. omega. apply range_split. omega. @@ -127,8 +128,8 @@ Lemma frame_env_aligned: (8 | fe_ofs_arg) /\ (8 | fe_ofs_local fe) /\ (8 | fe_stack_data fe) - /\ (align_chunk Mptr | fe_ofs_link fe) - /\ (align_chunk Mptr | fe_ofs_retaddr fe). + /\ (align_chunk Mptr_any | fe_ofs_link fe) + /\ (align_chunk Mptr_any | fe_ofs_retaddr fe). Proof. intros; simpl. set (w := if Archi.ptr64 then 8 else 4). @@ -138,10 +139,19 @@ Proof. set (ol := align (size_callee_save_area b ocs) 8). set (ostkdata := align (ol + 4 * b.(bound_local)) 8). assert (0 < w) by (unfold w; destruct Archi.ptr64; omega). - replace (align_chunk Mptr) with w by (rewrite align_chunk_Mptr; auto). + replace (align_chunk Mptr_any) with 4 by (unfold Mptr_any; destruct Archi.ptr64; auto). split. apply Z.divide_0_r. split. apply align_divides; omega. split. apply align_divides; omega. - split. apply align_divides; omega. + split. + unfold olink, w. destruct Archi.ptr64. + apply Z.divide_trans with (m := 8). + exists 2; auto. + apply align_divides; omega. + apply align_divides; omega. + unfold oretaddr, w. destruct Archi.ptr64. + apply Z.divide_trans with (m := 8). + exists 2; auto. + apply Z.divide_add_r. apply align_divides; omega. apply Z.divide_refl. apply Z.divide_add_r. apply align_divides; omega. apply Z.divide_refl. Qed. diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml index e3fbfe365f..00dc2d0994 100644 --- a/riscV/TargetPrinter.ml +++ b/riscV/TargetPrinter.ml @@ -396,9 +396,9 @@ module Target : TARGET = fprintf oc " fmv.x.d %a, %a\n" ireg rd freg fs (* 32-bit (single-precision) floating point *) - | Pfls (fd, ra, ofs) -> + | Pfls (fd, ra, ofs) | Pfls_a (fd, ra, ofs) -> fprintf oc " flw %a, %a(%a)\n" freg fd offset ofs ireg ra - | Pfss (fs, ra, ofs) -> + | Pfss (fs, ra, ofs) | Pfss_a (fs, ra, ofs) -> fprintf oc " fsw %a, %a(%a)\n" freg fs offset ofs ireg ra | Pfnegs (fd, fs) -> diff --git a/x86/Asm.v b/x86/Asm.v index 7b4a29f421..81daa45a57 100644 --- a/x86/Asm.v +++ b/x86/Asm.v @@ -38,11 +38,32 @@ Proof. decide equality. Defined. Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}. Proof. decide equality. Defined. +Definition ireg_index (i: ireg): Z := + 2 + 2 * match i with + | RAX => 0 | RBX => 1 | RCX => 2 | RDX => 3 + | RSI => 4 | RDI => 5 | RBP => 6 | RSP => 7 + | R8 => 8 | R9 => 9 | R10 => 10 | R11 => 11 + | R12 => 12 | R13 => 13 | R14 => 14 | R15 => 15 + end. + +Definition freg_index (f: freg): Z := + 34 + 2 * match f with + | XMM0 => 0 | XMM1 => 1 | XMM2 => 2 | XMM3 => 3 + | XMM4 => 4 | XMM5 => 5 | XMM6 => 6 | XMM7 => 7 + | XMM8 => 8 | XMM9 => 9 | XMM10 => 10 | XMM11 => 11 + | XMM12 => 12 | XMM13 => 13 | XMM14 => 14 | XMM15 => 15 + end. + (** Bits of the flags register. *) Inductive crbit: Type := | ZF | CF | PF | SF | OF. +Definition crbit_index (c: crbit): Z := + 66 + 2 * match c with + | ZF => 0 | CF => 1 | PF => 2 | SF => 3 | OF => 4 + end. + (** All registers modeled here. *) Inductive preg: Type := @@ -121,11 +142,17 @@ Inductive instruction: Type := | Pmovsd_mf (a: addrmode) (r1: freg) | Pmovss_fi (rd: freg) (n: float32) (**r [movss] (single 32-bit float) *) | Pmovss_fm (rd: freg) (a: addrmode) + | Pmovss_fm_a (rd: freg) (a: addrmode) | Pmovss_mf (a: addrmode) (r1: freg) + | Pmovss_mf_a (a: addrmode) (r1: freg) | Pfldl_m (a: addrmode) (**r [fld] double precision *) + | Pfldl_m_a (a: addrmode) (**r [fld] double precision *) | Pfstpl_m (a: addrmode) (**r [fstp] double precision *) + | Pfstpl_m_a (a: addrmode) (**r [fstp] double precision *) | Pflds_m (a: addrmode) (**r [fld] simple precision *) + | Pflds_m_a (a: addrmode) (**r [fld] simple precision *) | Pfstps_m (a: addrmode) (**r [fstp] simple precision *) + | Pfstps_m_a (a: addrmode) (**r [fstp] simple precision *) (** Moves with conversion *) | Pmovb_mr (a: addrmode) (rs: ireg) (**r [mov] (8-bit int) *) | Pmovw_mr (a: addrmode) (rs: ireg) (**r [mov] (16-bit int) *) @@ -242,8 +269,10 @@ Inductive instruction: Type := | Pcall_r (r: ireg) (sg: signature) | Pret (** Saving and restoring registers *) - | Pmov_rm_a (rd: ireg) (a: addrmode) (**r like [Pmov_rm], using [Many64] chunk *) - | Pmov_mr_a (a: addrmode) (rs: ireg) (**r like [Pmov_mr], using [Many64] chunk *) + | Pmovl_rm_a (rd: ireg) (a: addrmode) (**r like [Pmovl_rm], using [Many32] chunk *) + | Pmovq_rm_a (rd: ireg) (a: addrmode) (**r like [Pmovq_rm], using [Many64] chunk *) + | Pmovl_mr_a (a: addrmode) (rs: ireg) (**r like [Pmovl_mr], using [Many32] chunk *) + | Pmovq_mr_a (a: addrmode) (rs: ireg) (**r like [Pmovq_mr], using [Many64] chunk *) | Pmovsd_fm_a (rd: freg) (a: addrmode) (**r like [Pmovsd_fm], using [Many64] chunk *) | Pmovsd_mf_a (a: addrmode) (r1: freg) (**r like [Pmovsd_mf], using [Many64] chunk *) (** Pseudo-instructions *) @@ -300,19 +329,166 @@ Definition program := AST.program fundef unit. Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}. Proof. decide equality. apply ireg_eq. apply freg_eq. decide equality. Defined. -Module PregEq. - Definition t := preg. - Definition eq := preg_eq. -End PregEq. +Definition preg_index (p: preg): Z := + match p with + | IR i => ireg_index i + | FR f => freg_index f + | CR c => crbit_index c + | PC => 76 + | ST0 => 78 + | RA => 80 + end. + +Lemma preg_index_bounds: + forall p: preg, + match p with + | IR _ => 0 < preg_index p /\ preg_index p <= 32 + | FR _ => 33 < preg_index p /\ preg_index p <= 64 + | CR _ => 65 < preg_index p /\ preg_index p <= 74 + | _ => 75 < preg_index p + end. +Proof. + destruct p; unfold preg_index; try omega. + destruct i; unfold ireg_index; omega. + destruct f; unfold freg_index; omega. + destruct c; unfold crbit_index; omega. +Qed. -Module Pregmap := EMap(PregEq). +Module Preg <: REGISTER_MODEL. + + Definition reg := preg. + Definition eq_dec := preg_eq. + + Definition type p := + match p with + | IR _ => if Archi.ptr64 then Tany64 else Tany32 + | PC | RA => if Archi.ptr64 then Tany64 else Tany32 + | FR _ => Tany64 + | CR _ => Tany32 + | ST0 => Tany64 + end. + + Definition quantity_of p := + match p with + | IR _ => if Archi.ptr64 then Q64 else Q32 + | PC | RA => if Archi.ptr64 then Q64 else Q32 + | FR _ => Q64 + | CR _ => Q32 + | ST0 => Q64 + end. + + Definition chunk_of p := + match p with + | IR _ => if Archi.ptr64 then Many64 else Many32 + | PC | RA => if Archi.ptr64 then Many64 else Many32 + | FR _ => Many64 + | CR _ => Many32 + | ST0 => Many64 + end. + + Lemma type_cases: forall r, type r = Tany32 \/ type r = Tany64. + Proof. + destruct r; simpl; auto. + Qed. + + Definition ofs (r: preg): Z := + preg_index r. + + (* A register's address: The index of its first byte. *) + Definition addr (r: preg): Z := + preg_index r * 4. + + (* The address one byte past the end of register [r]. The next nonoverlapping + register may start here. *) + Definition next_addr (r: preg): Z := addr r + AST.typesize (type r). + + Remark addr_pos: forall p, addr p > 0. + Proof. + intros. unfold addr. + destruct p; simpl; try omega. + destruct i; simpl; omega. + destruct f; simpl; omega. + destruct c; simpl; omega. + Qed. + + Lemma addr_compat: forall p, FragBlock.addr (ofs p) = addr p. + Proof. + reflexivity. + Qed. + + Lemma size_compat: + forall p, + AST.typesize (type p) = FragBlock.quantity_size (quantity_of p). + Proof. + intros. unfold quantity_of. + destruct p; simpl; auto. + Qed. + + Lemma next_addr_compat: forall p, FragBlock.next_addr (ofs p) (quantity_of p) = next_addr p. + Proof. + unfold next_addr, addr, ofs, FragBlock.next_addr, FragBlock.addr; intros. + rewrite size_compat. auto. + Qed. + + Lemma quantity_of_compat: + forall p, + quantity_of p = quantity_of_typ (type p). + Proof. + destruct p; simpl; destruct Archi.ptr64; try destruct i; reflexivity. + Qed. + + Lemma chunk_of_reg_compat: + forall p, + chunk_of p = chunk_of_quantity (quantity_of p). + Proof. + destruct p; simpl; destruct Archi.ptr64; try destruct i; reflexivity. + Qed. + + Lemma chunk_of_reg_type: + forall p, + chunk_of p = chunk_of_type (type p). + Proof. + destruct p; simpl; destruct Archi.ptr64; try destruct i; reflexivity. + Qed. + + Lemma diff_outside_interval: + forall r s, r <> s -> next_addr r <= addr s \/ next_addr s <= addr r. + Proof. + intros. + unfold next_addr, addr. + assert (TS: forall p, AST.typesize (type p) = 4 \/ AST.typesize (type p) = 8). + { intro p. destruct (type_cases p) as [T | T]; rewrite T; auto. } + generalize (preg_index_bounds r), (preg_index_bounds s); intros RB SB. + destruct r eqn:R, s eqn:S; + try congruence; + try (destruct (TS r), (TS s); subst; omega); + simpl AST.typesize; + try (destruct Archi.ptr64; simpl; omega). + - (* two iregs *) + destruct Archi.ptr64; simpl; destruct i, i0; simpl; try omega; congruence. + - (* two fregs *) + destruct f, f0; simpl; try omega; congruence. + - (* two crbits *) + destruct Archi.ptr64; destruct c, c0; simpl; try omega; congruence. + Qed. + +End Preg. + +Lemma pc_type: subtype Tptr (Preg.type PC) = true. +Proof. + unfold Tptr. simpl Preg.type. destruct Archi.ptr64; auto. +Qed. + +Module Pregmap := Regfile(Preg). -Definition regset := Pregmap.t val. +Definition regset := Pregmap.t. Definition genv := Genv.t fundef unit. -Notation "a # b" := (a b) (at level 1, only parsing) : asm. +Notation "a # b" := (Pregmap.get b a) (at level 1, only parsing) : asm. Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm. +Definition pregmap_read (rs: regset) := fun r => Pregmap.get r rs. + Open Scope asm. (** Undefining some registers *) @@ -320,9 +496,26 @@ Open Scope asm. Fixpoint undef_regs (l: list preg) (rs: regset) : regset := match l with | nil => rs - | r :: l' => undef_regs l' (rs#r <- Vundef) + | r :: l' => (undef_regs l' rs)#r <- Vundef end. +Lemma undef_regs_other: + forall r rl rs, ~In r rl -> (undef_regs rl rs) # r = rs # r. +Proof. + induction rl; simpl; intros. auto. rewrite Pregmap.gso. apply IHrl. intuition. intuition. +Qed. + +Lemma undef_regs_same: + forall r rl rs, In r rl -> (undef_regs rl rs) # r = Vundef. +Proof. + induction rl; simpl; intros. tauto. + destruct H. + - subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + - destruct (Preg.eq_dec r a). + + subst a. rewrite Pregmap.gss. destruct (Preg.chunk_of r); auto. + + rewrite Pregmap.gso; auto. +Qed. + (** Assigning a register pair *) Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := @@ -333,11 +526,11 @@ Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := (** Assigning the result of a builtin *) -Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := +Definition set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := match res with | BR r => rs#r <- v | BR_none => rs - | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) + | BR_splitlong hi lo => rs #hi <- (Val.hiword v) #lo <- (Val.loword v) end. Section RELSEM. @@ -381,14 +574,14 @@ Definition eval_addrmode32 (a: addrmode) (rs: regset) : val := let '(Addrmode base ofs const) := a in Val.add (match base with | None => Vint Int.zero - | Some r => rs r + | Some r => rs # r end) (Val.add (match ofs with | None => Vint Int.zero | Some(r, sc) => if zeq sc 1 - then rs r - else Val.mul (rs r) (Vint (Int.repr sc)) + then rs # r + else Val.mul (rs # r) (Vint (Int.repr sc)) end) (match const with | inl ofs => Vint (Int.repr ofs) @@ -399,14 +592,14 @@ Definition eval_addrmode64 (a: addrmode) (rs: regset) : val := let '(Addrmode base ofs const) := a in Val.addl (match base with | None => Vlong Int64.zero - | Some r => rs r + | Some r => rs # r end) (Val.addl (match ofs with | None => Vlong Int64.zero | Some(r, sc) => if zeq sc 1 - then rs r - else Val.mull (rs r) (Vlong (Int64.repr sc)) + then rs # r + else Val.mull (rs # r) (Vlong (Int64.repr sc)) end) (match const with | inl ofs => Vlong (Int64.repr ofs) @@ -476,62 +669,62 @@ Definition compare_floats32 (vx vy: val) (rs: regset) : regset := Definition eval_testcond (c: testcond) (rs: regset) : option bool := match c with | Cond_e => - match rs ZF with + match rs # ZF with | Vint n => Some (Int.eq n Int.one) | _ => None end | Cond_ne => - match rs ZF with + match rs # ZF with | Vint n => Some (Int.eq n Int.zero) | _ => None end | Cond_b => - match rs CF with + match rs # CF with | Vint n => Some (Int.eq n Int.one) | _ => None end | Cond_be => - match rs CF, rs ZF with + match rs # CF, rs # ZF with | Vint c, Vint z => Some (Int.eq c Int.one || Int.eq z Int.one) | _, _ => None end | Cond_ae => - match rs CF with + match rs # CF with | Vint n => Some (Int.eq n Int.zero) | _ => None end | Cond_a => - match rs CF, rs ZF with + match rs # CF, rs # ZF with | Vint c, Vint z => Some (Int.eq c Int.zero && Int.eq z Int.zero) | _, _ => None end | Cond_l => - match rs OF, rs SF with + match rs # OF, rs # SF with | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.one) | _, _ => None end | Cond_le => - match rs OF, rs SF, rs ZF with + match rs # OF, rs # SF, rs # ZF with | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.one || Int.eq z Int.one) | _, _, _ => None end | Cond_ge => - match rs OF, rs SF with + match rs # OF, rs # SF with | Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.zero) | _, _ => None end | Cond_g => - match rs OF, rs SF, rs ZF with + match rs # OF, rs # SF, rs # ZF with | Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.zero && Int.eq z Int.zero) | _, _, _ => None end | Cond_p => - match rs PF with + match rs # PF with | Vint n => Some (Int.eq n Int.one) | _ => None end | Cond_np => - match rs PF with + match rs # PF with | Vint n => Some (Int.eq n Int.zero) | _ => None end @@ -580,7 +773,7 @@ Definition exec_load (chunk: memory_chunk) (m: mem) Definition exec_store (chunk: memory_chunk) (m: mem) (a: addrmode) (rs: regset) (r1: preg) (destroyed: list preg) := - match Mem.storev chunk m (eval_addrmode a rs) (rs r1) with + match Mem.storev chunk m (eval_addrmode a rs) (rs # r1) with | Some m' => Next (nextinstr_nf (undef_regs destroyed rs)) m' | None => Stuck end. @@ -608,7 +801,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out match i with (** Moves *) | Pmov_rr rd r1 => - Next (nextinstr (rs#rd <- (rs r1))) m + Next (nextinstr (rs#rd <- (rs # r1))) m | Pmovl_ri rd n => Next (nextinstr_nf (rs#rd <- (Vint n))) m | Pmovq_ri rd n => @@ -624,7 +817,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pmovq_mr a r1 => exec_store Mint64 m a rs r1 nil | Pmovsd_ff rd r1 => - Next (nextinstr (rs#rd <- (rs r1))) m + Next (nextinstr (rs#rd <- (rs # r1))) m | Pmovsd_fi rd n => Next (nextinstr (rs#rd <- (Vfloat n))) m | Pmovsd_fm rd a => @@ -635,16 +828,28 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out Next (nextinstr (rs#rd <- (Vsingle n))) m | Pmovss_fm rd a => exec_load Mfloat32 m a rs rd + | Pmovss_fm_a rd a => + exec_load Many32 m a rs rd | Pmovss_mf a r1 => exec_store Mfloat32 m a rs r1 nil + | Pmovss_mf_a a r1 => + exec_store Many32 m a rs r1 nil | Pfldl_m a => exec_load Mfloat64 m a rs ST0 + | Pfldl_m_a a => + exec_load Many64 m a rs ST0 | Pfstpl_m a => exec_store Mfloat64 m a rs ST0 (ST0 :: nil) + | Pfstpl_m_a a => + exec_store Many64 m a rs ST0 (ST0 :: nil) | Pflds_m a => exec_load Mfloat32 m a rs ST0 + | Pflds_m_a a => + exec_load Many32 m a rs ST0 | Pfstps_m a => exec_store Mfloat32 m a rs ST0 (ST0 :: nil) + | Pfstps_m_a a => + exec_store Many32 m a rs ST0 (ST0 :: nil) (** Moves with conversion *) | Pmovb_mr a r1 => exec_store Mint8unsigned m a rs r1 nil @@ -834,21 +1039,21 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Prorq_ri rd n => Next (nextinstr_nf (rs#rd <- (Val.rorl rs#rd (Vint n)))) m | Pcmpl_rr r1 r2 => - Next (nextinstr (compare_ints (rs r1) (rs r2) rs m)) m + Next (nextinstr (compare_ints (rs # r1) (rs # r2) rs m)) m | Pcmpq_rr r1 r2 => - Next (nextinstr (compare_longs (rs r1) (rs r2) rs m)) m + Next (nextinstr (compare_longs (rs # r1) (rs # r2) rs m)) m | Pcmpl_ri r1 n => - Next (nextinstr (compare_ints (rs r1) (Vint n) rs m)) m + Next (nextinstr (compare_ints (rs # r1) (Vint n) rs m)) m | Pcmpq_ri r1 n => - Next (nextinstr (compare_longs (rs r1) (Vlong n) rs m)) m + Next (nextinstr (compare_longs (rs # r1) (Vlong n) rs m)) m | Ptestl_rr r1 r2 => - Next (nextinstr (compare_ints (Val.and (rs r1) (rs r2)) Vzero rs m)) m + Next (nextinstr (compare_ints (Val.and (rs # r1) (rs # r2)) Vzero rs m)) m | Ptestq_rr r1 r2 => - Next (nextinstr (compare_longs (Val.andl (rs r1) (rs r2)) (Vlong Int64.zero) rs m)) m + Next (nextinstr (compare_longs (Val.andl (rs # r1) (rs # r2)) (Vlong Int64.zero) rs m)) m | Ptestl_ri r1 n => - Next (nextinstr (compare_ints (Val.and (rs r1) (Vint n)) Vzero rs m)) m + Next (nextinstr (compare_ints (Val.and (rs # r1) (Vint n)) Vzero rs m)) m | Ptestq_ri r1 n => - Next (nextinstr (compare_longs (Val.andl (rs r1) (Vlong n)) (Vlong Int64.zero) rs m)) m + Next (nextinstr (compare_longs (Val.andl (rs # r1) (Vlong n)) (Vlong Int64.zero) rs m)) m | Pcmov c rd r1 => match eval_testcond c rs with | Some true => Next (nextinstr (rs#rd <- (rs#r1))) m @@ -871,7 +1076,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pabsd rd => Next (nextinstr (rs#rd <- (Val.absf rs#rd))) m | Pcomisd_ff r1 r2 => - Next (nextinstr (compare_floats (rs r1) (rs r2) rs)) m + Next (nextinstr (compare_floats (rs # r1) (rs # r2) rs)) m | Pxorpd_f rd => Next (nextinstr_nf (rs#rd <- (Vfloat Float.zero))) m (** Arithmetic operations over single-precision floats *) @@ -888,7 +1093,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pabss rd => Next (nextinstr (rs#rd <- (Val.absfs rs#rd))) m | Pcomiss_ff r1 r2 => - Next (nextinstr (compare_floats32 (rs r1) (rs r2) rs)) m + Next (nextinstr (compare_floats32 (rs # r1) (rs # r2) rs)) m | Pxorps_f rd => Next (nextinstr_nf (rs#rd <- (Vsingle Float32.zero))) m (** Branches and calls *) @@ -897,7 +1102,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pjmp_s id sg => Next (rs#PC <- (Genv.symbol_address ge id Ptrofs.zero)) m | Pjmp_r r sg => - Next (rs#PC <- (rs r)) m + Next (rs#PC <- (rs # r)) m | Pjcc cond lbl => match eval_testcond cond rs with | Some true => goto_label f lbl rs m @@ -922,14 +1127,18 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pcall_s id sg => Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (Genv.symbol_address ge id Ptrofs.zero)) m | Pcall_r r sg => - Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (rs r)) m + Next (rs#RA <- (Val.offset_ptr rs#PC Ptrofs.one) #PC <- (rs # r)) m | Pret => Next (rs#PC <- (rs#RA)) m (** Saving and restoring registers *) - | Pmov_rm_a rd a => - exec_load (if Archi.ptr64 then Many64 else Many32) m a rs rd - | Pmov_mr_a a r1 => - exec_store (if Archi.ptr64 then Many64 else Many32) m a rs r1 nil + | Pmovl_rm_a rd a => + exec_load Many32 m a rs rd + | Pmovq_rm_a rd a => + exec_load Many64 m a rs rd + | Pmovl_mr_a a r1 => + exec_store Many32 m a rs r1 nil + | Pmovq_mr_a a r1 => + exec_store Many64 m a rs r1 nil | Pmovsd_fm_a rd a => exec_load Many64 m a rs rd | Pmovsd_mf_a a r1 => @@ -940,19 +1149,19 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pallocframe sz ofs_ra ofs_link => let (m1, stk) := Mem.alloc m 0 sz in let sp := Vptr stk Ptrofs.zero in - match Mem.storev Mptr m1 (Val.offset_ptr sp ofs_link) rs#RSP with + match Mem.storev Mptr_any m1 (Val.offset_ptr sp ofs_link) rs#RSP with | None => Stuck | Some m2 => - match Mem.storev Mptr m2 (Val.offset_ptr sp ofs_ra) rs#RA with + match Mem.storev Mptr_any m2 (Val.offset_ptr sp ofs_ra) rs#RA with | None => Stuck | Some m3 => Next (nextinstr (rs #RAX <- (rs#RSP) #RSP <- sp)) m3 end end | Pfreeframe sz ofs_ra ofs_link => - match Mem.loadv Mptr m (Val.offset_ptr rs#RSP ofs_ra) with + match Mem.loadv Mptr_any m (Val.offset_ptr rs#RSP ofs_ra) with | None => Stuck | Some ra => - match Mem.loadv Mptr m (Val.offset_ptr rs#RSP ofs_link) with + match Mem.loadv Mptr_any m (Val.offset_ptr rs#RSP ofs_link) with | None => Stuck | Some sp => match rs#RSP with @@ -1049,25 +1258,54 @@ Definition preg_of (r: mreg) : preg := (** Undefine all registers except SP and callee-save registers *) -Definition undef_caller_save_regs (rs: regset) : regset := +Definition undef_caller_save_regs_spec (rs: preg -> val) : preg -> val := fun r => if preg_eq r SP || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) then rs r else Vundef. +Definition pregs_destroyed_at_call := + PC :: ST0 :: CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: RA + :: (map preg_of (filter (fun m => negb (is_callee_save m)) all_mregs)). + +Lemma pregs_destroyed_at_call_correct: + forall r, + preg_eq r SP + || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) + = + negb (In_dec preg_eq r pregs_destroyed_at_call). +Proof. + intros. + destruct r as [|x|x| |x|]; try destruct x; auto. +Qed. + +Definition undef_caller_save_regs (rs: regset) : regset := + undef_regs pregs_destroyed_at_call rs. + +Lemma undef_caller_save_regs_correct: + forall rs r, + (undef_caller_save_regs rs) # r = undef_caller_save_regs_spec (fun r' => rs # r) r. +Proof. + intros. unfold undef_caller_save_regs, undef_caller_save_regs_spec. + rewrite pregs_destroyed_at_call_correct. + destruct (In_dec preg_eq r pregs_destroyed_at_call) as [IN | NOTIN]. + - rewrite undef_regs_same; auto. + - rewrite undef_regs_other; auto. +Qed. + (** Extract the values of the arguments of an external call. We exploit the calling conventions from module [Conventions], except that we use machine registers instead of locations. *) Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop := | extcall_arg_reg: forall r, - extcall_arg rs m (R r) (rs (preg_of r)) - | extcall_arg_stack: forall ofs ty bofs v, + extcall_arg rs m (R r) (rs # (preg_of r)) + | extcall_arg_stack: forall ofs q bofs v, bofs = Stacklayout.fe_ofs_arg + 4 * ofs -> - Mem.loadv (chunk_of_type ty) m - (Val.offset_ptr (rs (IR RSP)) (Ptrofs.repr bofs)) = Some v -> - extcall_arg rs m (S Outgoing ofs ty) v. + Mem.loadv (chunk_of_quantity q) m + (Val.offset_ptr (rs # (IR RSP)) (Ptrofs.repr bofs)) = Some v -> + extcall_arg rs m (S Outgoing ofs q) v. Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop := | extcall_arg_one: forall l v, @@ -1093,17 +1331,17 @@ Inductive state: Type := Inductive step: state -> trace -> state -> Prop := | exec_step_internal: forall b ofs f i rs m rs' m', - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some i -> exec_instr f i rs m = Next rs' m' -> step (State rs m) E0 (State rs' m') | exec_step_builtin: forall b ofs f ef args res rs m vargs t vres rs' m', - rs PC = Vptr b ofs -> + rs # PC = Vptr b ofs -> Genv.find_funct_ptr ge b = Some (Internal f) -> find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) -> - eval_builtin_args ge rs (rs RSP) m args vargs -> + eval_builtin_args ge (pregmap_read rs) (rs # RSP) m args vargs -> external_call ef ge vargs m t vres m' -> rs' = nextinstr_nf (set_res res vres @@ -1111,11 +1349,11 @@ Inductive step: state -> trace -> state -> Prop := step (State rs m) t (State rs' m') | exec_step_external: forall b ef args res rs m t rs' m', - rs PC = Vptr b Ptrofs.zero -> + rs # PC = Vptr b Ptrofs.zero -> Genv.find_funct_ptr ge b = Some (External ef) -> extcall_arguments rs m (ef_sig ef) args -> external_call ef ge args m t res m' -> - rs' = (set_pair (loc_external_result (ef_sig ef)) res (undef_caller_save_regs rs)) #PC <- (rs RA) -> + rs' = (set_pair (loc_external_result (ef_sig ef)) res (undef_caller_save_regs rs)) #PC <- (rs#RA) -> step (State rs m) t (State rs' m'). End RELSEM. @@ -1127,7 +1365,7 @@ Inductive initial_state (p: program): state -> Prop := Genv.init_mem p = Some m0 -> let ge := Genv.globalenv p in let rs0 := - (Pregmap.init Vundef) + Pregmap.init # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero) # RA <- Vnullptr # RSP <- Vnullptr in diff --git a/x86/Asmexpand.ml b/x86/Asmexpand.ml index 99666920bf..8f73be2f2e 100644 --- a/x86/Asmexpand.ml +++ b/x86/Asmexpand.ml @@ -173,7 +173,7 @@ let expand_builtin_vload_common chunk addr res = emit (Pmovl_rm (res,addr)) | Mint64, BR(IR res) -> emit (Pmovq_rm (res,addr)) - | Mint64, BR_splitlong(BR(IR res1), BR(IR res2)) -> + | Mint64, BR_splitlong(IR res1, IR res2) -> let addr' = offset_addressing addr _4z in if not (Asmgen.addressing_mentions addr res2) then begin emit (Pmovl_rm (res2,addr)); @@ -323,7 +323,7 @@ let expand_builtin_inline name args res = emit (Pmov_rr (res,a1)); emit (Pbswap64 res) | "__builtin_bswap64", [BA_splitlong(BA(IR ah), BA(IR al))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> assert (ah = RAX && al = RDX && rh = RDX && rl = RAX); emit (Pbswap32 RAX); emit (Pbswap32 RDX) @@ -424,25 +424,25 @@ let expand_builtin_inline name args res = (fun r1 r2 r3 -> Pfnmsub231(r1, r2, r3)) (* 64-bit integer arithmetic *) | "__builtin_negl", [BA_splitlong(BA(IR ah), BA(IR al))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> assert (ah = RDX && al = RAX && rh = RDX && rl = RAX); emit (Pnegl RAX); emit (Padcl_ri (RDX,_0)); emit (Pnegl RDX) | "__builtin_addl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> assert (ah = RDX && al = RAX && bh = RCX && bl = RBX && rh = RDX && rl = RAX); emit (Paddl_rr (RAX,RBX)); emit (Padcl_rr (RDX,RCX)) | "__builtin_subl", [BA_splitlong(BA(IR ah), BA(IR al)); BA_splitlong(BA(IR bh), BA(IR bl))], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> assert (ah = RDX && al = RAX && bh = RCX && bl = RBX && rh = RDX && rl = RAX); emit (Psubl_rr (RAX,RBX)); emit (Psbbl_rr (RDX,RCX)) | "__builtin_mull", [BA(IR a); BA(IR b)], - BR_splitlong(BR(IR rh), BR(IR rl)) -> + BR_splitlong(IR rh, IR rl) -> assert (a = RAX && b = RDX && rh = RDX && rl = RAX); emit (Pmull_r RDX) (* Memory accesses *) diff --git a/x86/Asmgen.v b/x86/Asmgen.v index dedbfbe6b8..7d33550f20 100644 --- a/x86/Asmgen.v +++ b/x86/Asmgen.v @@ -41,7 +41,9 @@ Definition freg_of (r: mreg) : res freg := Definition mk_mov (rd rs: preg) (k: code) : res code := match rd, rs with - | IR rd, IR rs => OK (Pmov_rr rd rs :: k) + | IR rd, IR rs => + assertion (subtype (Preg.type rs) (Preg.type rd)); + OK (Pmov_rr rd rs :: k) | FR rd, FR rs => OK (Pmovsd_ff rd rs :: k) | _, _ => Error(msg "Asmgen.mk_mov") end. @@ -54,6 +56,7 @@ Definition mk_shrximm (n: int) (k: code) : res code := Psarl_ri RAX n :: k). Definition mk_shrxlimm (n: int) (k: code) : res code := + assertion (Archi.ptr64); OK (if Int.eq n Int.zero then Pmov_rr RAX RAX :: k else Pcqto :: Pshrq_ri RDX (Int.sub (Int.repr 64) n) :: @@ -86,33 +89,27 @@ Definition mk_storebyte (addr: addrmode) (rs: ireg) (k: code) := (** Accessing slots in the stack frame. *) -Definition loadind (base: ireg) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) := +Definition loadind (base: ireg) (ofs: ptrofs) (q: quantity) (dst: mreg) (k: code) := let a := Addrmode (Some base) None (inl _ (Ptrofs.unsigned ofs)) in - match ty, preg_of dst with - | Tint, IR r => OK (Pmovl_rm r a :: k) - | Tlong, IR r => OK (Pmovq_rm r a :: k) - | Tsingle, FR r => OK (Pmovss_fm r a :: k) - | Tsingle, ST0 => OK (Pflds_m a :: k) - | Tfloat, FR r => OK (Pmovsd_fm r a :: k) - | Tfloat, ST0 => OK (Pfldl_m a :: k) - | Tany32, IR r => if Archi.ptr64 then Error (msg "Asmgen.loadind1") else OK (Pmov_rm_a r a :: k) - | Tany64, IR r => if Archi.ptr64 then OK (Pmov_rm_a r a :: k) else Error (msg "Asmgen.loadind2") - | Tany64, FR r => OK (Pmovsd_fm_a r a :: k) + match q, preg_of dst with + | Q32, IR r => OK (Pmovl_rm_a r a :: k) + | Q64, IR r => if Archi.ptr64 then OK (Pmovq_rm_a r a :: k) else Error (msg "Asmgen.loadind2") + | Q32, FR r => OK (Pmovss_fm_a r a :: k) + | Q32, ST0 => OK (Pflds_m_a a :: k) + | Q64, FR r => OK (Pmovsd_fm_a r a :: k) + | Q64, ST0 => OK (Pfldl_m_a a :: k) | _, _ => Error (msg "Asmgen.loadind") end. -Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) := +Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (q: quantity) (k: code) := let a := Addrmode (Some base) None (inl _ (Ptrofs.unsigned ofs)) in - match ty, preg_of src with - | Tint, IR r => OK (Pmovl_mr a r :: k) - | Tlong, IR r => OK (Pmovq_mr a r :: k) - | Tsingle, FR r => OK (Pmovss_mf a r :: k) - | Tsingle, ST0 => OK (Pfstps_m a :: k) - | Tfloat, FR r => OK (Pmovsd_mf a r :: k) - | Tfloat, ST0 => OK (Pfstpl_m a :: k) - | Tany32, IR r => if Archi.ptr64 then Error (msg "Asmgen.storeind1") else OK (Pmov_mr_a a r :: k) - | Tany64, IR r => if Archi.ptr64 then OK (Pmov_mr_a a r :: k) else Error (msg "Asmgen.storeind2") - | Tany64, FR r => OK (Pmovsd_mf_a a r :: k) + match q, preg_of src with + | Q32, IR r => OK (Pmovl_mr_a a r :: k) + | Q64, IR r => if Archi.ptr64 then OK (Pmovq_mr_a a r :: k) else Error (msg "Asmgen.storeind2") + | Q32, FR r => OK (Pmovss_mf_a a r :: k) + | Q32, ST0 => OK (Pfstps_m_a a :: k) + | Q64, FR r => OK (Pmovsd_mf_a a r :: k) + | Q64, ST0 => OK (Pfstpl_m_a a :: k) | _, _ => Error (msg "Asmgen.storeind") end. @@ -317,6 +314,7 @@ Definition transl_op do r <- ireg_of res; OK ((if Int.eq_dec n Int.zero then Pxorl_r r else Pmovl_ri r n) :: k) | Olongconst n, nil => + assertion Archi.ptr64; do r <- ireg_of res; OK ((if Int64.eq_dec n Int64.zero then Pxorq_r r else Pmovq_ri r n) :: k) | Ofloatconst f, nil => @@ -433,105 +431,135 @@ Definition transl_op OK (Pleal r (normalize_addrmode_32 am) :: k) (* 64-bit integer operations *) | Olowlong, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pmovls_rr r :: k) | Ocast32signed, a1 :: nil => + assertion Archi.ptr64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pmovsl_rr r r1 :: k) | Ocast32unsigned, a1 :: nil => + assertion Archi.ptr64; do r1 <- ireg_of a1; do r <- ireg_of res; OK (Pmovzl_rr r r1 :: k) | Onegl, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pnegq r :: k) | Oaddlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Paddq_ri r n :: k) | Osubl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; do r2 <- ireg_of a2; OK (Psubq_rr r r2 :: k) | Omull, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pimulq_rr r r2 :: k) | Omullimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pimulq_ri r n :: k) | Omullhs, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq res DX); do r2 <- ireg_of a2; OK (Pimulq_r r2 :: k) | Omullhu, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq res DX); do r2 <- ireg_of a2; OK (Pmulq_r r2 :: k) | Odivl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq a2 CX); assertion (mreg_eq res AX); OK(Pcqto :: Pidivq RCX :: k) | Odivlu, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq a2 CX); assertion (mreg_eq res AX); OK(Pxorq_r RDX :: Pdivq RCX :: k) | Omodl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq a2 CX); assertion (mreg_eq res DX); OK(Pcqto :: Pidivq RCX :: k) | Omodlu, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq a2 CX); assertion (mreg_eq res DX); OK(Pxorq_r RDX :: Pdivq RCX :: k) | Oandl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pandq_rr r r2 :: k) | Oandlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pandq_ri r n :: k) | Oorl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; do r2 <- ireg_of a2; OK (Porq_rr r r2 :: k) | Oorlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Porq_ri r n :: k) | Oxorl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; do r2 <- ireg_of a2; OK (Pxorq_rr r r2 :: k) | Oxorlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pxorq_ri r n :: k) | Onotl, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pnotq r :: k) | Oshll, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); assertion (mreg_eq a2 CX); do r <- ireg_of res; OK (Psalq_rcl r :: k) | Oshllimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Psalq_ri r n :: k) | Oshrl, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); assertion (mreg_eq a2 CX); do r <- ireg_of res; OK (Psarq_rcl r :: k) | Oshrlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Psarq_ri r n :: k) | Oshrxlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 AX); assertion (mreg_eq res AX); mk_shrxlimm n k | Oshrlu, a1 :: a2 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); assertion (mreg_eq a2 CX); do r <- ireg_of res; OK (Pshrq_rcl r :: k) | Oshrluimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Pshrq_ri r n :: k) | Ororlimm n, a1 :: nil => + assertion Archi.ptr64; assertion (mreg_eq a1 res); do r <- ireg_of res; OK (Prorq_ri r n :: k) | Oleal addr, _ => + assertion Archi.ptr64; do am <- transl_addressing addr args; do r <- ireg_of res; OK (match normalize_addrmode_64 am with | (am', None) => Pleaq r am' :: k @@ -587,10 +615,12 @@ Definition transl_op | Osingleofint, a1 :: nil => do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsi2ss_fr r r1 :: k) | Olongoffloat, a1 :: nil => + assertion Archi.ptr64; do r <- ireg_of res; do r1 <- freg_of a1; OK (Pcvttsd2sl_rf r r1 :: k) | Ofloatoflong, a1 :: nil => do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsl2sd_fr r r1 :: k) | Olongofsingle, a1 :: nil => + assertion Archi.ptr64; do r <- ireg_of res; do r1 <- freg_of a1; OK (Pcvttss2sl_rf r r1 :: k) | Osingleoflong, a1 :: nil => do r <- freg_of res; do r1 <- ireg_of a1; OK (Pcvtsl2ss_fr r r1 :: k) @@ -606,6 +636,7 @@ Definition transl_op Definition transl_load (chunk: memory_chunk) (addr: addressing) (args: list mreg) (dest: mreg) (k: code) : res code := + assertion (subtype (type_of_chunk chunk) (Mreg.type dest)); do am <- transl_addressing addr args; match chunk with | Mint8unsigned => @@ -654,16 +685,16 @@ Definition transl_store (chunk: memory_chunk) Definition transl_instr (f: Mach.function) (i: Mach.instruction) (ax_is_parent: bool) (k: code) := match i with - | Mgetstack ofs ty dst => - loadind RSP ofs ty dst k - | Msetstack src ofs ty => - storeind src RSP ofs ty k - | Mgetparam ofs ty dst => + | Mgetstack ofs q dst => + loadind RSP ofs q dst k + | Msetstack src ofs q => + storeind src RSP ofs q k + | Mgetparam ofs q dst => if ax_is_parent then - loadind RAX ofs ty dst k + loadind RAX ofs q dst k else - (do k1 <- loadind RAX ofs ty dst k; - loadind RSP f.(fn_link_ofs) Tptr AX k1) + (do k1 <- loadind RAX ofs q dst k; + loadind RSP f.(fn_link_ofs) (quantity_of_typ Tptr) AX k1) | Mop op args res => transl_op op args res k | Mload chunk addr args dst => @@ -700,8 +731,8 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction) Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool := match i with - | Msetstack src ofs ty => before - | Mgetparam ofs ty dst => negb (mreg_eq dst AX) + | Msetstack src ofs q => before + | Mgetparam ofs q dst => negb (mreg_eq dst AX) | _ => false end. diff --git a/x86/Asmgenproof.v b/x86/Asmgenproof.v index 3aa87a4c10..d408880139 100644 --- a/x86/Asmgenproof.v +++ b/x86/Asmgenproof.v @@ -72,7 +72,7 @@ Qed. Lemma exec_straight_exec: forall fb f c ep tf tc c' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs#PC) fb f c ep tf tc -> exec_straight tge tf tc rs m c' rs' m' -> plus step tge (State rs m) E0 (State rs' m'). Proof. @@ -84,10 +84,10 @@ Qed. Lemma exec_straight_at: forall fb f c ep tf tc c' ep' tc' rs m rs' m', - transl_code_at_pc ge (rs PC) fb f c ep tf tc -> + transl_code_at_pc ge (rs#PC) fb f c ep tf tc -> transl_code f c' ep' = OK tc' -> exec_straight tge tf tc rs m tc' rs' m' -> - transl_code_at_pc ge (rs' PC) fb f c' ep' tf tc'. + transl_code_at_pc ge (rs'#PC) fb f c' ep' tf tc'. Proof. intros. inv H. exploit exec_straight_steps_2; eauto. @@ -214,7 +214,22 @@ Remark transl_op_label: transl_op op args r k = OK c -> tail_nolabel k c. Proof. - unfold transl_op; intros. destruct op; TailNoLabel. + unfold transl_op; intros. + + Ltac MyTailNoLabel := + match goal with + | [ |- tail_nolabel _ (_ :: _) ] => apply tail_nolabel_cons; [auto; exact I | MyTailNoLabel] + | [ H: Error _ = OK _ |- _ ] => discriminate + | [ H: assertion_failed = OK _ |- _ ] => discriminate + | [ H: OK _ = OK _ |- _ ] => inv H; MyTailNoLabel + | [ H: bind _ _ = OK _ |- _ ] => monadInv H; MyTailNoLabel + | [ H: (if ?x then _ else _) = OK _ |- _ ] => destruct x; MyTailNoLabel + | [ H: match ?x with nil => _ | _ :: _ => _ end = OK _ |- _ ] => destruct x; MyTailNoLabel + | _ => eauto with labels + end. + + destruct op; MyTailNoLabel. + destruct (Int.eq_dec n Int.zero); TailNoLabel. destruct (Int64.eq_dec n Int64.zero); TailNoLabel. destruct (Float.eq_dec n Float.zero); TailNoLabel. @@ -306,11 +321,11 @@ Lemma find_label_goto_label: forall f tf lbl rs m c' b ofs, Genv.find_funct_ptr ge b = Some (Internal f) -> transf_function f = OK tf -> - rs PC = Vptr b ofs -> + rs#PC = Vptr b ofs -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> exists tc', exists rs', goto_label tf lbl rs m = Next rs' m - /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc' + /\ transl_code_at_pc ge (rs'#PC) b f c' false tf tc' /\ forall r, r <> PC -> rs'#r = rs#r. Proof. intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2. @@ -366,7 +381,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (FIND: Genv.find_funct_ptr ge fb = Some (Internal f)) (MEXT: Mem.extends m m') - (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc) + (AT: transl_code_at_pc ge (rs#PC) fb f c ep tf tc) (AG: agree ms sp rs) (AXP: ep = true -> rs#RAX = parent_sp s), match_states (Mach.State s fb sp c ms m) @@ -376,8 +391,8 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = Vptr fb Ptrofs.zero) - (ATLR: rs RA = parent_ra s), + (ATPC: rs#PC = Vptr fb Ptrofs.zero) + (ATLR: rs#RA = parent_ra s), match_states (Mach.Callstate s fb ms m) (Asm.State rs m') | match_states_return: @@ -385,7 +400,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (STACKS: match_stack ge s) (MEXT: Mem.extends m m') (AG: agree ms (parent_sp s) rs) - (ATPC: rs PC = parent_ra s), + (ATPC: rs#PC = parent_ra s), match_states (Mach.Returnstate s ms m) (Asm.State rs m'). @@ -394,7 +409,7 @@ Lemma exec_straight_steps: match_stack ge s -> Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> (forall k c (TR: transl_instr f i ep k = OK c), exists rs2, exec_straight tge tf c rs1 m1' k rs2 m2' @@ -417,7 +432,7 @@ Lemma exec_straight_steps_goto: Mem.extends m2 m2' -> Genv.find_funct_ptr ge fb = Some (Internal f) -> Mach.find_label lbl f.(Mach.fn_code) = Some c' -> - transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc -> + transl_code_at_pc ge (rs1#PC) fb f (i :: c) ep tf tc -> it1_is_parent ep i = false -> (forall k c (TR: transl_instr f i ep k = OK c), exists jmp, exists k', exists rs2, @@ -475,7 +490,7 @@ Proof. - (* Mlabel *) left; eapply exec_straight_steps; eauto; intros. - monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto. + monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto with asmgen. split. apply agree_nextinstr; auto. simpl; congruence. - (* Mgetstack *) @@ -490,7 +505,7 @@ Proof. - (* Msetstack *) unfold store_stack in H. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (Regmap.get src rs) (rs0#(preg_of src))). eapply preg_val; eauto. 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. @@ -499,7 +514,7 @@ Proof. split. eapply agree_undef_regs; eauto. simpl; intros. rewrite Q; auto with asmgen. Local Transparent destroyed_by_setstack. - destruct ty; simpl; intuition congruence. + destruct q; simpl; intuition congruence. - (* Mgetparam *) assert (f0 = f) by congruence; subst f0. @@ -539,7 +554,7 @@ Opaque loadind. 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]]]. - assert (S: Val.lessdef v (rs2 (preg_of res))) by (eapply Val.lessdef_trans; eauto). + assert (S: Val.lessdef v (rs2#(preg_of res))) by (eapply Val.lessdef_trans; eauto). exists rs2; split. eauto. split. eapply agree_set_undef_mreg; eauto. simpl; congruence. @@ -561,7 +576,7 @@ Opaque loadind. rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1. intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A. - assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto. + assert (Val.lessdef (Regmap.get src rs) (rs0#(preg_of src))). eapply preg_val; eauto. exploit Mem.storev_extends; eauto. intros [m2' [C D]]. left; eapply exec_straight_steps; eauto. intros. simpl in TR. @@ -577,10 +592,10 @@ Opaque loadind. eapply transf_function_no_overflow; eauto. destruct ros as [rf|fid]; simpl in H; monadInv H5. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0#x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto. generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x). @@ -592,9 +607,10 @@ Opaque loadind. simpl. eauto. econstructor; eauto. econstructor; eauto. - eapply agree_sp_def; eauto. + eapply agree_sp_def; eauto. eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simplifs. - Simplifs. rewrite <- H2. auto. + Simplifs; rewrite H7; reflexivity. + rewrite <- H2. Simplifs. + (* Direct call *) generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1. assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x). @@ -606,9 +622,9 @@ Opaque loadind. simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. eauto. econstructor; eauto. econstructor; eauto. - eapply agree_sp_def; eauto. + eapply agree_sp_def; eauto. eapply agree_sp_type; eauto. simpl. eapply agree_exten; eauto. intros. Simplifs. - Simplifs. rewrite <- H2. auto. + Simplifs. rewrite <- H2. Simplifs. - (* Mtailcall *) assert (f0 = f) by congruence. subst f0. @@ -623,46 +639,62 @@ Opaque loadind. exploit Mem.free_parallel_extends; eauto. intros [m2' [E F]]. destruct ros as [rf|fid]; simpl in H; monadInv H7. + (* Indirect call *) - assert (rs rf = Vptr f' Ptrofs.zero). - destruct (rs rf); try discriminate. + assert (regmap_get rf rs = Vptr f' Ptrofs.zero). + destruct (regmap_get rf rs); try discriminate. revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. - assert (rs0 x0 = Vptr f' Ptrofs.zero). + assert (rs0#x0 = Vptr f' Ptrofs.zero). exploit ireg_val; eauto. rewrite H7; intros LD; inv LD; auto. generalize (code_tail_next_int _ _ _ _ NOOV H8). intro CT1. left; econstructor; split. eapply plus_left. eapply exec_step_internal. eauto. eapply functions_transl; eauto. eapply find_instr_tail; eauto. - simpl. replace (chunk_of_type Tptr) with Mptr in * by (unfold Tptr, Mptr; destruct Archi.ptr64; auto). + simpl. + replace (chunk_of_quantity (quantity_of_typ Tptr)) with Mptr_any in * by (unfold Tptr, Mptr_any; destruct Archi.ptr64; auto). rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. apply star_one. eapply exec_step_internal. - transitivity (Val.offset_ptr rs0#PC Ptrofs.one). auto. rewrite <- H4. simpl. eauto. + transitivity (Val.offset_ptr rs0#PC Ptrofs.one). + rewrite nextinstr_pc. f_equal; Simplifs. + rewrite <- H4. simpl. eauto. eapply functions_transl; eauto. eapply find_instr_tail; eauto. simpl. eauto. traceEq. econstructor; eauto. apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto. - eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. - Simplifs. rewrite Pregmap.gso; auto. + eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. eapply parent_sp_type; eauto. + eapply Val.has_subtype with (ty1 := Tptr), parent_sp_type; eauto. + Simplifs. + rewrite nextinstr_inv, !Pregmap.gso by auto with asmgen. + rewrite Pregmap.gso. try rewrite H9; reflexivity. generalize (preg_of_not_SP rf). rewrite (ireg_of_eq _ _ EQ1). congruence. + Simplifs; + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto; + eapply Val.has_subtype with (ty1 := Tptr), parent_ra_type; eauto. + (* Direct call *) generalize (code_tail_next_int _ _ _ _ NOOV H8). intro CT1. left; econstructor; split. eapply plus_left. eapply exec_step_internal. eauto. eapply functions_transl; eauto. eapply find_instr_tail; eauto. - simpl. replace (chunk_of_type Tptr) with Mptr in * by (unfold Tptr, Mptr; destruct Archi.ptr64; auto). + simpl. + replace (chunk_of_quantity (quantity_of_typ Tptr)) with Mptr_any in * by (unfold Tptr, Mptr_any; destruct Archi.ptr64; auto). rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. apply star_one. eapply exec_step_internal. - transitivity (Val.offset_ptr rs0#PC Ptrofs.one). auto. rewrite <- H4. simpl. eauto. + transitivity (Val.offset_ptr rs0#PC Ptrofs.one). + rewrite nextinstr_pc. f_equal; Simplifs. + rewrite <- H4. simpl. eauto. eapply functions_transl; eauto. eapply find_instr_tail; eauto. simpl. eauto. traceEq. econstructor; eauto. apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto. - eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. + eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. eapply parent_sp_type; eauto. + eapply Val.has_subtype with (ty1 := Tptr), parent_sp_type; eauto. rewrite Pregmap.gss. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. auto. + Simplifs; + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto; + eapply Val.has_subtype with (ty1 := Tptr), parent_ra_type; eauto. - (* Mbuiltin *) - inv AT. monadInv H4. + inv AT. monadInv H5. exploit functions_transl; eauto. intro FN. - generalize (transf_function_no_overflow _ _ H3); intro NOOV. + generalize (transf_function_no_overflow _ _ H4); intro NOOV. exploit builtin_args_match; eauto. intros [vargs' [P Q]]. exploit external_call_mem_extends; eauto. intros [vres' [m2' [A [B [C D]]]]]. @@ -677,13 +709,20 @@ Opaque loadind. instantiate (2 := tf); instantiate (1 := x). unfold nextinstr_nf, nextinstr. rewrite Pregmap.gss. rewrite undef_regs_other. rewrite set_res_other. rewrite undef_regs_other_2. - rewrite <- H1. simpl. econstructor; eauto. + rewrite <- H2. simpl. econstructor; eauto. eapply code_tail_next_int; eauto. rewrite preg_notin_charact. intros. auto with asmgen. auto with asmgen. simpl; intros. intuition congruence. apply agree_nextinstr_nf. eapply agree_set_res; auto. eapply agree_undef_regs; eauto. intros; apply undef_regs_other_2; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + destruct res; simpl in H1; InvBooleans; simpl; auto. + eapply Val.has_subtype; eauto. + split. + destruct vres'; simpl; auto; destruct (mreg_type lo); auto; congruence. + destruct vres'; simpl; auto; destruct (mreg_type hi); auto; congruence. congruence. - (* Mgoto *) @@ -706,7 +745,7 @@ Opaque loadind. intros. simpl in TR. destruct (transl_cond_correct tge tf cond args _ _ rs0 m' TR) as [rs' [A [B C]]]. - rewrite EC in B. + unfold pregmap_read in EC. rewrite EC in B. destruct (testcond_for_condition cond); simpl in *. (* simple jcc *) exists (Pjcc c1 lbl); exists k; exists rs'. @@ -725,7 +764,7 @@ Opaque loadind. (* second jcc jumps *) exists (Pjcc c2 lbl); exists k; exists (nextinstr rs'). split. eapply exec_straight_trans. eexact A. - eapply exec_straight_one. simpl. rewrite TC1. auto. auto. + eapply exec_straight_one. simpl. rewrite TC1. auto. auto with asmgen. split. eapply agree_exten; eauto. intros; Simplifs. simpl. rewrite eval_testcond_nextinstr. rewrite TC2. @@ -744,12 +783,12 @@ Opaque loadind. left; eapply exec_straight_steps; eauto. intros. simpl in TR. destruct (transl_cond_correct tge tf cond args _ _ rs0 m' TR) as [rs' [A [B C]]]. - rewrite EC in B. + unfold pregmap_read in EC. rewrite EC in B. destruct (testcond_for_condition cond); simpl in *. (* simple jcc *) econstructor; split. eapply exec_straight_trans. eexact A. - apply exec_straight_one. simpl. rewrite B. eauto. auto. + apply exec_straight_one. simpl. rewrite B. eauto. auto with asmgen. split. apply agree_nextinstr. eapply agree_exten; eauto. simpl; congruence. (* jcc ; jcc *) @@ -759,7 +798,7 @@ Opaque loadind. econstructor; split. eapply exec_straight_trans. eexact A. eapply exec_straight_two. simpl. rewrite TC1. eauto. auto. - simpl. rewrite eval_testcond_nextinstr. rewrite TC2. eauto. auto. auto. + simpl. rewrite eval_testcond_nextinstr. rewrite TC2. eauto. auto with asmgen. auto with asmgen. split. apply agree_nextinstr. apply agree_nextinstr. eapply agree_exten; eauto. simpl; congruence. (* jcc2 *) @@ -770,7 +809,7 @@ Opaque loadind. apply exec_straight_one. simpl. rewrite TC1; rewrite TC2. destruct b1. simpl in *. subst b2. auto. auto. - auto. + auto with asmgen. split. apply agree_nextinstr. eapply agree_exten; eauto. rewrite H1; congruence. @@ -781,6 +820,7 @@ Opaque loadind. generalize (transf_function_no_overflow _ _ H5); intro NOOV. set (rs1 := rs0 #RAX <- Vundef #RDX <- Vundef). exploit (find_label_goto_label f tf lbl rs1); eauto. + unfold rs1. erewrite H3. Simplifs. intros [tc' [rs' [A [B C]]]]. exploit ireg_val; eauto. rewrite H. intros LD; inv LD. left; econstructor; split. @@ -810,14 +850,23 @@ Transparent destroyed_by_jumptable. left; econstructor; split. eapply plus_left. eapply exec_step_internal. eauto. eapply functions_transl; eauto. eapply find_instr_tail; eauto. - simpl. rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. + simpl. + replace (chunk_of_quantity (quantity_of_typ Tptr)) with Mptr_any in * by (unfold Tptr, Mptr_any; destruct Archi.ptr64; auto). + rewrite C. rewrite A. rewrite <- (sp_val _ _ _ AG). rewrite E. eauto. apply star_one. eapply exec_step_internal. - transitivity (Val.offset_ptr rs0#PC Ptrofs.one). auto. rewrite <- H3. simpl. eauto. + transitivity (Val.offset_ptr rs0#PC Ptrofs.one). + rewrite nextinstr_pc. f_equal; Simplifs. + rewrite <- H3. simpl. eauto with asmgen. eapply functions_transl; eauto. eapply find_instr_tail; eauto. simpl. eauto. traceEq. constructor; auto. apply agree_set_other; auto. apply agree_nextinstr. apply agree_set_other; auto. - eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. + eapply agree_change_sp; eauto. eapply parent_sp_def; eauto. eapply parent_sp_type; eauto. + eapply Val.has_subtype with (ty1 := Tptr), parent_sp_type; eauto. + Simplifs; + rewrite nextinstr_inv, Pregmap.gss by discriminate; auto; + rewrite !Preg.chunk_of_reg_type, Val.load_result_same; try rewrite Val.load_result_same; + auto; eapply Val.has_subtype with (ty1 := Tptr), parent_ra_type; eauto. - (* internal function *) exploit functions_translated; eauto. intros [tf [A B]]. monadInv B. @@ -835,7 +884,7 @@ Transparent destroyed_by_jumptable. apply plus_one. econstructor; eauto. simpl. rewrite Ptrofs.unsigned_zero. simpl. eauto. simpl. rewrite C. simpl in F, P. - replace (chunk_of_type Tptr) with Mptr in F, P by (unfold Tptr, Mptr; destruct Archi.ptr64; auto). + replace (chunk_of_quantity (quantity_of_typ Tptr)) with Mptr_any in * by (unfold Tptr, Mptr_any; destruct Archi.ptr64; auto). rewrite (sp_val _ _ _ AG) in F. rewrite F. rewrite ATLR. rewrite P. eauto. econstructor; eauto. @@ -847,8 +896,12 @@ Transparent destroyed_by_jumptable. Transparent destroyed_at_function_entry. apply agree_undef_regs with rs0; eauto. simpl; intros. apply Pregmap.gso; auto with asmgen. tauto. - congruence. - intros. Simplifs. eapply agree_sp; eauto. + congruence. apply Val.Vptr_has_type. + eapply Val.has_subtype, Val.Vptr_has_type; auto. + intros. Simplifs. + erewrite agree_sp; eauto. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + eapply Val.has_subtype with (ty1 := Tptr), parent_sp_type; eauto. - (* external function *) exploit functions_translated; eauto. @@ -863,6 +916,19 @@ Transparent destroyed_at_function_entry. econstructor; eauto. unfold loc_external_result. apply agree_set_other; auto. apply agree_set_pair; auto. apply agree_undef_caller_save_regs; auto. + exploit external_call_well_typed; eauto; intro. + exploit loc_result_has_type; eauto; intro. + unfold Val.has_type_rpair in H3. + unfold loc_result, proj_sig_res, loc_result_64, loc_result_32 in *. + destruct Archi.ptr64, (sig_res (ef_sig ef)). + destruct t0; simpl in *; auto. + simpl in *; auto. + destruct t0; simpl in *; auto. + simpl in *; auto. + rewrite Pregmap.gss. rewrite ATLR. + exploit parent_ra_type; eauto; intro; + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto; + eapply Val.has_subtype with (ty1 := Tptr), parent_ra_type; eauto. - (* return *) inv STACKS. simpl in *. @@ -885,7 +951,9 @@ Proof. apply Mem.extends_refl. split. reflexivity. simpl. unfold Vnullptr; destruct Archi.ptr64; congruence. - intros. rewrite Regmap.gi. auto. + constructor. + intros. unfold regmap_get, Regmap.get. rewrite FragBlock.gi. auto. + Simplifs. unfold Genv.symbol_address. rewrite (match_program_main TRANSF). rewrite symbols_preserved. diff --git a/x86/Asmgenproof1.v b/x86/Asmgenproof1.v index c5b0342680..e879abfc08 100644 --- a/x86/Asmgenproof1.v +++ b/x86/Asmgenproof1.v @@ -55,24 +55,96 @@ Lemma nextinstr_nf_set_preg: (nextinstr_nf (rs#(preg_of m) <- v))#PC = Val.offset_ptr rs#PC Ptrofs.one. Proof. intros. unfold nextinstr_nf. - transitivity (nextinstr (rs#(preg_of m) <- v) PC). auto. + transitivity ((nextinstr (rs#(preg_of m) <- v))#PC). + simpl. unfold nextinstr. + rewrite Pregmap.gss, 6 Pregmap.gso, Pregmap.gss; auto; try discriminate. + auto with asmgen. apply nextinstr_set_preg. Qed. +Lemma nextinstr_data_pc: + forall rs r v, + r <> PC -> + (nextinstr (rs#r <- v))#PC = Val.offset_ptr rs#PC Ptrofs.one. +Proof. + intros. rewrite nextinstr_pc, Pregmap.gso; auto. +Qed. + +Lemma nextinstr_nf_data_pc: + forall rs r v, + r <> PC -> + (nextinstr_nf (rs#r <- v))#PC = Val.offset_ptr rs#PC Ptrofs.one. +Proof. + intros. unfold nextinstr_nf. + transitivity ((nextinstr (rs#r <- v))#PC). + simpl. unfold nextinstr. + rewrite Pregmap.gss, 6 Pregmap.gso, Pregmap.gss; auto; try discriminate. + auto using nextinstr_data_pc. +Qed. + +Lemma nextinstr_nf_pc: + forall rs, + (nextinstr_nf rs)#PC = Val.offset_ptr rs#PC Ptrofs.one. +Proof. + intros. unfold nextinstr_nf. + simpl. repeat (rewrite nextinstr_data_pc by discriminate; f_equal). + repeat rewrite Pregmap.gso by discriminate; auto. +Qed. + +Hint Resolve nextinstr_pc nextinstr_nf_pc nextinstr_data_pc + nextinstr_nf_data_pc nextinstr_nf_pc load_result_offset_pc: asmgen. + +Lemma nextinstr_compare_ints_pc: + forall i1 i2 rs m, + Pregmap.get PC (nextinstr (compare_ints i1 i2 rs m)) = + Val.offset_ptr (Pregmap.get PC rs) Ptrofs.one. +Proof. + intros. unfold nextinstr, compare_ints. + rewrite Pregmap.gss. repeat rewrite Pregmap.gso by discriminate. + apply load_result_offset_pc. +Qed. + +Lemma nextinstr_compare_longs_pc: + forall i1 i2 rs m, + Pregmap.get PC (nextinstr (compare_longs i1 i2 rs m)) = + Val.offset_ptr (Pregmap.get PC rs) Ptrofs.one. +Proof. + intros. unfold nextinstr, compare_longs. + rewrite Pregmap.gss. repeat rewrite Pregmap.gso by discriminate. + apply load_result_offset_pc. +Qed. + +(** A subtraction's result always fits into an integer register *) + +Lemma load_result_sub: + forall (i: ireg) v1 v2, + Val.load_result (Preg.chunk_of i) (Val.sub v1 v2) = (Val.sub v1 v2). +Proof. + intros. destruct Archi.ptr64 eqn:PTR64. + simpl; rewrite PTR64; destruct i; auto. + destruct v1, v2; simpl; auto; simpl; + rewrite PTR64; destruct i; simpl; auto; + rewrite PTR64; auto; destruct (eq_block _ _); auto. +Qed. + (** Useful simplification tactic *) Ltac Simplif := match goal with - | [ |- nextinstr_nf _ _ = _ ] => + | [ |- Pregmap.get PC (nextinstr _ # _ <- _) = Val.offset_ptr (Pregmap.get PC _) Ptrofs.one ] => + rewrite nextinstr_pc, Pregmap.gso by auto with asmgen; auto using load_result_offset_pc + | [ |- Pregmap.get PC (nextinstr_nf _ # _ <- _) = Val.offset_ptr (Pregmap.get PC _) Ptrofs.one ] => + rewrite nextinstr_nf_data_pc by auto with asmgen; auto + | [ |- (nextinstr_nf _) # _ = _ ] => ((rewrite nextinstr_nf_inv by auto with asmgen) || (rewrite nextinstr_nf_inv1 by auto with asmgen)); auto - | [ |- nextinstr _ _ = _ ] => + | [ |- (nextinstr _) # _ = _ ] => ((rewrite nextinstr_inv by auto with asmgen) || (rewrite nextinstr_inv1 by auto with asmgen)); auto + | [ |- Pregmap.get ?x (nextinstr _) = Pregmap.get ?x _ ] => + unfold nextinstr | [ |- Pregmap.get ?x (Pregmap.set ?x _ _) = _ ] => - rewrite Pregmap.gss; auto - | [ |- Pregmap.set ?x _ _ ?x = _ ] => - rewrite Pregmap.gss; auto + rewrite Pregmap.gss; auto using Pregmap.get_load_result | [ |- Pregmap.get _ (Pregmap.set _ _ _) = _ ] => rewrite Pregmap.gso by (auto with asmgen); auto | [ |- Pregmap.set _ _ _ _ = _ ] => @@ -101,10 +173,14 @@ Proof. unfold mk_mov; intros. destruct rd; try (monadInv H); destruct rs; monadInv H. (* mov *) - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. Simplifs. intros; Simplifs. + econstructor. split. apply exec_straight_one. simpl. eauto. Simplifs. + split. unfold nextinstr. + Simplifs; + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto; + eauto using Val.has_subtype, Pregmap.get_has_type. + intros; Simplifs. (* movsd *) - econstructor. split. apply exec_straight_one. simpl. eauto. auto. + econstructor. split. apply exec_straight_one. simpl. eauto. Simplifs. split. Simplifs. intros; Simplifs. Qed. @@ -210,19 +286,38 @@ Proof. set (rs3 := nextinstr (rs2#RCX <- (Vint x'))). set (rs4 := nextinstr (if Int.lt x Int.zero then rs3#RAX <- (Vint x') else rs3)). set (rs5 := nextinstr_nf (rs4#RAX <- (Val.shr rs4#RAX (Vint n)))). - assert (rs3#RAX = Vint x). unfold rs3. Simplifs. + assert (rs3#RAX = Vint x). unfold rs3, rs2, compare_ints. Simplifs. assert (rs3#RCX = Vint x'). unfold rs3. Simplifs. exists rs5. split. - apply exec_straight_step with rs2 m. simpl. rewrite A. simpl. rewrite Int.and_idem. auto. auto. + apply exec_straight_step with rs2 m. simpl. rewrite A. simpl. rewrite Int.and_idem. auto. + unfold rs2. apply nextinstr_compare_ints_pc. apply exec_straight_step with rs3 m. simpl. - change (rs2 RAX) with (rs1 RAX). rewrite A. simpl. - rewrite Int.repr_unsigned. rewrite Int.add_zero_l. auto. auto. + replace (rs2#RAX) with (rs1#RAX). rewrite A. simpl. + rewrite Int.repr_unsigned. rewrite Int.add_zero_l. auto. + unfold rs2. unfold compare_ints. unfold nextinstr. Simplifs. + rewrite !Pregmap.gso by discriminate; reflexivity. + unfold rs3. apply nextinstr_data_pc; discriminate. apply exec_straight_step with rs4 m. simpl. - rewrite Int.lt_sub_overflow. unfold rs4. destruct (Int.lt x Int.zero); simpl; auto. + assert (O: Pregmap.get OF rs3 = Vint (Int.sub_overflow x Int.zero Int.zero)). + { unfold rs3, rs2. Simplifs. unfold compare_ints. Simplifs. } + assert (S: Pregmap.get SF rs3 = Val.negative (Val.sub (Vint x) (Vint Int.zero))). + { unfold rs3, rs2. Simplifs. unfold compare_ints. Simplifs. } + rewrite O, S; simpl. + rewrite Int.lt_sub_overflow. unfold rs4. + destruct (Int.lt x Int.zero); try rewrite H0; auto. unfold rs4. destruct (Int.lt x Int.zero); simpl; auto. - apply exec_straight_one. auto. auto. + apply nextinstr_data_pc; discriminate. + apply nextinstr_pc. + apply exec_straight_one. auto. + unfold rs5. rewrite nextinstr_nf_data_pc; auto with asmgen. split. unfold rs5. Simplifs. unfold rs4. rewrite nextinstr_inv; auto with asmgen. - destruct (Int.lt x Int.zero). rewrite Pregmap.gss. rewrite A; auto. rewrite A; rewrite H; auto. + destruct (Int.lt x Int.zero). rewrite Pregmap.gss. + rewrite Preg.chunk_of_reg_type, Val.load_result_same + by (simpl; destruct Archi.ptr64; simpl; destruct (Int.ltu _ _); simpl; auto). + rewrite A; auto. rewrite A; rewrite H. + rewrite Preg.chunk_of_reg_type, Val.load_result_same + by (simpl; destruct Archi.ptr64; simpl; destruct (Int.ltu _ _); simpl; auto). + auto. intros. unfold rs5. Simplifs. unfold rs4. Simplifs. transitivity (rs3#r). destruct (Int.lt x Int.zero). Simplifs. auto. unfold rs3. Simplifs. unfold rs2. Simplifs. @@ -241,12 +336,12 @@ Lemma mk_shrxlimm_correct: /\ forall r, data_preg r = true -> r <> RAX -> r <> RDX -> rs2#r = rs1#r. Proof. unfold mk_shrxlimm; intros. exploit Val.shrxl_shrl_2; eauto. intros EQ. - destruct (Int.eq n Int.zero); inv H. -- econstructor; split. apply exec_straight_one. simpl; reflexivity. auto. + destruct (Int.eq n Int.zero); monadInv H. +- econstructor; split. apply exec_straight_one. simpl; reflexivity. auto with asmgen. split. Simplifs. intros; Simplifs. -- set (v1 := Val.shrl (rs1 RAX) (Vint (Int.repr 63))) in *. +- set (v1 := Val.shrl (rs1#RAX) (Vint (Int.repr 63))) in *. set (v2 := Val.shrlu v1 (Vint (Int.sub (Int.repr 64) n))) in *. - set (v3 := Val.addl (rs1 RAX) v2) in *. + set (v3 := Val.addl (rs1#RAX) v2) in *. set (v4 := Val.shrl v3 (Vint n)) in *. set (rs2 := nextinstr_nf (rs1#RDX <- v1)). set (rs3 := nextinstr_nf (rs2#RDX <- v2)). @@ -261,38 +356,70 @@ Proof. rewrite Int64.add_zero; auto. } exists rs5; split. eapply exec_straight_trans with (rs2 := rs3). - eapply exec_straight_two with (rs2 := rs2); reflexivity. + eapply exec_straight_two with (rs2 := rs2). + reflexivity. + unfold rs3, rs2, v2, v1. simpl. + repeat f_equal. Simplifs; simpl; rewrite Heqb; auto. + unfold rs2. Simplifs. + unfold rs3. Simplifs. eapply exec_straight_two with (rs2 := rs4). - simpl. rewrite X. reflexivity. reflexivity. reflexivity. reflexivity. - split. unfold rs5; Simplifs. + simpl. rewrite X. unfold rs4, rs3, rs2, v3, v2. + do 4 f_equal; Simplifs; simpl; rewrite Heqb; auto. + simpl. unfold rs5, rs4, v4, v3. + do 4 f_equal; Simplifs; simpl; rewrite Heqb; auto. + unfold rs4; auto with asmgen. + unfold rs5; auto with asmgen. + split. unfold rs5; Simplifs; simpl; rewrite Heqb; auto. intros. unfold rs5; Simplifs. unfold rs4; Simplifs. unfold rs3; Simplifs. unfold rs2; Simplifs. Qed. (** Smart constructor for integer conversions *) -Lemma mk_intconv_correct: - forall mk sem rd rs k c rs1 m, - mk_intconv mk rd rs k = OK c -> - (forall c rd rs r m, - exec_instr ge c (mk rd rs) r m = Next (nextinstr (r#rd <- (sem r#rs))) m) -> +Lemma mk_intconv_Pmovsb_rr_correct: + forall rd rs k c rs1 m, + mk_intconv Pmovsb_rr rd rs k = OK c -> + exists rs2, + exec_straight ge fn c rs1 m k rs2 m + /\ rs2#rd = Val.sign_ext 8 rs1#rs + /\ forall r, data_preg r = true -> r <> rd -> r <> RAX -> rs2#r = rs1#r. +Proof. + unfold mk_intconv; intros. destruct (Archi.ptr64 || low_ireg rs); monadInv H. + econstructor. split. apply exec_straight_one. simpl. eauto. + auto with asmgen. + split. Simplifs. destruct (rs1#rs), rd; auto; intros; Simplifs. + try (intros; Simplifs). + econstructor. split. eapply exec_straight_two. + simpl. eauto. simpl. auto. auto with asmgen. auto with asmgen. + split. Simplifs. + rewrite nextinstr_inv, Pregmap.gss by discriminate. simpl. destruct (rs1#rs), rd; auto. + intros. Simplifs. +Qed. + +Lemma mk_intconv_Pmovzb_rr_correct: + forall rd rs k c rs1 m, + mk_intconv Pmovzb_rr rd rs k = OK c -> exists rs2, exec_straight ge fn c rs1 m k rs2 m - /\ rs2#rd = sem rs1#rs + /\ rs2#rd = Val.zero_ext 8 rs1#rs /\ forall r, data_preg r = true -> r <> rd -> r <> RAX -> rs2#r = rs1#r. Proof. unfold mk_intconv; intros. destruct (Archi.ptr64 || low_ireg rs); monadInv H. - econstructor. split. apply exec_straight_one. rewrite H0. eauto. auto. - split. Simplifs. intros. Simplifs. + econstructor. split. apply exec_straight_one. simpl. eauto. + auto with asmgen. + split. Simplifs. destruct (rs1#rs), rd; auto; intros; Simplifs. + try (intros; Simplifs). econstructor. split. eapply exec_straight_two. - simpl. eauto. apply H0. auto. auto. - split. Simplifs. intros. Simplifs. + simpl. eauto. simpl. auto. auto with asmgen. auto with asmgen. + split. Simplifs. + rewrite nextinstr_inv, Pregmap.gss by discriminate. simpl. destruct (rs1#rs), rd; auto. + intros. Simplifs. Qed. (** Smart constructor for small stores *) Lemma addressing_mentions_correct: forall a r (rs1 rs2: regset), - (forall (r': ireg), r' <> r -> rs1 r' = rs2 r') -> + (forall (r': ireg), r' <> r -> rs1#r' = rs2#r') -> addressing_mentions a r = false -> eval_addrmode32 ge a rs1 = eval_addrmode32 ge a rs2. Proof. @@ -302,10 +429,18 @@ Proof. decEq. destruct ofs as [[r' sc] | ]; auto. rewrite AG; auto. destruct (ireg_eq r r'); congruence. Qed. +Lemma eval_addrmode32_type: + forall ge addr rs, + Val.has_type (eval_addrmode32 ge addr rs) Tany32. +Proof. + intros. destruct addr; simpl. + apply Val.has_subtype with (ty1 := Tint); auto using type_add. +Qed. + Lemma mk_storebyte_correct: forall addr r k c rs1 m1 m2, mk_storebyte addr r k = OK c -> - Mem.storev Mint8unsigned m1 (eval_addrmode ge addr rs1) (rs1 r) = Some m2 -> + Mem.storev Mint8unsigned m1 (eval_addrmode ge addr rs1) (rs1#r) = Some m2 -> exists rs2, exec_straight ge fn c rs1 m1 k rs2 m2 /\ forall r, data_preg r = true -> preg_notin r (if Archi.ptr64 then nil else AX :: CX :: nil) -> rs2#r = rs1#r. @@ -314,39 +449,52 @@ Proof. destruct (Archi.ptr64 || low_ireg r) eqn:E. (* low reg *) monadInv H. econstructor; split. apply exec_straight_one. - simpl. unfold exec_store. rewrite H0. eauto. auto. + simpl. unfold exec_store. rewrite H0. eauto. + unfold nextinstr_nf, nextinstr. simpl undef_regs. + rewrite Pregmap.gss, !Pregmap.gso by discriminate. + apply load_result_offset_pc. intros; Simplifs. (* high reg *) InvBooleans. rewrite H1; simpl. destruct (addressing_mentions addr RAX) eqn:E; monadInv H. (* RAX is mentioned. *) assert (r <> RCX). { red; intros; subst r; discriminate H2. } set (rs2 := nextinstr (rs1#RCX <- (eval_addrmode32 ge addr rs1))). - set (rs3 := nextinstr (rs2#RAX <- (rs1 r))). + set (rs3 := nextinstr (rs2#RAX <- (rs1#r))). econstructor; split. apply exec_straight_three with rs2 m1 rs3 m1. simpl. auto. - simpl. replace (rs2 r) with (rs1 r). auto. symmetry. unfold rs2; Simplifs. + simpl. replace (rs2#r) with (rs1#r). auto. symmetry. unfold rs2; Simplifs. simpl. unfold exec_store. unfold eval_addrmode; rewrite H1; simpl. rewrite Int.add_zero. - change (rs3 RAX) with (rs1 r). - change (rs3 RCX) with (eval_addrmode32 ge addr rs1). + replace (rs3#RAX) with (rs1#r). + replace (rs3#RCX) with (eval_addrmode32 ge addr rs1). replace (Val.add (eval_addrmode32 ge addr rs1) (Vint Int.zero)) with (eval_addrmode ge addr rs1). rewrite H0. eauto. unfold eval_addrmode in *; rewrite H1 in *. destruct (eval_addrmode32 ge addr rs1); simpl in H0; try discriminate H0. simpl. rewrite H1. rewrite Ptrofs.add_zero; auto. - auto. auto. auto. + unfold rs3, rs2. rewrite nextinstr_inv, Pregmap.gso, nextinstr_inv, Pregmap.gss by discriminate. + rewrite Preg.chunk_of_reg_type. simpl. rewrite H1. + rewrite Val.load_result_same; auto using eval_addrmode32_type. + unfold rs3, rs2. symmetry. + Simplifs; change (Preg.chunk_of RAX) with (Preg.chunk_of r); apply Pregmap.get_load_result. + unfold rs2; auto with asmgen. + unfold rs3; auto with asmgen. + auto with asmgen. intros. destruct H4. Simplifs. unfold rs3; Simplifs. unfold rs2; Simplifs. (* RAX is not mentioned *) - set (rs2 := nextinstr (rs1#RAX <- (rs1 r))). + set (rs2 := nextinstr (rs1#RAX <- (rs1#r))). econstructor; split. apply exec_straight_two with rs2 m1. simpl. auto. simpl. unfold exec_store. unfold eval_addrmode in *; rewrite H1 in *. rewrite (addressing_mentions_correct addr RAX rs2 rs1); auto. - change (rs2 RAX) with (rs1 r). rewrite H0. eauto. + replace (rs2#RAX) with (rs1#r). rewrite H0. eauto. + unfold rs2. rewrite nextinstr_inv, Pregmap.gss by discriminate. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + simpl. destruct r eqn:R; try inversion H2; try apply Pregmap.get_has_type. intros. unfold rs2; Simplifs. - auto. auto. + unfold rs2. auto with asmgen. auto with asmgen. intros. destruct H3. simpl. Simplifs. unfold rs2; Simplifs. Qed. @@ -372,9 +520,9 @@ Ltac loadind_correct_solve := end. Lemma loadind_correct: - forall (base: ireg) ofs ty dst k (rs: regset) c m v, - loadind base ofs ty dst k = OK c -> - Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v -> + forall (base: ireg) ofs q dst k (rs: regset) c m v, + loadind base ofs q dst k = OK c -> + Mem.loadv (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) = Some v -> exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of dst) = v @@ -383,30 +531,35 @@ Proof. unfold loadind; intros. set (addr := Addrmode (Some base) None (inl (ident * ptrofs) (Ptrofs.unsigned ofs))) in *. assert (eval_addrmode ge addr rs = Val.offset_ptr rs#base ofs). - { apply eval_addrmode_indexed. destruct (rs base); auto || discriminate. } + { apply eval_addrmode_indexed. destruct (rs#base); auto || discriminate. } rewrite <- H1 in H0. exists (nextinstr_nf (rs#(preg_of dst) <- v)); split. -- loadind_correct_solve; apply exec_straight_one; auto; simpl in *; unfold exec_load; rewrite ?Heqb, ?H0; auto. +- loadind_correct_solve; apply exec_straight_one; auto; simpl in *; unfold exec_load; rewrite ?Heqb, ?H0; auto with asmgen. - intuition Simplifs. + apply Mem.loadv_type in H0. + destruct q, (preg_of dst); try inversion H; auto; + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. Qed. Lemma storeind_correct: - forall (base: ireg) ofs ty src k (rs: regset) c m m', - storeind src base ofs ty k = OK c -> - Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' -> + forall (base: ireg) ofs q src k (rs: regset) c m m', + storeind src base ofs q k = OK c -> + Mem.storev (chunk_of_quantity q) m (Val.offset_ptr rs#base ofs) (rs#(preg_of src)) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' - /\ forall r, data_preg r = true -> preg_notin r (destroyed_by_setstack ty) -> rs'#r = rs#r. + /\ forall r, data_preg r = true -> preg_notin r (destroyed_by_setstack q) -> rs'#r = rs#r. Proof. unfold storeind; intros. set (addr := Addrmode (Some base) None (inl (ident * ptrofs) (Ptrofs.unsigned ofs))) in *. assert (eval_addrmode ge addr rs = Val.offset_ptr rs#base ofs). - { apply eval_addrmode_indexed. destruct (rs base); auto || discriminate. } + { apply eval_addrmode_indexed. destruct (rs#base); auto || discriminate. } rewrite <- H1 in H0. loadind_correct_solve; simpl in H0; (econstructor; split; - [apply exec_straight_one; [simpl; unfold exec_store; rewrite ?Heqb, H0;eauto|auto] + [apply exec_straight_one; [simpl; unfold exec_store; rewrite ?Heqb, H0;eauto|auto with asmgen] |simpl; intros; unfold undef_regs; repeat Simplifs]). + simpl; Simplifs. + simpl; Simplifs. Qed. (** Translation of addressing modes *) @@ -414,7 +567,7 @@ Qed. Lemma transl_addressing_mode_32_correct: forall addr args am (rs: regset) v, transl_addressing addr args = OK am -> - eval_addressing32 ge (rs RSP) addr (List.map rs (List.map preg_of args)) = Some v -> + eval_addressing32 ge (rs#RSP) addr (List.map (fun r => rs#r) (List.map preg_of args)) = Some v -> Val.lessdef v (eval_addrmode32 ge am rs). Proof. assert (A: forall id ofs, Archi.ptr64 = false -> @@ -437,7 +590,7 @@ Proof. - rewrite Val.add_commut. rewrite A by auto. auto. - rewrite Val.add_permut. rewrite Val.add_commut. apply Val.add_lessdef; auto. rewrite A; auto. - simpl. unfold Val.add; rewrite Heqb. - destruct (rs RSP); simpl; auto. + destruct (rs#RSP); simpl; auto. rewrite Int.add_zero_l. apply Val.lessdef_same; f_equal; f_equal. symmetry. transitivity (Ptrofs.repr (Ptrofs.signed i)). auto with ptrofs. auto with ints. Qed. @@ -445,7 +598,7 @@ Qed. Lemma transl_addressing_mode_64_correct: forall addr args am (rs: regset) v, transl_addressing addr args = OK am -> - eval_addressing64 ge (rs RSP) addr (List.map rs (List.map preg_of args)) = Some v -> + eval_addressing64 ge (rs#RSP) addr (List.map (fun r => rs#r) (List.map preg_of args)) = Some v -> Val.lessdef v (eval_addrmode64 ge am rs). Proof. assert (A: forall id ofs, Archi.ptr64 = true -> @@ -465,7 +618,7 @@ Proof. - rewrite Val.addl_permut. apply Val.addl_lessdef; auto. simpl; rewrite Int64.add_zero_l; auto. - apply Val.addl_lessdef; auto. apply Val.addl_lessdef; auto. - rewrite ! A by auto. auto. -- unfold Val.addl; rewrite Heqb. destruct (rs RSP); auto. simpl. +- unfold Val.addl; rewrite Heqb. destruct (rs#RSP); auto. simpl. rewrite Int64.add_zero_l. apply Val.lessdef_same; f_equal; f_equal. symmetry. transitivity (Ptrofs.repr (Ptrofs.signed i)). auto with ptrofs. auto with ints. Qed. @@ -473,7 +626,7 @@ Qed. Lemma transl_addressing_mode_correct: forall addr args am (rs: regset) v, transl_addressing addr args = OK am -> - eval_addressing ge (rs RSP) addr (List.map rs (List.map preg_of args)) = Some v -> + eval_addressing ge (rs#RSP) addr (List.map (fun r => rs#r) (List.map preg_of args)) = Some v -> Val.lessdef v (eval_addrmode ge am rs). Proof. unfold eval_addressing, eval_addrmode; intros. destruct Archi.ptr64. @@ -517,10 +670,11 @@ Lemma compare_ints_spec: /\ (forall r, data_preg r = true -> rs'#r = rs#r). Proof. intros. unfold rs'; unfold compare_ints. - split. auto. - split. auto. - split. auto. - split. auto. + split. Simplifs. apply Val.load_result_of_optbool; auto. + split. Simplifs. apply Val.load_result_of_optbool; auto. + split. Simplifs. + destruct v1, v2; simpl; auto; destruct Archi.ptr64; auto; destruct (eq_block _ _); auto. + split. Simplifs. destruct v1, v2; simpl; auto. intros. Simplifs. Qed. @@ -604,10 +758,17 @@ Lemma compare_longs_spec: /\ (forall r, data_preg r = true -> rs'#r = rs#r). Proof. intros. unfold rs'; unfold compare_longs. - split. auto. - split. auto. - split. auto. - split. auto. + split. Simplifs. + unfold Val.maketotal, Val.cmplu, option_map. + destruct (Val.cmplu_bool _ _ _ _); auto. destruct b; auto. + split. Simplifs. + unfold Val.maketotal, Val.cmplu, option_map. + destruct (Val.cmplu_bool _ _ _ _); auto. destruct b; auto. + split. Simplifs. + unfold Val.negativel. unfold Val.subl. + destruct v1, v2; simpl; auto; destruct Archi.ptr64; simpl; auto; destruct (eq_block b b0); auto. + split. Simplifs. + destruct v1, v2; simpl; auto. intros. Simplifs. Qed. @@ -706,9 +867,9 @@ Lemma compare_floats_spec: /\ (forall r, data_preg r = true -> rs'#r = rs#r). Proof. intros. unfold rs'; unfold compare_floats. - split. auto. - split. auto. - split. auto. + split. Simplifs. destruct (Float.cmp _ _ _); auto. + split. Simplifs. destruct (Float.cmp _ _ _); auto. + split. Simplifs. repeat destruct (Float.cmp _ _ _); auto. intros. Simplifs. Qed. @@ -721,9 +882,9 @@ Lemma compare_floats32_spec: /\ (forall r, data_preg r = true -> rs'#r = rs#r). Proof. intros. unfold rs'; unfold compare_floats32. - split. auto. - split. auto. - split. auto. + split. Simplifs. destruct (Float32.cmp _ _ _); auto. + split. Simplifs. destruct (Float32.cmp _ _ _); auto. + split. Simplifs. repeat destruct (Float32.cmp _ _ _); auto. intros. Simplifs. Qed. @@ -950,10 +1111,10 @@ Qed. Remark compare_floats_inv: forall vx vy rs r, r <> CR ZF -> r <> CR CF -> r <> CR PF -> r <> CR SF -> r <> CR OF -> - compare_floats vx vy rs r = rs r. + (compare_floats vx vy rs)#r = rs#r. Proof. intros. - assert (DFL: undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs r = rs r). + assert (DFL: (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs)#r = rs#r). simpl. Simplifs. unfold compare_floats; destruct vx; destruct vy; auto. Simplifs. Qed. @@ -961,10 +1122,10 @@ Qed. Remark compare_floats32_inv: forall vx vy rs r, r <> CR ZF -> r <> CR CF -> r <> CR PF -> r <> CR SF -> r <> CR OF -> - compare_floats32 vx vy rs r = rs r. + (compare_floats32 vx vy rs)#r = rs#r. Proof. intros. - assert (DFL: undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs r = rs r). + assert (DFL: (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs)#r = rs#r). simpl. Simplifs. unfold compare_floats32; destruct vx; destruct vy; auto. Simplifs. Qed. @@ -974,117 +1135,117 @@ Lemma transl_cond_correct: transl_cond cond args k = OK c -> exists rs', exec_straight ge fn c rs m k rs' m - /\ match eval_condition cond (map rs (map preg_of args)) m with + /\ match eval_condition cond (map (fun r => rs#r) (map preg_of args)) m with | None => True | Some b => eval_extcond (testcond_for_condition cond) rs' = Some b end - /\ forall r, data_preg r = true -> rs'#r = rs r. + /\ forall r, data_preg r = true -> rs'#r = rs#r. Proof. unfold transl_cond; intros. destruct cond; repeat (destruct args; try discriminate); monadInv H. - (* comp *) simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1). - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmp_bool c0 (rs x) (rs x0)) eqn:?; auto. + econstructor. split. apply exec_straight_one. simpl. eauto. apply nextinstr_compare_ints_pc. + split. destruct (Val.cmp_bool c0 (rs#x) (rs#x0)) eqn:?; auto. eapply testcond_for_signed_comparison_32_correct; eauto. intros. unfold compare_ints. Simplifs. - (* compu *) simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1). - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (rs x0)) eqn:?; auto. + econstructor. split. apply exec_straight_one. simpl. eauto. apply nextinstr_compare_ints_pc. + split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs#x) (rs#x0)) eqn:?; auto. eapply testcond_for_unsigned_comparison_32_correct; eauto. intros. unfold compare_ints. Simplifs. - (* compimm *) simpl. rewrite (ireg_of_eq _ _ EQ). destruct (Int.eq_dec n Int.zero). - econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split. destruct (rs x); simpl; auto. subst. rewrite Int.and_idem. + econstructor; split. apply exec_straight_one. simpl; eauto. apply nextinstr_compare_ints_pc. + split. destruct (rs#x); simpl; auto. subst. rewrite Int.and_idem. eapply testcond_for_signed_comparison_32_correct; eauto. intros. unfold compare_ints. Simplifs. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split. destruct (Val.cmp_bool c0 (rs x) (Vint n)) eqn:?; auto. + econstructor; split. apply exec_straight_one. simpl; eauto. apply nextinstr_compare_ints_pc. + split. destruct (Val.cmp_bool c0 (rs#x) (Vint n)) eqn:?; auto. eapply testcond_for_signed_comparison_32_correct; eauto. intros. unfold compare_ints. Simplifs. - (* compuimm *) simpl. rewrite (ireg_of_eq _ _ EQ). - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs x) (Vint n)) eqn:?; auto. + econstructor. split. apply exec_straight_one. simpl. eauto. apply nextinstr_compare_ints_pc. + split. destruct (Val.cmpu_bool (Mem.valid_pointer m) c0 (rs#x) (Vint n)) eqn:?; auto. eapply testcond_for_unsigned_comparison_32_correct; eauto. intros. unfold compare_ints. Simplifs. - (* compl *) simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1). - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmpl_bool c0 (rs x) (rs x0)) eqn:?; auto. + econstructor. split. apply exec_straight_one. simpl. eauto. apply nextinstr_compare_longs_pc. + split. destruct (Val.cmpl_bool c0 (rs#x) (rs#x0)) eqn:?; auto. eapply testcond_for_signed_comparison_64_correct; eauto. intros. unfold compare_longs. Simplifs. - (* complu *) simpl. rewrite (ireg_of_eq _ _ EQ). rewrite (ireg_of_eq _ _ EQ1). - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmplu_bool (Mem.valid_pointer m) c0 (rs x) (rs x0)) eqn:?; auto. + econstructor. split. apply exec_straight_one. simpl. eauto. apply nextinstr_compare_longs_pc. + split. destruct (Val.cmplu_bool (Mem.valid_pointer m) c0 (rs#x) (rs#x0)) eqn:?; auto. eapply testcond_for_unsigned_comparison_64_correct; eauto. intros. unfold compare_longs. Simplifs. - (* compimm *) simpl. rewrite (ireg_of_eq _ _ EQ). destruct (Int64.eq_dec n Int64.zero). - econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split. destruct (rs x); simpl; auto. subst. rewrite Int64.and_idem. + econstructor; split. apply exec_straight_one. simpl; eauto. apply nextinstr_compare_longs_pc. + split. destruct (rs#x); simpl; auto. subst. rewrite Int64.and_idem. eapply testcond_for_signed_comparison_64_correct; eauto. intros. unfold compare_longs. Simplifs. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split. destruct (Val.cmpl_bool c0 (rs x) (Vlong n)) eqn:?; auto. + econstructor; split. apply exec_straight_one. simpl; eauto. apply nextinstr_compare_longs_pc. + split. destruct (Val.cmpl_bool c0 (rs#x) (Vlong n)) eqn:?; auto. eapply testcond_for_signed_comparison_64_correct; eauto. intros. unfold compare_longs. Simplifs. - (* compuimm *) simpl. rewrite (ireg_of_eq _ _ EQ). - econstructor. split. apply exec_straight_one. simpl. eauto. auto. - split. destruct (Val.cmplu_bool (Mem.valid_pointer m) c0 (rs x) (Vlong n)) eqn:?; auto. + econstructor. split. apply exec_straight_one. simpl. eauto. apply nextinstr_compare_longs_pc. + split. destruct (Val.cmplu_bool (Mem.valid_pointer m) c0 (rs#x) (Vlong n)) eqn:?; auto. eapply testcond_for_unsigned_comparison_64_correct; eauto. intros. unfold compare_longs. Simplifs. - (* compf *) simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1). - exists (nextinstr (compare_floats (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)). + exists (nextinstr (compare_floats (swap_floats c0 (rs#x) (rs#x0)) (swap_floats c0 (rs#x0) (rs#x)) rs)). split. apply exec_straight_one. destruct c0; simpl; auto. unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats_inv; auto with asmgen. - split. destruct (rs x); destruct (rs x0); simpl; auto. + split. destruct (rs#x); destruct (rs#x0); simpl; auto. repeat rewrite swap_floats_commut. apply testcond_for_float_comparison_correct. intros. Simplifs. apply compare_floats_inv; auto with asmgen. - (* notcompf *) simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1). - exists (nextinstr (compare_floats (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)). + exists (nextinstr (compare_floats (swap_floats c0 (rs#x) (rs#x0)) (swap_floats c0 (rs#x0) (rs#x)) rs)). split. apply exec_straight_one. destruct c0; simpl; auto. unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats_inv; auto with asmgen. - split. destruct (rs x); destruct (rs x0); simpl; auto. + split. destruct (rs#x); destruct (rs#x0); simpl; auto. repeat rewrite swap_floats_commut. apply testcond_for_neg_float_comparison_correct. intros. Simplifs. apply compare_floats_inv; auto with asmgen. - (* compfs *) simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1). - exists (nextinstr (compare_floats32 (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)). + exists (nextinstr (compare_floats32 (swap_floats c0 (rs#x) (rs#x0)) (swap_floats c0 (rs#x0) (rs#x)) rs)). split. apply exec_straight_one. destruct c0; simpl; auto. unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats32_inv; auto with asmgen. - split. destruct (rs x); destruct (rs x0); simpl; auto. + split. destruct (rs#x); destruct (rs#x0); simpl; auto. repeat rewrite swap_floats_commut. apply testcond_for_float32_comparison_correct. intros. Simplifs. apply compare_floats32_inv; auto with asmgen. - (* notcompfs *) simpl. rewrite (freg_of_eq _ _ EQ). rewrite (freg_of_eq _ _ EQ1). - exists (nextinstr (compare_floats32 (swap_floats c0 (rs x) (rs x0)) (swap_floats c0 (rs x0) (rs x)) rs)). + exists (nextinstr (compare_floats32 (swap_floats c0 (rs#x) (rs#x0)) (swap_floats c0 (rs#x0) (rs#x)) rs)). split. apply exec_straight_one. destruct c0; simpl; auto. unfold nextinstr. rewrite Pregmap.gss. rewrite compare_floats32_inv; auto with asmgen. - split. destruct (rs x); destruct (rs x0); simpl; auto. + split. destruct (rs#x); destruct (rs#x0); simpl; auto. repeat rewrite swap_floats_commut. apply testcond_for_neg_float32_comparison_correct. intros. Simplifs. apply compare_floats32_inv; auto with asmgen. - (* maskzero *) simpl. rewrite (ireg_of_eq _ _ EQ). - econstructor. split. apply exec_straight_one. simpl; eauto. auto. - split. destruct (rs x); simpl; auto. + econstructor. split. apply exec_straight_one. simpl; eauto. apply nextinstr_compare_ints_pc. + split. destruct (rs#x); simpl; auto. generalize (compare_ints_spec rs (Vint (Int.and i n)) Vzero m). intros [A B]. rewrite A. unfold Val.cmpu; simpl. destruct (Int.eq (Int.and i n) Int.zero); auto. intros. unfold compare_ints. Simplifs. - (* masknotzero *) simpl. rewrite (ireg_of_eq _ _ EQ). - econstructor. split. apply exec_straight_one. simpl; eauto. auto. - split. destruct (rs x); simpl; auto. + econstructor. split. apply exec_straight_one. simpl; eauto. apply nextinstr_compare_ints_pc. + split. destruct (rs#x); simpl; auto. generalize (compare_ints_spec rs (Vint (Int.and i n)) Vzero m). intros [A B]. rewrite A. unfold Val.cmpu; simpl. destruct (Int.eq (Int.and i n) Int.zero); auto. intros. unfold compare_ints. Simplifs. @@ -1112,8 +1273,9 @@ Proof. intros. destruct cond; simpl in *. - (* base *) econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. Simplifs. intros; Simplifs. + apply exec_straight_one. simpl; eauto. auto with asmgen. + split. Simplifs; apply Val.load_result_of_optbool; simpl; destruct rd, Archi.ptr64; auto. + intros; Simplifs. - (* or *) assert (Val.of_optbool match eval_testcond c1 rs1 with @@ -1136,15 +1298,19 @@ Proof. simpl; eauto. simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto. simpl; eauto. - auto. auto. auto. + Simplifs. Simplifs. Simplifs. intuition Simplifs. + rewrite Val.load_result_or by (simpl; destruct Archi.ptr64; auto). + f_equal; Simplifs; apply Val.load_result_of_optbool; auto. econstructor; split. eapply exec_straight_three. simpl; eauto. simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto. simpl. eauto. - auto. auto. auto. - split. Simplifs. rewrite Val.or_commut. decEq; Simplifs. + Simplifs. Simplifs. Simplifs. + split. Simplifs. rewrite Val.or_commut. + rewrite Val.load_result_or by (simpl; destruct Archi.ptr64, rd; auto). + f_equal; Simplifs; apply Val.load_result_of_optbool; auto. intros. destruct H0; Simplifs. - (* and *) assert (Val.of_optbool @@ -1170,15 +1336,19 @@ Proof. simpl; eauto. simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto. simpl; eauto. - auto. auto. auto. + Simplifs. Simplifs. Simplifs. intuition Simplifs. + rewrite Val.load_result_and by (simpl; destruct Archi.ptr64; auto). + f_equal; Simplifs; apply Val.load_result_of_optbool; auto. econstructor; split. eapply exec_straight_three. simpl; eauto. simpl. rewrite eval_testcond_nextinstr. repeat rewrite eval_testcond_set_ireg. eauto. simpl. eauto. - auto. auto. auto. - split. Simplifs. rewrite Val.and_commut. decEq; Simplifs. + Simplifs. Simplifs. Simplifs. + split. Simplifs. rewrite Val.and_commut. + rewrite Val.load_result_and by (simpl; destruct Archi.ptr64, rd; auto). + f_equal; Simplifs; apply Val.load_result_of_optbool; auto. intros. destruct H0; Simplifs. Qed. @@ -1193,8 +1363,9 @@ Proof. - apply mk_setcc_base_correct. - exploit mk_setcc_base_correct. intros [rs2 [A [B C]]]. econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. - simpl. eauto. simpl. auto. - intuition Simplifs. + simpl. eauto. Simplifs. + intuition Simplifs; + rewrite B; apply Val.load_result_of_optbool; simpl; destruct Archi.ptr64, rd; auto. Qed. (** Translation of arithmetic operations. *) @@ -1211,19 +1382,41 @@ Ltac ArgsInv := | _ => idtac end. +Ltac LoadResultMatch := + match goal with + | [ |- Val.load_result _ (if ?c then _ else _) = _ ] => + destruct c; try destruct Archi.ptr64; try reflexivity; try LoadResultMatch + | [ |- Val.load_result (match ?c with _ => _ end) _ = _ ] => + destruct c; try destruct Archi.ptr64; reflexivity + end. + +Ltac InvertLoadResult := + try apply load_result_sub; + match goal with + | [ H: Archi.ptr64 = true |- Val.load_result (Preg.chunk_of (IR ?i)) _ = _ ] => + simpl; destruct i; try rewrite H; auto + | [ |- Val.load_result _ (?f ?x ?y) = ?f ?x ?y ] => + destruct x, y; simpl; try LoadResultMatch + | [ |- Val.load_result _ (?f ?x) = ?f ?x ] => + destruct x; simpl; try LoadResultMatch + | [ |- Val.load_result (Preg.chunk_of (IR ?i)) _ = _ ] => + simpl; destruct i; try destruct Archi.ptr64; auto + | _ => idtac + end. + Ltac TranslOp := econstructor; split; - [ apply exec_straight_one; [ simpl; eauto | auto ] - | split; [ Simplifs | intros; Simplifs ]]. + [ apply exec_straight_one; [ simpl; eauto | auto; Simplif ] + | split; [ Simplifs; InvertLoadResult | intros; Simplifs ]]. Lemma transl_op_correct: forall op args res k c (rs: regset) m v, transl_op op args res k = OK c -> - eval_operation ge (rs#RSP) op (map rs (map preg_of args)) m = Some v -> + eval_operation ge (rs#RSP) op (map (fun r => rs#r) (map preg_of args)) m = Some v -> exists rs', exec_straight ge fn c rs m k rs' m /\ Val.lessdef v rs'#(preg_of res) - /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r. + /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs'#r = rs#r. Proof. Transparent destroyed_by_op. intros until v; intros TR EV. @@ -1231,11 +1424,11 @@ Transparent destroyed_by_op. (exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of res) = v - /\ 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) -> 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). { intros [rs' [A [B C]]]. subst v. exists rs'; auto. } @@ -1252,118 +1445,199 @@ Transparent destroyed_by_op. apply SAME. destruct (Float.eq_dec n Float.zero). subst n. TranslOp. TranslOp. (* singleconst *) apply SAME. destruct (Float32.eq_dec n Float32.zero). subst n. TranslOp. TranslOp. +(* indirectsymbol (only in 32-bit mode) *) + match goal with + | [ |- exists _, _ /\ Val.lessdef (Genv.symbol_address _ _ _) _ /\ _ ] => + TranslOp; + rewrite nextinstr_nf_inv, Pregmap.gss by auto; + unfold Genv.symbol_address; destruct (Genv.find_symbol _ _); auto + | _ => idtac + end. (* cast8signed *) - apply SAME. eapply mk_intconv_correct; eauto. + apply SAME. eapply mk_intconv_Pmovsb_rr_correct; eauto. (* cast8unsigned *) - apply SAME. eapply mk_intconv_correct; eauto. + apply SAME. eapply mk_intconv_Pmovzb_rr_correct; eauto. (* mulhs *) - apply SAME. TranslOp. destruct H1. Simplifs. + apply SAME. TranslOp. f_equal; Simplif. destruct H1. Simplifs. (* mulhu *) - apply SAME. TranslOp. destruct H1. Simplifs. + apply SAME. TranslOp. f_equal; Simplif. destruct H1. Simplifs. (* div *) apply SAME. - exploit (divs_mods_exists (rs RAX) (rs RCX)). left; congruence. + exploit (divs_mods_exists (rs#RAX) (rs#RCX)). left; congruence. intros (nh & nl & d & q & r & A & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- (Vint nh))). econstructor; split. eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vint q = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; destruct Archi.ptr64; simpl. + rewrite D. reflexivity. rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vint q) with v by congruence; simpl. + destruct Archi.ptr64 eqn:PTR64. + Simplifs; simpl; rewrite PTR64; auto. + Simplifs; replace v with (Vint q) by congruence; simpl; rewrite PTR64; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* divu *) apply SAME. - exploit (divu_modu_exists (rs RAX) (rs RCX)). left; congruence. + exploit (divu_modu_exists (rs#RAX) (rs#RCX)). left; congruence. intros (n & d & q & r & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- Vzero)). econstructor; split. eapply exec_straight_two with (rs2 := rs1). reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vint q = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; destruct Archi.ptr64; simpl. + rewrite D. reflexivity. rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vint q) with v by congruence; simpl. + destruct Archi.ptr64 eqn:PTR64. + Simplifs; simpl; rewrite PTR64; auto. + Simplifs; replace v with (Vint q) by congruence; simpl; rewrite PTR64; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* mod *) apply SAME. - exploit (divs_mods_exists (rs RAX) (rs RCX)). right; congruence. + exploit (divs_mods_exists (rs#RAX) (rs#RCX)). right; congruence. intros (nh & nl & d & q & r & A & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- (Vint nh))). econstructor; split. eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vint r = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; destruct Archi.ptr64; simpl. + rewrite D. reflexivity. rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vint r) with v by congruence; simpl. + destruct Archi.ptr64 eqn:PTR64. + Simplifs; simpl; rewrite PTR64; auto. + Simplifs; replace v with (Vint r) by congruence; simpl; rewrite PTR64; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* modu *) apply SAME. - exploit (divu_modu_exists (rs RAX) (rs RCX)). right; congruence. + exploit (divu_modu_exists (rs#RAX) (rs#RCX)). right; congruence. intros (n & d & q & r & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- Vzero)). econstructor; split. eapply exec_straight_two with (rs2 := rs1). reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vint r = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; destruct Archi.ptr64; simpl. + rewrite D. reflexivity. rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vint r) with v by congruence; simpl. + destruct Archi.ptr64 eqn:PTR64. + Simplifs; simpl; rewrite PTR64; auto. + Simplifs; replace v with (Vint r) by congruence; simpl; rewrite PTR64; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* shrximm *) apply SAME. eapply mk_shrximm_correct; eauto. (* lea *) exploit transl_addressing_mode_32_correct; eauto. intros EA. - TranslOp. rewrite nextinstr_inv; auto with asmgen. rewrite Pregmap.gss. rewrite normalize_addrmode_32_correct; auto. + TranslOp. rewrite nextinstr_inv; auto with asmgen. + rewrite Pregmap.gss, normalize_addrmode_32_correct. + rewrite Preg.chunk_of_reg_type, Val.load_result_same; auto. + apply Val.has_subtype with (ty1 := Tany32). + destruct (Preg.type_cases x0) as [H|H]; rewrite H; auto. + apply eval_addrmode32_type. (* mullhs *) - apply SAME. TranslOp. destruct H1. Simplifs. + apply SAME. TranslOp. f_equal; Simplifs. try (destruct H1; Simplifs). (* mullhu *) - apply SAME. TranslOp. destruct H1. Simplifs. + apply SAME. TranslOp. f_equal; Simplifs. try (destruct H1; Simplifs). (* divl *) apply SAME. - exploit (divls_modls_exists (rs RAX) (rs RCX)). left; congruence. + exploit (divls_modls_exists (rs#RAX) (rs#RCX)). left; congruence. intros (nh & nl & d & q & r & A & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- (Vlong nh))). econstructor; split. eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vlong q = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; rewrite Heqb; simpl. + rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vlong q) with v by congruence. simpl. Simplifs; simpl; rewrite Heqb; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* divlu *) apply SAME. - exploit (divlu_modlu_exists (rs RAX) (rs RCX)). left; congruence. + exploit (divlu_modlu_exists (rs#RAX) (rs#RCX)). left; congruence. intros (n & d & q & r & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- (Vlong Int64.zero))). econstructor; split. eapply exec_straight_two with (rs2 := rs1). reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vlong q = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; rewrite Heqb; simpl. + rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vlong q) with v by congruence. simpl. Simplifs; simpl; rewrite Heqb; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* modl *) apply SAME. - exploit (divls_modls_exists (rs RAX) (rs RCX)). right; congruence. + exploit (divls_modls_exists (rs#RAX) (rs#RCX)). right; congruence. intros (nh & nl & d & q & r & A & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- (Vlong nh))). econstructor; split. eapply exec_straight_two with (rs2 := rs1). simpl. rewrite A. reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vlong r = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; rewrite Heqb; simpl. + rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vlong r) with v by congruence. simpl. Simplifs; simpl; rewrite Heqb; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* modlu *) apply SAME. - exploit (divlu_modlu_exists (rs RAX) (rs RCX)). right; congruence. + exploit (divlu_modlu_exists (rs#RAX) (rs#RCX)). right; congruence. intros (n & d & q & r & B & C & D & E & F). set (rs1 := nextinstr_nf (rs#RDX <- (Vlong Int64.zero))). econstructor; split. eapply exec_straight_two with (rs2 := rs1). reflexivity. - simpl. change (rs1 RAX) with (rs RAX); rewrite B. - change (rs1 RCX) with (rs RCX); rewrite C. - rewrite D. reflexivity. auto. auto. - split. change (Vlong r = v). congruence. + simpl. replace (rs1#RAX) with (rs#RAX); rewrite B. + replace (rs1#RCX) with (rs#RCX); rewrite C. + unfold rs1. rewrite nextinstr_nf_inv, Pregmap.gss. simpl; rewrite Heqb; simpl. + rewrite D. reflexivity. auto. + unfold rs1. symmetry. Simplifs. + unfold rs1. symmetry. Simplifs. + unfold rs1. auto with asmgen. + Simplifs. + unfold rs1. f_equal. Simplifs. + split. + replace (Vlong r) with v by congruence. simpl. Simplifs; simpl; rewrite Heqb; auto. simpl; intros. destruct H2. unfold rs1; Simplifs. (* shrxlimm *) apply SAME. eapply mk_shrxlimm_correct; eauto. @@ -1371,25 +1645,28 @@ Transparent destroyed_by_op. exploit transl_addressing_mode_64_correct; eauto. intros EA. generalize (normalize_addrmode_64_correct x rs). destruct (normalize_addrmode_64 x) as [am' [delta|]]; intros EV. econstructor; split. eapply exec_straight_two. - simpl. reflexivity. simpl. reflexivity. auto. auto. + simpl. reflexivity. simpl. reflexivity. Simplifs. Simplifs. split. rewrite nextinstr_nf_inv by auto. rewrite Pregmap.gss. rewrite nextinstr_inv by auto with asmgen. - rewrite Pregmap.gss. rewrite <- EV; auto. + rewrite Pregmap.gss. replace (Preg.chunk_of x0) with Many64; simpl. rewrite <- EV; auto. + rewrite Heqb; destruct x0; auto. intros; Simplifs. - TranslOp. rewrite nextinstr_inv; auto with asmgen. rewrite Pregmap.gss; auto. rewrite <- EV; auto. + TranslOp. rewrite nextinstr_inv; auto with asmgen. + rewrite Pregmap.gss. replace (Preg.chunk_of x0) with Many64; simpl. rewrite <- EV; auto. + rewrite Heqb; destruct x0; auto. (* intoffloat *) - apply SAME. TranslOp. rewrite H0; auto. + apply SAME. TranslOp; auto using Val.load_result_intoffloat; rewrite H0; auto. (* floatofint *) apply SAME. TranslOp. rewrite H0; auto. (* intofsingle *) - apply SAME. TranslOp. rewrite H0; auto. + apply SAME. TranslOp; auto using Val.load_result_intofsingle; rewrite H0; auto. (* singleofint *) apply SAME. TranslOp. rewrite H0; auto. (* longoffloat *) - apply SAME. TranslOp. rewrite H0; auto. + apply SAME. TranslOp; rewrite H0; auto. (* floatoflong *) apply SAME. TranslOp. rewrite H0; auto. (* longofsingle *) - apply SAME. TranslOp. rewrite H0; auto. + apply SAME. TranslOp; rewrite H0; auto. (* singleoflong *) apply SAME. TranslOp. rewrite H0; auto. (* condition *) @@ -1397,10 +1674,10 @@ Transparent destroyed_by_op. exploit mk_setcc_correct; eauto. intros [rs3 [S [T U]]]. exists rs3. split. eapply exec_straight_trans. eexact P. eexact S. - split. rewrite T. destruct (eval_condition cond rs ## (preg_of ## args) m). + split. rewrite T. destruct (eval_condition cond (map (fun r => rs#r) (map preg_of args)) m). rewrite Q. auto. simpl; auto. - intros. transitivity (rs2 r); auto. + intros. transitivity (rs2#r); auto. Qed. (** Translation of memory loads. *) @@ -1408,7 +1685,7 @@ Qed. Lemma transl_load_correct: forall chunk addr args dest k c (rs: regset) m a v, transl_load chunk addr args dest k = OK c -> - eval_addressing ge (rs#RSP) addr (map rs (map preg_of args)) = Some a -> + eval_addressing ge (rs#RSP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> exists rs', exec_straight ge fn c rs m k rs' m @@ -1421,20 +1698,25 @@ Proof. set (rs2 := nextinstr_nf (rs#(preg_of dest) <- v)). assert (exec_load ge chunk m x rs (preg_of dest) = Next rs2 m). unfold exec_load. rewrite EA'. rewrite H1. auto. - assert (rs2 PC = Val.offset_ptr (rs PC) Ptrofs.one). - transitivity (Val.offset_ptr ((rs#(preg_of dest) <- v) PC) Ptrofs.one). - auto. decEq. apply Pregmap.gso; auto with asmgen. + assert (rs2#PC = Val.offset_ptr (rs#PC) Ptrofs.one). + transitivity (Val.offset_ptr ((rs#(preg_of dest) <- v)#PC) Ptrofs.one). + unfold rs2. Simplifs. rewrite Pregmap.gso; auto with asmgen. + decEq. apply Pregmap.gso; auto with asmgen. exists rs2. split. destruct chunk; ArgsInv; apply exec_straight_one; auto. - split. unfold rs2. rewrite nextinstr_nf_inv1. Simplifs. apply preg_of_data. + split. unfold rs2. rewrite nextinstr_nf_inv1. Simplifs. + replace (Preg.chunk_of (preg_of dest)) with (chunk_of_type (Mreg.type dest)). + apply Val.load_result_same. eapply Val.has_subtype; eauto using Mem.loadv_type. + destruct dest; auto. + apply preg_of_data. intros. unfold rs2. Simplifs. Qed. Lemma transl_store_correct: forall chunk addr args src k c (rs: regset) m a m', transl_store chunk addr args src k = OK c -> - eval_addressing ge (rs#RSP) addr (map rs (map preg_of args)) = Some a -> - Mem.storev chunk m a (rs (preg_of src)) = Some m' -> + eval_addressing ge (rs#RSP) addr (map (fun r => rs#r) (map preg_of args)) = Some a -> + Mem.storev chunk m a (rs#(preg_of src)) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' /\ forall r, data_preg r = true -> preg_notin r (destroyed_by_store chunk addr) -> rs'#r = rs#r. @@ -1451,31 +1733,31 @@ Proof. (* int16signed *) econstructor; split. apply exec_straight_one. simpl. unfold exec_store. - replace (Mem.storev Mint16unsigned m (eval_addrmode ge x rs) (rs x0)) - with (Mem.storev Mint16signed m (eval_addrmode ge x rs) (rs x0)). + replace (Mem.storev Mint16unsigned m (eval_addrmode ge x rs) (rs#x0)) + with (Mem.storev Mint16signed m (eval_addrmode ge x rs) (rs#x0)). rewrite H1. eauto. destruct (eval_addrmode ge x rs); simpl; auto. rewrite Mem.store_signed_unsigned_16; auto. - auto. + apply nextinstr_nf_pc. intros. Simplifs. (* int16unsigned *) econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto with asmgen. intros. Simplifs. (* int32 *) econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto with asmgen. intros. Simplifs. (* int64 *) econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto with asmgen. intros. Simplifs. (* float32 *) econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto with asmgen. intros. Transparent destroyed_by_store. simpl in H2. simpl. Simplifs. (* float64 *) econstructor; split. - apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto. + apply exec_straight_one. simpl. unfold exec_store. rewrite H1. eauto. auto with asmgen. intros. Simplifs. Qed. diff --git a/x86/Conventions1.v b/x86/Conventions1.v index 646c4afb31..49c61b8b54 100644 --- a/x86/Conventions1.v +++ b/x86/Conventions1.v @@ -14,7 +14,7 @@ machine registers and stack slots. *) Require Import Coqlib Decidableplus. -Require Import AST Machregs Locations. +Require Import AST Values Machregs Locations. (** * Classification of machine registers *) @@ -130,6 +130,17 @@ Proof. destruct Archi.ptr64; destruct (sig_res sig) as [[]|]; auto. Qed. +Lemma loc_result_has_type: + forall res sig, + Val.has_type res (proj_sig_res sig) -> + Val.has_type_rpair res (loc_result sig) Val.loword Val.hiword mreg_type. +Proof. + intros. unfold Val.has_type_rpair, loc_result, proj_sig_res in *. + destruct sig; simpl. destruct sig_res; simpl in H. + destruct t, res; simpl; auto; destruct Archi.ptr64 eqn:P; simpl; try rewrite P; auto. + destruct res; simpl; auto; destruct Archi.ptr64 eqn:P; simpl; try rewrite P; auto. +Qed. + (** The result locations are caller-save registers *) Lemma loc_result_caller_save: @@ -177,8 +188,8 @@ Fixpoint loc_arguments_32 | nil => nil | ty :: tys => match ty with - | Tlong => Twolong (S Outgoing (ofs + 1) Tint) (S Outgoing ofs Tint) - | _ => One (S Outgoing ofs ty) + | Tlong => Twolong (S Outgoing (ofs + 1) Q32) (S Outgoing ofs Q32) + | _ => One (S Outgoing ofs (quantity_of_typ ty)) end :: loc_arguments_32 tys (ofs + typesize ty) end. @@ -201,14 +212,14 @@ Fixpoint loc_arguments_64 | (Tint | Tlong | Tany32 | Tany64) as ty :: tys => match list_nth_z int_param_regs ir with | None => - One (S Outgoing ofs ty) :: loc_arguments_64 tys ir fr (ofs + 2) + One (S Outgoing ofs (quantity_of_typ ty)) :: loc_arguments_64 tys ir fr (ofs + 2) | Some ireg => One (R ireg) :: loc_arguments_64 tys (ir + 1) fr ofs end | (Tfloat | Tsingle) as ty :: tys => match list_nth_z float_param_regs fr with | None => - One (S Outgoing ofs ty) :: loc_arguments_64 tys ir fr (ofs + 2) + One (S Outgoing ofs (quantity_of_typ ty)) :: loc_arguments_64 tys ir fr (ofs + 2) | Some freg => One (R freg) :: loc_arguments_64 tys ir (fr + 1) ofs end @@ -258,20 +269,20 @@ Definition size_arguments (s: signature) : Z := Definition loc_argument_acceptable (l: loc) : Prop := match l with | R r => is_callee_save r = false - | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs) + | S Outgoing ofs q => ofs >= 0 /\ (quantity_align q | ofs) | _ => False end. Definition loc_argument_32_charact (ofs: Z) (l: loc) : Prop := match l with - | S Outgoing ofs' ty => ofs' >= ofs /\ typealign ty = 1 + | S Outgoing ofs' q => ofs' >= ofs /\ quantity_align q = 1 | _ => False end. Definition loc_argument_64_charact (ofs: Z) (l: loc) : Prop := match l with | R r => In r int_param_regs \/ In r float_param_regs - | S Outgoing ofs' ty => ofs' >= ofs /\ (2 | ofs') + | S Outgoing ofs' q => ofs' >= ofs /\ (2 | ofs') | _ => False end. @@ -284,7 +295,7 @@ Proof. induction tyl as [ | ty tyl]; simpl loc_arguments_32; intros. - contradiction. - destruct H. -+ destruct ty; subst p; simpl; omega. ++ destruct ty; subst p; simpl; unfold quantity_align; omega. + apply IHtyl in H. generalize (typesize_pos ty); intros. destruct p; simpl in *. * eapply X; eauto; omega. * destruct H; split; eapply X; eauto; omega. @@ -303,10 +314,10 @@ Proof. Opaque list_nth_z. induction tyl; simpl loc_arguments_64; intros. elim H. - assert (A: forall ty, In p + assert (A: forall q, In p match list_nth_z int_param_regs ir with | Some ireg => One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs - | None => One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2) + | None => One (S Outgoing ofs q) :: loc_arguments_64 tyl ir fr (ofs + 2) end -> forall_rpair (loc_argument_64_charact ofs) p). { intros. destruct (list_nth_z int_param_regs ir) as [r|] eqn:E; destruct H1. @@ -314,10 +325,10 @@ Opaque list_nth_z. eapply IHtyl; eauto. subst. split. omega. assumption. eapply Y; eauto. omega. } - assert (B: forall ty, In p + assert (B: forall q, In p match list_nth_z float_param_regs fr with | Some ireg => One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs - | None => One (S Outgoing ofs ty) :: loc_arguments_64 tyl ir fr (ofs + 2) + | None => One (S Outgoing ofs q) :: loc_arguments_64 tyl ir fr (ofs + 2) end -> forall_rpair (loc_argument_64_charact ofs) p). { intros. destruct (list_nth_z float_param_regs fr) as [r|] eqn:E; destruct H1. @@ -338,15 +349,15 @@ Proof. assert (B: forall r, In r float_param_regs -> is_callee_save r = false) by decide_goal. assert (X: forall l, loc_argument_64_charact 0 l -> loc_argument_acceptable l). { unfold loc_argument_64_charact, loc_argument_acceptable. - destruct l as [r | [] ofs ty]; auto. intros [C|C]; auto. + destruct l as [r | [] ofs q]; auto. intros [C|C]; auto. intros [C D]. split; auto. apply Z.divide_trans with 2; auto. - exists (2 / typealign ty); destruct ty; reflexivity. + exists (2 / quantity_align q); destruct q; reflexivity. } exploit loc_arguments_64_charact; eauto using Z.divide_0_r. unfold forall_rpair; destruct p; intuition auto. - (* 32 bits *) assert (X: forall l, loc_argument_32_charact 0 l -> loc_argument_acceptable l). - { destruct l as [r | [] ofs ty]; simpl; intuition auto. rewrite H2; apply Z.divide_1_l. } + { destruct l as [r | [] ofs q]; simpl; intuition auto. rewrite H2; apply Z.divide_1_l. } exploit loc_arguments_32_charact; eauto. unfold forall_rpair; destruct p; intuition auto. Qed. @@ -395,9 +406,9 @@ Proof. 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. + forall ofs q tyl ofs0, + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_32 tyl ofs0)) -> + ofs + typesize (typ_of_quantity q) <= size_arguments_32 tyl ofs0. Proof. induction tyl as [ | t l]; simpl; intros x IN. - contradiction. @@ -414,22 +425,22 @@ Proof. 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. + forall ofs q tyl ir fr ofs0, + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments_64 tyl ir fr ofs0)) -> + ofs + typesize (typ_of_quantity q) <= 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 + In (S Outgoing ofs q) (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 <= + ofs + typesize (typ_of_quantity q) <= 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) @@ -437,16 +448,16 @@ Proof. { 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. + - inv H0. apply Z.le_trans with (ofs + 2). specialize (T (typ_of_quantity q)). omega. apply size_arguments_64_above. - eapply IHtyl; eauto. } assert (B: forall ty0, - In (S Outgoing ofs ty) (regs_of_rpairs + In (S Outgoing ofs q) (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 <= + ofs + typesize (typ_of_quantity q) <= 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) @@ -454,15 +465,15 @@ Proof. { 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. + - inv H0. apply Z.le_trans with (ofs + 2). specialize (T (typ_of_quantity q)). 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. + forall (s: signature) (ofs: Z) (q: quantity), + In (S Outgoing ofs q) (regs_of_rpairs (loc_arguments s)) -> + ofs + typesize (typ_of_quantity q) <= size_arguments s. Proof. unfold loc_arguments, size_arguments; intros. destruct Archi.ptr64; eauto using loc_arguments_32_bounded, loc_arguments_64_bounded. diff --git a/x86/Machregs.v b/x86/Machregs.v index bdf492ed4c..82a4da28ee 100644 --- a/x86/Machregs.v +++ b/x86/Machregs.v @@ -17,6 +17,7 @@ Require Import Maps. Require Import AST. Require Import Integers. Require Import Op. +Require Import Memdata. (** ** Machine registers *) @@ -67,12 +68,21 @@ Instance Finite_mreg : Finite mreg := { Definition mreg_type (r: mreg): typ := match r with | AX | BX | CX | DX | SI | DI | BP => if Archi.ptr64 then Tany64 else Tany32 - | R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15 => Tany64 + (* These registers don't exist in 32-bit mode. When they do exist, they + are 64 bits wide. For some proofs it's nonetheless easier to pretend + that *if* they exist, they have the same type as the other integer + registers. *) + | R8 | R9 | R10 | R11 | R12 | R13 | R14 | R15 => if Archi.ptr64 then Tany64 else Tany32 | X0 | X1 | X2 | X3 | X4 | X5 | X6 | X7 => Tany64 | X8 | X9 | X10 | X11 | X12 | X13 | X14 | X15 => Tany64 | FP0 => Tany64 end. +Lemma mreg_type_cases: forall r, mreg_type r = Tany32 \/ mreg_type r = Tany64. +Proof. + destruct r; simpl; auto. +Qed. + Local Open Scope positive_scope. Module IndexedMreg <: INDEXED_TYPE. @@ -80,17 +90,36 @@ Module IndexedMreg <: INDEXED_TYPE. Definition eq := mreg_eq. Definition index (r: mreg): positive := match r with - | AX => 1 | BX => 2 | CX => 3 | DX => 4 | SI => 5 | DI => 6 | BP => 7 - | R8 => 8 | R9 => 9 | R10 => 10 | R11 => 11 | R12 => 12 | R13 => 13 | R14 => 14 | R15 => 15 - | X0 => 16 | X1 => 17 | X2 => 18 | X3 => 19 | X4 => 20 | X5 => 21 | X6 => 22 | X7 => 23 - | X8 => 24 | X9 => 25 | X10 => 26 | X11 => 27 | X12 => 28 | X13 => 29 | X14 => 30 | X15 => 31 - | FP0 => 32 + | AX => 2 | BX => 4 | CX => 6 | DX => 8 | SI => 10 | DI => 12 | BP => 14 + | R8 => 16 | R9 => 18 | R10 => 20 | R11 => 22 | R12 => 24 | R13 => 26 | R14 => 28 | R15 => 30 + | X0 => 32 | X1 => 34 | X2 => 36 | X3 => 38 | X4 => 40 | X5 => 42 | X6 => 44 | X7 => 46 + | X8 => 48 | X9 => 50 | X10 => 52 | X11 => 54 | X12 => 56 | X13 => 58 | X14 => 60 | X15 => 62 + | FP0 => 64 end. Lemma index_inj: forall r1 r2, index r1 = index r2 -> r1 = r2. Proof. decide_goal. Qed. + + Open Scope Z_scope. + + Lemma scaled_index_with_size_aux: + forall r1 r2, Zpos (index r1) < Zpos (index r2) -> Zpos (index r1) + 2 <= Zpos (index r2). + Proof. + decide_goal. + Qed. + + Lemma scaled_index_with_size: + forall r1 r2, + Zpos (index r1) < Zpos (index r2) -> + Zpos (index r1) * 4 + AST.typesize (mreg_type r1) <= Zpos (index r2) * 4. + Proof. + intros. + generalize (scaled_index_with_size_aux r1 r2); intro. + assert (AST.typesize (mreg_type r1) <= 8) by (destruct (mreg_type r1); simpl; omega). + omega. + Qed. End IndexedMreg. Definition is_stack_reg (r: mreg) : bool := @@ -189,11 +218,7 @@ Definition destroyed_at_function_entry: list mreg := (* must include [destroyed_by_setstack ty] *) AX :: FP0 :: nil. -Definition destroyed_by_setstack (ty: typ): list mreg := - match ty with - | Tfloat | Tsingle => FP0 :: nil - | _ => nil - end. +Definition destroyed_by_setstack (q: quantity): list mreg := FP0 :: nil. Definition destroyed_at_indirect_call: list mreg := AX :: nil. diff --git a/x86/Op.v b/x86/Op.v index 79c84ca204..5e3290a306 100644 --- a/x86/Op.v +++ b/x86/Op.v @@ -762,6 +762,17 @@ Proof. intros; discriminate. Qed. +Lemma is_not_move_operation: + forall (F V A: Type) (genv: Genv.t F V) (sp: val) + (op: operation) (f: A -> val) (args: list A) (m: mem) (v: val), + eval_operation genv sp op (map f args) m = Some v -> + is_move_operation op args = None -> + op <> Omove. +Proof. + intros. destruct (eq_operation op Omove); auto. + subst. destruct args; try destruct args; simpl in *; congruence. +Qed. + (** [negate_condition cond] returns a condition that is logically equivalent to the negation of [cond]. *) diff --git a/x86/Stacklayout.v b/x86/Stacklayout.v index d375febfa1..73164d2154 100644 --- a/x86/Stacklayout.v +++ b/x86/Stacklayout.v @@ -53,8 +53,8 @@ Lemma frame_env_separated: m |= range sp 0 (fe_stack_data fe) ** range sp (fe_stack_data fe + bound_stack_data b) (fe_size fe) ** P -> m |= range sp (fe_ofs_local fe) (fe_ofs_local fe + 4 * bound_local b) ** range sp fe_ofs_arg (fe_ofs_arg + 4 * bound_outgoing b) - ** range sp (fe_ofs_link fe) (fe_ofs_link fe + size_chunk Mptr) - ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + size_chunk Mptr) + ** range sp (fe_ofs_link fe) (fe_ofs_link fe + size_chunk Mptr_any) + ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + size_chunk Mptr_any) ** range sp (fe_ofs_callee_save fe) (size_callee_save_area b (fe_ofs_callee_save fe)) ** P. Proof. @@ -66,7 +66,7 @@ Local Opaque Z.add Z.mul sepconj range. set (ol := align (size_callee_save_area b ocs) 8). set (ostkdata := align (ol + 4 * b.(bound_local)) 8). set (oretaddr := align (ostkdata + b.(bound_stack_data)) w). - replace (size_chunk Mptr) with w by (rewrite size_chunk_Mptr; auto). + replace (size_chunk Mptr_any) with w by (rewrite size_chunk_Mptr_any; auto). assert (0 < w) by (unfold w; destruct Archi.ptr64; omega). generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros. assert (0 <= 4 * b.(bound_outgoing)) by omega. @@ -128,8 +128,8 @@ Lemma frame_env_aligned: (8 | fe_ofs_arg) /\ (8 | fe_ofs_local fe) /\ (8 | fe_stack_data fe) - /\ (align_chunk Mptr | fe_ofs_link fe) - /\ (align_chunk Mptr | fe_ofs_retaddr fe). + /\ (align_chunk Mptr_any | fe_ofs_link fe) + /\ (align_chunk Mptr_any | fe_ofs_retaddr fe). Proof. intros; simpl. set (w := if Archi.ptr64 then 8 else 4). @@ -139,10 +139,19 @@ Proof. set (ostkdata := align (ol + 4 * b.(bound_local)) 8). set (oretaddr := align (ostkdata + b.(bound_stack_data)) w). assert (0 < w) by (unfold w; destruct Archi.ptr64; omega). - replace (align_chunk Mptr) with w by (rewrite align_chunk_Mptr; auto). + rewrite align_chunk_Mptr_any. split. apply Z.divide_0_r. split. apply align_divides; omega. split. apply align_divides; omega. - split. apply align_divides; omega. + split. + unfold olink, w. destruct Archi.ptr64. + apply Z.divide_trans with (m := 8). + exists 2; auto. + apply align_divides; omega. + apply align_divides; omega. + unfold oretaddr, w. destruct Archi.ptr64. + apply Z.divide_trans with (m := 8). + exists 2; auto. + apply align_divides; omega. apply align_divides; omega. Qed. diff --git a/x86/TargetPrinter.ml b/x86/TargetPrinter.ml index 1bb8c22626..8c33bb7258 100644 --- a/x86/TargetPrinter.ml +++ b/x86/TargetPrinter.ml @@ -418,22 +418,14 @@ module Target(System: SYSTEM):TARGET = fprintf oc " movabsq $%Ld, %a\n" n1 ireg64 rd | Pmov_rs(rd, id) -> print_mov_rs oc rd id - | Pmovl_rm(rd, a) -> + | Pmovl_rm(rd, a) | Pmovl_rm_a(rd, a) -> fprintf oc " movl %a, %a\n" addressing a ireg32 rd - | Pmovq_rm(rd, a) -> + | Pmovq_rm(rd, a) | Pmovq_rm_a(rd, a) -> fprintf oc " movq %a, %a\n" addressing a ireg64 rd - | Pmov_rm_a(rd, a) -> - if Archi.ptr64 - then fprintf oc " movq %a, %a\n" addressing a ireg64 rd - else fprintf oc " movl %a, %a\n" addressing a ireg32 rd - | Pmovl_mr(a, r1) -> + | Pmovl_mr(a, r1) | Pmovl_mr_a(a, r1) -> fprintf oc " movl %a, %a\n" ireg32 r1 addressing a - | Pmovq_mr(a, r1) -> + | Pmovq_mr(a, r1) | Pmovq_mr_a(a, r1) -> fprintf oc " movq %a, %a\n" ireg64 r1 addressing a - | Pmov_mr_a(a, r1) -> - if Archi.ptr64 - then fprintf oc " movq %a, %a\n" ireg64 r1 addressing a - else fprintf oc " movl %a, %a\n" ireg32 r1 addressing a | Pmovsd_ff(rd, r1) -> fprintf oc " movapd %a, %a\n" freg r1 freg rd | Pmovsd_fi(rd, n) -> @@ -452,17 +444,17 @@ module Target(System: SYSTEM):TARGET = fprintf oc " movss %a%s, %a %s %.18g\n" label lbl rip_rel freg rd comment (camlfloat_of_coqfloat32 n) - | Pmovss_fm(rd, a) -> + | Pmovss_fm(rd, a) | Pmovss_fm_a(rd, a) -> fprintf oc " movss %a, %a\n" addressing a freg rd - | Pmovss_mf(a, r1) -> + | Pmovss_mf(a, r1) | Pmovss_mf_a(a, r1) -> fprintf oc " movss %a, %a\n" freg r1 addressing a - | Pfldl_m(a) -> + | Pfldl_m(a) | Pfldl_m_a(a) -> fprintf oc " fldl %a\n" addressing a - | Pfstpl_m(a) -> + | Pfstpl_m(a) | Pfstpl_m_a(a) -> fprintf oc " fstpl %a\n" addressing a - | Pflds_m(a) -> + | Pflds_m(a) | Pflds_m_a(a) -> fprintf oc " flds %a\n" addressing a - | Pfstps_m(a) -> + | Pfstps_m(a) | Pfstps_m_a(a) -> fprintf oc " fstps %a\n" addressing a (* Moves with conversion *) | Pmovb_mr(a, r1) ->