From 7698300cfe2d3f944ce2e1d4a60a263620487718 Mon Sep 17 00:00:00 2001 From: xleroy Date: Fri, 20 Dec 2013 13:05:53 +0000 Subject: Merge of branch value-analysis. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2381 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/Allocation.v | 2 +- backend/Allocproof.v | 2 +- backend/CSE.v | 428 +++--- backend/CSEdomain.v | 147 ++ backend/CSEproof.v | 1608 ++++++++++---------- backend/Constprop.v | 447 ++---- backend/Constpropproof.v | 782 +++------- backend/Deadcode.v | 192 +++ backend/Deadcodeproof.v | 1014 +++++++++++++ backend/Kildall.v | 1167 +++++++++------ backend/Linearize.v | 2 +- backend/Linearizeproof.v | 4 +- backend/Liveness.v | 3 +- backend/NeedDomain.v | 1515 +++++++++++++++++++ backend/PrintRTL.ml | 3 + backend/Regalloc.ml | 2 +- backend/Splitting.ml | 2 +- backend/ValueAnalysis.v | 1812 +++++++++++++++++++++++ backend/ValueDomain.v | 3692 ++++++++++++++++++++++++++++++++++++++++++++++ 19 files changed, 10535 insertions(+), 2289 deletions(-) create mode 100644 backend/CSEdomain.v create mode 100644 backend/Deadcode.v create mode 100644 backend/Deadcodeproof.v create mode 100644 backend/NeedDomain.v create mode 100644 backend/ValueAnalysis.v create mode 100644 backend/ValueDomain.v (limited to 'backend') diff --git a/backend/Allocation.v b/backend/Allocation.v index e53c5aa..0851b77 100644 --- a/backend/Allocation.v +++ b/backend/Allocation.v @@ -1104,7 +1104,7 @@ Definition successors_block_shape (bsh: block_shape) : list node := end. Definition analyze (f: RTL.function) (env: regenv) (bsh: PTree.t block_shape) := - DS.fixpoint bsh successors_block_shape (transfer f env bsh) nil. + DS.fixpoint_allnodes bsh successors_block_shape (transfer f env bsh). (** * Validating and translating functions and programs *) diff --git a/backend/Allocproof.v b/backend/Allocproof.v index e91be74..1e637f9 100644 --- a/backend/Allocproof.v +++ b/backend/Allocproof.v @@ -1342,7 +1342,7 @@ Lemma analyze_successors: an!!pc = OK e -> exists e', transfer f env bsh s an!!s = OK e' /\ EqSet.Subset e' e. Proof. - unfold analyze; intros. exploit DS.fixpoint_solution; eauto. + unfold analyze; intros. exploit DS.fixpoint_allnodes_solution; eauto. rewrite H2. unfold DS.L.ge. destruct (transfer f env bsh s an#s); intros. exists e0; auto. contradiction. diff --git a/backend/CSE.v b/backend/CSE.v index 205d446..373acce 100644 --- a/backend/CSE.v +++ b/backend/CSE.v @@ -24,12 +24,12 @@ Require Import Memory. Require Import Op. Require Import Registers. Require Import RTL. -Require Import RTLtyping. +Require Import ValueDomain. +Require Import ValueAnalysis. +Require Import CSEdomain. Require Import Kildall. Require Import CombineOp. -(** * Value numbering *) - (** The idea behind value numbering algorithms is to associate abstract identifiers (``value numbers'') to the contents of registers at various program points, and record equations between these @@ -45,49 +45,10 @@ Require Import CombineOp. [r1 = add(r2, r3)] by a move instruction [r1 = r4], therefore eliminating a common subexpression and reusing the result of an earlier addition. - Abstract identifiers / value numbers are represented by positive integers. - Equations are of the form [valnum = rhs], where the right-hand sides - [rhs] are either arithmetic operations or memory loads. *) - -(* -Definition valnum := positive. + The representation of value numbers and equations is described in + module [CSEdomain]. *) -Inductive rhs : Type := - | Op: operation -> list valnum -> rhs - | Load: memory_chunk -> addressing -> list valnum -> rhs. -*) - -Definition eq_valnum: forall (x y: valnum), {x=y}+{x<>y} := peq. - -Definition eq_list_valnum (x y: list valnum) : {x=y}+{x<>y}. -Proof. decide equality. apply eq_valnum. Defined. - -Definition eq_rhs (x y: rhs) : {x=y}+{x<>y}. -Proof. - generalize Int.eq_dec; intro. - generalize Float.eq_dec; intro. - generalize eq_operation; intro. - generalize eq_addressing; intro. - assert (forall (x y: memory_chunk), {x=y}+{x<>y}). decide equality. - generalize eq_valnum; intro. - generalize eq_list_valnum; intro. - decide equality. -Defined. - -(** A value numbering is a collection of equations between value numbers - plus a partial map from registers to value numbers. Additionally, - we maintain the next unused value number, so as to easily generate - fresh value numbers. *) - -Record numbering : Type := mknumbering { - num_next: valnum; (**r first unused value number *) - num_eqs: list (valnum * rhs); (**r valid equations *) - num_reg: PTree.t valnum; (**r mapping register to valnum *) - num_val: PMap.t (list reg) (**r reverse mapping valnum to regs containing it *) -}. - -Definition empty_numbering := - mknumbering 1%positive nil (PTree.empty valnum) (PMap.init nil). +(** * Operations on value numberings *) (** [valnum_reg n r] returns the value number for the contents of register [r]. If none exists, a fresh value number is returned @@ -100,10 +61,10 @@ Definition valnum_reg (n: numbering) (r: reg) : numbering * valnum := | Some v => (n, v) | None => let v := n.(num_next) in - (mknumbering (Psucc v) - n.(num_eqs) - (PTree.set r v n.(num_reg)) - (PMap.set v (r :: nil) n.(num_val)), + ( {| num_next := Psucc v; + num_eqs := n.(num_eqs); + num_reg := PTree.set r v n.(num_reg); + num_val := PMap.set v (r :: nil) n.(num_val) |}, v) end. @@ -122,24 +83,67 @@ Fixpoint valnum_regs (n: numbering) (rl: list reg) for an equation of the form [vn = rhs] for some value number [vn]. If found, [Some vn] is returned, otherwise [None] is returned. *) -Fixpoint find_valnum_rhs (r: rhs) (eqs: list (valnum * rhs)) +Fixpoint find_valnum_rhs (r: rhs) (eqs: list equation) {struct eqs} : option valnum := match eqs with | nil => None - | (v, r') :: eqs1 => - if eq_rhs r r' then Some v else find_valnum_rhs r eqs1 + | Eq v str r' :: eqs1 => + if str && eq_rhs r r' then Some v else find_valnum_rhs r eqs1 + end. + +(** [find_valnum_rhs' rhs eqs] is similar, but also accepts equations + of the form [vn >= rhs]. *) + +Fixpoint find_valnum_rhs' (r: rhs) (eqs: list equation) + {struct eqs} : option valnum := + match eqs with + | nil => None + | Eq v str r' :: eqs1 => + if eq_rhs r r' then Some v else find_valnum_rhs' r eqs1 end. (** [find_valnum_num vn eqs] searches the list of equations [eqs] for an equation of the form [vn = rhs] for some equation [rhs]. If found, [Some rhs] is returned, otherwise [None] is returned. *) -Fixpoint find_valnum_num (v: valnum) (eqs: list (valnum * rhs)) +Fixpoint find_valnum_num (v: valnum) (eqs: list equation) {struct eqs} : option rhs := match eqs with | nil => None - | (v', r') :: eqs1 => - if peq v v' then Some r' else find_valnum_num v eqs1 + | Eq v' str r' :: eqs1 => + if str && peq v v' then Some r' else find_valnum_num v eqs1 + end. + +(** [reg_valnum n vn] returns a register that is mapped to value number + [vn], or [None] if no such register exists. *) + +Definition reg_valnum (n: numbering) (vn: valnum) : option reg := + match PMap.get vn n.(num_val) with + | nil => None + | r :: rs => Some r + end. + +(** [regs_valnums] is similar, for a list of value numbers. *) + +Fixpoint regs_valnums (n: numbering) (vl: list valnum) : option (list reg) := + match vl with + | nil => Some nil + | v1 :: vs => + match reg_valnum n v1, regs_valnums n vs with + | Some r1, Some rs => Some (r1 :: rs) + | _, _ => None + end + end. + +(** [find_rhs] return a register that already holds the result of the + given arithmetic operation or memory load, or a value more defined + than this result, according to the given + numbering. [None] is returned if no such register exists. *) + +Definition find_rhs (n: numbering) (rh: rhs) : option reg := + match find_valnum_rhs' rh n.(num_eqs) with + | None => None + | Some vres => reg_valnum n vres end. (** Update the [num_val] mapping prior to a redefinition of register [r]. *) @@ -163,14 +167,15 @@ Definition update_reg (n: numbering) (rd: reg) (vn: valnum) : PMap.t (list reg) Definition add_rhs (n: numbering) (rd: reg) (rh: rhs) : numbering := match find_valnum_rhs rh n.(num_eqs) with | Some vres => - mknumbering n.(num_next) n.(num_eqs) - (PTree.set rd vres n.(num_reg)) - (update_reg n rd vres) + {| num_next := n.(num_next); + num_eqs := n.(num_eqs); + num_reg := PTree.set rd vres n.(num_reg); + num_val := update_reg n rd vres |} | None => - mknumbering (Psucc n.(num_next)) - ((n.(num_next), rh) :: n.(num_eqs)) - (PTree.set rd n.(num_next) n.(num_reg)) - (update_reg n rd n.(num_next)) + {| num_next := Psucc n.(num_next); + num_eqs := Eq n.(num_next) true rh :: n.(num_eqs); + num_reg := PTree.set rd n.(num_next) n.(num_reg); + num_val := update_reg n rd n.(num_next) |} end. (** [add_op n rd op rs] specializes [add_rhs] for the case of an @@ -193,8 +198,10 @@ Definition add_op (n: numbering) (rd: reg) (op: operation) (rs: list reg) := match is_move_operation op rs with | Some r => let (n1, v) := valnum_reg n r in - mknumbering n1.(num_next) n1.(num_eqs) - (PTree.set rd v n1.(num_reg)) (update_reg n1 rd v) + {| num_next := n1.(num_next); + num_eqs := n1.(num_eqs); + num_reg := PTree.set rd v n1.(num_reg); + num_val := update_reg n1 rd v |} | None => let (n1, vs) := valnum_regs n rs in add_rhs n1 rd (Op op vs) @@ -211,30 +218,32 @@ Definition add_load (n: numbering) (rd: reg) let (n1, vs) := valnum_regs n rs in add_rhs n1 rd (Load chunk addr vs). -(** [add_unknown n rd] returns a numbering where [rd] is mapped to a - fresh value number, and no equations are added. This is useful - to model instructions with unpredictable results such as [Ibuiltin]. *) +(** [set_unknown n rd] returns a numbering where [rd] is mapped to + no value number, and no equations are added. This is useful + to model instructions with unpredictable results such as [Ibuiltin]. *) -Definition add_unknown (n: numbering) (rd: reg) := - mknumbering (Psucc n.(num_next)) - n.(num_eqs) - (PTree.set rd n.(num_next) n.(num_reg)) - (forget_reg n rd). +Definition set_unknown (n: numbering) (rd: reg) := + {| num_next := n.(num_next); + num_eqs := n.(num_eqs); + num_reg := PTree.remove rd n.(num_reg); + num_val := forget_reg n rd |}. (** [kill_equations pred n] remove all equations satisfying predicate [pred]. *) -Fixpoint kill_eqs (pred: rhs -> bool) (eqs: list (valnum * rhs)) : list (valnum * rhs) := +Fixpoint kill_eqs (pred: rhs -> bool) (eqs: list equation) : list equation := match eqs with | nil => nil - | eq :: rem => if pred (snd eq) then kill_eqs pred rem else eq :: kill_eqs pred rem + | (Eq l strict r) as eq :: rem => + if pred r then kill_eqs pred rem else eq :: kill_eqs pred rem end. Definition kill_equations (pred: rhs -> bool) (n: numbering) : numbering := - mknumbering n.(num_next) - (kill_eqs pred n.(num_eqs)) - n.(num_reg) n.(num_val). + {| num_next := n.(num_next); + num_eqs := kill_eqs pred n.(num_eqs); + num_reg := n.(num_reg); + num_val := n.(num_val) |}. -(** [kill_loads n] removes all equations involving memory loads, +(** [kill_all_loads n] removes all equations involving memory loads, as well as those involving memory-dependent operators. It is used to reflect the effect of a builtin operation, which can change memory in unpredictable ways and potentially invalidate all such equations. *) @@ -245,63 +254,121 @@ Definition filter_loads (r: rhs) : bool := | Load _ _ _ => true end. -Definition kill_loads (n: numbering) : numbering := +Definition kill_all_loads (n: numbering) : numbering := kill_equations filter_loads n. -(** [add_store n chunk addr rargs rsrc] updates the numbering [n] to reflect - the effect of a store instruction [Istore chunk addr rargs rsrc]. - Equations involving the memory state are removed from [n], unless we - can prove that the store does not invalidate them. - Then, an equations [rsrc = Load chunk addr rargs] is added to reflect - the known content of the stored memory area, but only if [chunk] is - a "full-size" quantity ([Mint32] or [Mfloat64] or [Mint64]). *) +(** [kill_loads_after_store app n chunk addr args] removes all equations + involving loads that could be invalidated by a store of quantity [chunk] + at address determined by [addr] and [args]. Loads that are disjoint + from this store are preserved. Equations involving memory-dependent + operators are also removed. *) -Definition filter_after_store (chunk: memory_chunk) (addr: addressing) (vl: list valnum) (r: rhs) : bool := +Definition filter_after_store (app: VA.t) (n: numbering) (p: aptr) (sz: Z) (r: rhs) := match r with - | Op op vl' => op_depends_on_memory op - | Load chunk' addr' vl' => - negb(eq_list_valnum vl vl' && addressing_separated chunk addr chunk' addr') + | Op op vl => + op_depends_on_memory op + | Load chunk addr vl => + match regs_valnums n vl with + | None => true + | Some rl => + negb (pdisjoint (aaddressing app addr rl) (size_chunk chunk) p sz) + end end. -Definition add_store (n: numbering) (chunk: memory_chunk) (addr: addressing) - (rargs: list reg) (rsrc: reg) : numbering := - let (n1, vargs) := valnum_regs n rargs in - let n2 := kill_equations (filter_after_store chunk addr vargs) n1 in +Definition kill_loads_after_store + (app: VA.t) (n: numbering) + (chunk: memory_chunk) (addr: addressing) (args: list reg) := + let p := aaddressing app addr args in + kill_equations (filter_after_store app n p (size_chunk chunk)) n. + +(** [add_store_result n chunk addr rargs rsrc] updates the numbering [n] + to reflect the knowledge gained after executing an instruction + [Istore chunk addr rargs rsrc]. An equation [vsrc >= Load chunk addr vargs] + is added, but only if the value of [rsrc] is known to be normalized + with respect to [chunk]. *) + +Definition store_normalized_range (chunk: memory_chunk) : aval := match chunk with - | Mint32 | Mint64 | Mfloat64 | Mfloat64al32 => add_rhs n2 rsrc (Load chunk addr vargs) - | _ => n2 + | Mint8signed => Sgn 8 + | Mint8unsigned => Uns 8 + | Mint16signed => Sgn 16 + | Mint16unsigned => Uns 16 + | Mfloat32 => Fsingle + | _ => Vtop end. -(** [reg_valnum n vn] returns a register that is mapped to value number - [vn], or [None] if no such register exists. *) - -Definition reg_valnum (n: numbering) (vn: valnum) : option reg := - match PMap.get vn n.(num_val) with - | nil => None - | r :: rs => Some r +Definition add_store_result (app: VA.t) (n: numbering) (chunk: memory_chunk) (addr: addressing) + (rargs: list reg) (rsrc: reg) := + if vincl (avalue app rsrc) (store_normalized_range chunk) then + let (n1, vsrc) := valnum_reg n rsrc in + let (n2, vargs) := valnum_regs n1 rargs in + {| num_next := n2.(num_next); + num_eqs := Eq vsrc false (Load chunk addr vargs) :: n2.(num_eqs); + num_reg := n2.(num_reg); + num_val := n2.(num_val) |} + else n. + +(** [kill_loads_after_storebyte app n dst sz] removes all equations + involving loads that could be invalidated by a store of [sz] bytes + starting at address [dst]. Loads that are disjoint from this + store-bytes are preserved. Equations involving memory-dependent + operators are also removed. *) + +Definition kill_loads_after_storebytes + (app: VA.t) (n: numbering) (dst: reg) (sz: Z) := + let p := aaddr app dst in + kill_equations (filter_after_store app n p sz) n. + +(** [add_memcpy app n1 n2 rsrc rdst sz] adds equations to [n2] that + represent the effect of a [memcpy] block copy operation of [sz] bytes + from the address denoted by [rsrc] to the address denoted by [rdst]. + [n2] is the numbering returned by [kill_loads_after_storebytes] + and [n1] is the original numbering before the [memcpy] operation. + Valid equations (found in [n1]) involving loads within the source + area of the [memcpy] are translated as equations involving loads + within the destination area, and added to numbering [n2]. + Currently, we only track [memcpy] operations between stack + locations, as often occur when compiling assignments between local C + variables of struct type. *) + +Definition shift_memcpy_eq (src sz delta: Z) (e: equation) := + match e with + | Eq l strict (Load chunk (Ainstack i) _) => + let i := Int.unsigned i in + let j := i + delta in + if zle src i + && zle (i + size_chunk chunk) (src + sz) + && zeq (Zmod delta (align_chunk chunk)) 0 + && zle 0 j + && zle j Int.max_unsigned + then Some(Eq l strict (Load chunk (Ainstack (Int.repr j)) nil)) + else None + | _ => None end. -Fixpoint regs_valnums (n: numbering) (vl: list valnum) : option (list reg) := - match vl with - | nil => Some nil - | v1 :: vs => - match reg_valnum n v1, regs_valnums n vs with - | Some r1, Some rs => Some (r1 :: rs) - | _, _ => None +Fixpoint add_memcpy_eqs (src sz delta: Z) (eqs1 eqs2: list equation) := + match eqs1 with + | nil => eqs2 + | e :: eqs => + match shift_memcpy_eq src sz delta e with + | None => add_memcpy_eqs src sz delta eqs eqs2 + | Some e' => e' :: add_memcpy_eqs src sz delta eqs eqs2 end end. -(** [find_rhs] return a register that already holds the result of the given arithmetic - operation or memory load, according to the given numbering. - [None] is returned if no such register exists. *) - -Definition find_rhs (n: numbering) (rh: rhs) : option reg := - match find_valnum_rhs rh n.(num_eqs) with - | None => None - | Some vres => reg_valnum n vres +Definition add_memcpy (app: VA.t) (n1 n2: numbering) (rsrc rdst: reg) (sz: Z) := + match aaddr app rsrc, aaddr app rdst with + | Stk src, Stk dst => + {| num_next := n2.(num_next); + num_eqs := add_memcpy_eqs (Int.unsigned src) sz + (Int.unsigned dst - Int.unsigned src) + n1.(num_eqs) n2.(num_eqs); + num_reg := n2.(num_reg); + num_val := n2.(num_val) |} + | _, _ => n2 end. -(** Experimental: take advantage of known equations to select more efficient +(** Take advantage of known equations to select more efficient forms of operations, addressing modes, and conditions. *) Section REDUCE. @@ -336,52 +403,6 @@ End REDUCE. (** * The static analysis *) -(** We now define a notion of satisfiability of a numbering. This semantic - notion plays a central role in the correctness proof (see [CSEproof]), - but is defined here because we need it to define the ordering over - numberings used in the static analysis. - - A numbering is satisfiable in a given register environment and memory - state if there exists a valuation, mapping value numbers to actual values, - that validates both the equations and the association of registers - to value numbers. *) - -Definition equation_holds - (valuation: valnum -> val) - (ge: genv) (sp: val) (m: mem) - (vres: valnum) (rh: rhs) : Prop := - match rh with - | Op op vl => - eval_operation ge sp op (List.map valuation vl) m = - Some (valuation vres) - | Load chunk addr vl => - exists a, - eval_addressing ge sp addr (List.map valuation vl) = Some a /\ - Mem.loadv chunk m a = Some (valuation vres) - end. - -Definition numbering_holds - (valuation: valnum -> val) - (ge: genv) (sp: val) (rs: regset) (m: mem) (n: numbering) : Prop := - (forall vn rh, - In (vn, rh) n.(num_eqs) -> - equation_holds valuation ge sp m vn rh) - /\ (forall r vn, - PTree.get r n.(num_reg) = Some vn -> rs#r = valuation vn). - -Definition numbering_satisfiable - (ge: genv) (sp: val) (rs: regset) (m: mem) (n: numbering) : Prop := - exists valuation, numbering_holds valuation ge sp rs m n. - -Lemma empty_numbering_satisfiable: - forall ge sp rs m, numbering_satisfiable ge sp rs m empty_numbering. -Proof. - intros; red. - exists (fun (vn: valnum) => Vundef). split; simpl; intros. - elim H. - rewrite PTree.gempty in H. discriminate. -Qed. - (** We now equip the type [numbering] with a partial order and a greatest element. The partial order is based on entailment: [n1] is greater than [n2] if [n1] is satisfiable whenever [n2] is. The greatest element @@ -390,13 +411,13 @@ Qed. Module Numbering. Definition t := numbering. Definition ge (n1 n2: numbering) : Prop := - forall ge sp rs m, - numbering_satisfiable ge sp rs m n2 -> - numbering_satisfiable ge sp rs m n1. + forall valu ge sp rs m, + numbering_holds valu ge sp rs m n2 -> + numbering_holds valu ge sp rs m n1. Definition top := empty_numbering. Lemma top_ge: forall x, ge top x. Proof. - intros; red; intros. unfold top. apply empty_numbering_satisfiable. + intros; red; intros. unfold top. apply empty_numbering_holds. Qed. Lemma refl_ge: forall x, ge x x. Proof. @@ -414,8 +435,9 @@ Module Solver := BBlock_solver(Numbering). numbering ``before''. For [Iop] and [Iload] instructions, we add equations or reuse existing value numbers as described for [add_op] and [add_load]. For [Istore] instructions, we forget - all equations involving memory loads. For [Icall] instructions, - we could simply associate a fresh, unconstrained by equations value number + equations involving memory loads at possibly overlapping locations, + then add an equation for loads from the same location stored to. + For [Icall] instructions, we could simply associate a fresh, unconstrained by equations value number to the result register. However, it is often undesirable to eliminate common subexpressions across a function call (there is a risk of increasing too much the register pressure across the call), so we @@ -428,13 +450,13 @@ Module Solver := BBlock_solver(Numbering). turned into function calls ([EF_external], [EF_malloc], [EF_free]). - Forget equations involving loads but keep equations over registers. This is appropriate for builtins that can modify memory, - e.g. volatile stores, or [EF_memcpy], or [EF_builtin] + e.g. volatile stores, or [EF_builtin] - Keep all equations, taking advantage of the fact that neither memory nor registers are modified. This is appropriate for annotations and for volatile loads. *) -Definition transfer (f: function) (pc: node) (before: numbering) := +Definition transfer (f: function) (approx: PMap.t VA.t) (pc: node) (before: numbering) := match f.(fn_code)!pc with | None => before | Some i => @@ -446,7 +468,9 @@ Definition transfer (f: function) (pc: node) (before: numbering) := | Iload chunk addr args dst s => add_load before dst chunk addr args | Istore chunk addr args src s => - add_store before chunk addr args src + let app := approx!!pc in + let n := kill_loads_after_store app before chunk addr args in + add_store_result app n chunk addr args src | Icall sig ros args res s => empty_numbering | Itailcall sig ros args => @@ -455,12 +479,19 @@ Definition transfer (f: function) (pc: node) (before: numbering) := match ef with | EF_external _ _ | EF_malloc | EF_free | EF_inline_asm _ => empty_numbering - | EF_vstore _ | EF_vstore_global _ _ _ - | EF_builtin _ _ | EF_memcpy _ _ => - add_unknown (kill_loads before) res - | EF_vload _ | EF_vload_global _ _ _ - | EF_annot _ _ | EF_annot_val _ _ => - add_unknown before res + | EF_builtin _ _ | EF_vstore _ | EF_vstore_global _ _ _ => + set_unknown (kill_all_loads before) res + | EF_memcpy sz al => + match args with + | rdst :: rsrc :: nil => + let app := approx!!pc in + let n := kill_loads_after_storebytes app before rdst sz in + set_unknown (add_memcpy app before n rsrc rdst sz) res + | _ => + empty_numbering + end + | EF_vload _ | EF_vload_global _ _ _ | EF_annot _ _ | EF_annot_val _ _ => + set_unknown before res end | Icond cond args ifso ifnot => before @@ -476,8 +507,8 @@ Definition transfer (f: function) (pc: node) (before: numbering) := which produces sub-optimal solutions quickly. The result is a mapping from program points to numberings. *) -Definition analyze (f: RTL.function): option (PMap.t numbering) := - Solver.fixpoint (fn_code f) successors_instr (transfer f) f.(fn_entrypoint). +Definition analyze (f: RTL.function) (approx: PMap.t VA.t): option (PMap.t numbering) := + Solver.fixpoint (fn_code f) successors_instr (transfer f approx) f.(fn_entrypoint). (** * Code transformation *) @@ -526,25 +557,24 @@ Definition transf_instr (n: numbering) (instr: instruction) := Definition transf_code (approxs: PMap.t numbering) (instrs: code) : code := PTree.map (fun pc instr => transf_instr approxs!!pc instr) instrs. -Definition transf_function (f: function) : res function := - match type_function f with - | Error msg => Error msg - | OK tyenv => - match analyze f with - | None => Error (msg "CSE failure") - | Some approxs => - OK(mkfunction - f.(fn_sig) - f.(fn_params) - f.(fn_stacksize) - (transf_code approxs f.(fn_code)) - f.(fn_entrypoint)) - end +Definition vanalyze := ValueAnalysis.analyze. + +Definition transf_function (rm: romem) (f: function) : res function := + let approx := vanalyze rm f in + match analyze f approx with + | None => Error (msg "CSE failure") + | Some approxs => + OK(mkfunction + f.(fn_sig) + f.(fn_params) + f.(fn_stacksize) + (transf_code approxs f.(fn_code)) + f.(fn_entrypoint)) end. -Definition transf_fundef (f: fundef) : res fundef := - AST.transf_partial_fundef transf_function f. +Definition transf_fundef (rm: romem) (f: fundef) : res fundef := + AST.transf_partial_fundef (transf_function rm) f. Definition transf_program (p: program) : res program := - transform_partial_program transf_fundef p. + transform_partial_program (transf_fundef (romem_for_program p)) p. diff --git a/backend/CSEdomain.v b/backend/CSEdomain.v new file mode 100644 index 0000000..6a75d51 --- /dev/null +++ b/backend/CSEdomain.v @@ -0,0 +1,147 @@ +(* *********************************************************************) +(* *) +(* 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. *) +(* *) +(* *********************************************************************) + +(** The abstract domain for value numbering, used in common + subexpression elimination. *) + +Require Import Coqlib. +Require Import Maps. +Require Import AST. +Require Import Values. +Require Import Memory. +Require Import Op. +Require Import Registers. +Require Import RTL. + +(** Value numbers are represented by positive integers. Equations are + of the form [valnum = rhs] or [valnum >= rhs], where the right-hand + sides [rhs] are either arithmetic operations or memory loads, [=] is + strict equality of values, and [>=] is the "more defined than" relation + over values. *) + +Definition valnum := positive. + +Inductive rhs : Type := + | Op: operation -> list valnum -> rhs + | Load: memory_chunk -> addressing -> list valnum -> rhs. + +Inductive equation : Type := + | Eq (v: valnum) (strict: bool) (r: rhs). + +Definition eq_valnum: forall (x y: valnum), {x=y}+{x<>y} := peq. + +Definition eq_list_valnum: forall (x y: list valnum), {x=y}+{x<>y} := list_eq_dec peq. + +Definition eq_rhs (x y: rhs) : {x=y}+{x<>y}. +Proof. + generalize chunk_eq eq_operation eq_addressing eq_valnum eq_list_valnum. + decide equality. +Defined. + +(** A value numbering is a collection of equations between value numbers + plus a partial map from registers to value numbers. Additionally, + we maintain the next unused value number, so as to easily generate + fresh value numbers. We also maintain a reverse mapping from value + numbers to registers, redundant with the mapping from registers to + value numbers, in order to speed up some operations. *) + +Record numbering : Type := mknumbering { + num_next: valnum; (**r first unused value number *) + num_eqs: list equation; (**r valid equations *) + num_reg: PTree.t valnum; (**r mapping register to valnum *) + num_val: PMap.t (list reg) (**r reverse mapping valnum to regs containing it *) +}. + +Definition empty_numbering := + {| num_next := 1%positive; + num_eqs := nil; + num_reg := PTree.empty _; + num_val := PMap.init nil |}. + +(** A numbering is well formed if all value numbers mentioned are below + [num_next]. Moreover, the [num_val] reverse mapping must be consistent + with the [num_reg] direct mapping. *) + +Definition valnums_rhs (r: rhs): list valnum := + match r with + | Op op vl => vl + | Load chunk addr vl => vl + end. + +Definition wf_rhs (next: valnum) (r: rhs) : Prop := +forall v, In v (valnums_rhs r) -> Plt v next. + +Definition wf_equation (next: valnum) (e: equation) : Prop := + match e with Eq l str r => Plt l next /\ wf_rhs next r end. + +Record wf_numbering (n: numbering) : Prop := { + wf_num_eqs: forall e, + In e n.(num_eqs) -> wf_equation n.(num_next) e; + wf_num_reg: forall r v, + PTree.get r n.(num_reg) = Some v -> Plt v n.(num_next); + wf_num_val: forall r v, + In r (PMap.get v n.(num_val)) -> PTree.get r n.(num_reg) = Some v +}. + +Hint Resolve wf_num_eqs wf_num_reg wf_num_val: cse. + +(** Satisfiability of numberings. A numbering holds in a concrete + execution state if there exists a valuation assigning values to + value numbers that satisfies the equations and register mapping + of the numbering. *) + +Definition valuation := valnum -> val. + +Inductive rhs_eval_to (valu: valuation) (ge: genv) (sp: val) (m: mem): + rhs -> val -> Prop := + | op_eval_to: forall op vl v, + eval_operation ge sp op (map valu vl) m = Some v -> + rhs_eval_to valu ge sp m (Op op vl) v + | load_eval_to: forall chunk addr vl a v, + eval_addressing ge sp addr (map valu vl) = Some a -> + Mem.loadv chunk m a = Some v -> + rhs_eval_to valu ge sp m (Load chunk addr vl) v. + +Inductive equation_holds (valu: valuation) (ge: genv) (sp: val) (m: mem): + equation -> Prop := + | eq_holds_strict: forall l r, + rhs_eval_to valu ge sp m r (valu l) -> + equation_holds valu ge sp m (Eq l true r) + | eq_holds_lessdef: forall l r v, + rhs_eval_to valu ge sp m r v -> Val.lessdef v (valu l) -> + equation_holds valu ge sp m (Eq l false r). + +Record numbering_holds (valu: valuation) (ge: genv) (sp: val) + (rs: regset) (m: mem) (n: numbering) : Prop := { + num_holds_wf: + wf_numbering n; + num_holds_eq: forall eq, + In eq n.(num_eqs) -> equation_holds valu ge sp m eq; + num_holds_reg: forall r v, + n.(num_reg)!r = Some v -> rs#r = valu v +}. + +Hint Resolve num_holds_wf num_holds_eq num_holds_reg: cse. + +Lemma empty_numbering_holds: + forall valu ge sp rs m, + numbering_holds valu ge sp rs m empty_numbering. +Proof. + intros; split; simpl; intros. +- split; simpl; intros. + + contradiction. + + rewrite PTree.gempty in H; discriminate. + + rewrite PMap.gi in H; contradiction. +- contradiction. +- rewrite PTree.gempty in H; discriminate. +Qed. + diff --git a/backend/CSEproof.v b/backend/CSEproof.v index 478b6f0..6c9331a 100644 --- a/backend/CSEproof.v +++ b/backend/CSEproof.v @@ -16,6 +16,8 @@ Require Import Coqlib. Require Import Maps. Require Import AST. Require Import Errors. +Require Import Integers. +Require Import Floats. Require Import Values. Require Import Memory. Require Import Events. @@ -24,713 +26,692 @@ Require Import Smallstep. Require Import Op. Require Import Registers. Require Import RTL. -Require Import RTLtyping. Require Import Kildall. +Require Import ValueDomain. +Require Import ValueAOp. +Require Import ValueAnalysis. +Require Import CSEdomain. Require Import CombineOp. Require Import CombineOpproof. Require Import CSE. -(** * Semantic properties of value numberings *) +(** * Soundness of operations over value numberings *) -(** ** Well-formedness of numberings *) +Remark wf_equation_incr: + forall next1 next2 e, + wf_equation next1 e -> Ple next1 next2 -> wf_equation next2 e. +Proof. + unfold wf_equation; intros; destruct e. destruct H. split. + apply Plt_le_trans with next1; auto. + red; intros. apply Plt_le_trans with next1; auto. apply H1; auto. +Qed. -(** A numbering is well-formed if all registers mentioned in equations - are less than the ``next'' register number given in the numbering. - This guarantees that the next register is fresh with respect to - the equations. *) +(** Extensionality with respect to valuations. *) -Definition wf_rhs (next: valnum) (rh: rhs) : Prop := - match rh with - | Op op vl => forall v, In v vl -> Plt v next - | Load chunk addr vl => forall v, In v vl -> Plt v next - end. +Definition valu_agree (valu1 valu2: valuation) (upto: valnum) := + forall v, Plt v upto -> valu2 v = valu1 v. -Definition wf_equation (next: valnum) (vr: valnum) (rh: rhs) : Prop := - Plt vr next /\ wf_rhs next rh. +Section EXTEN. -Inductive wf_numbering (n: numbering) : Prop := - wf_numbering_intro - (EQS: forall v rh, - In (v, rh) n.(num_eqs) -> wf_equation n.(num_next) v rh) - (REG: forall r v, - PTree.get r n.(num_reg) = Some v -> Plt v n.(num_next)) - (VAL: forall r v, - In r (PMap.get v n.(num_val)) -> PTree.get r n.(num_reg) = Some v). +Variable valu1: valuation. +Variable upto: valnum. +Variable valu2: valuation. +Hypothesis AGREE: valu_agree valu1 valu2 upto. +Variable ge: genv. +Variable sp: val. +Variable rs: regset. +Variable m: mem. -Lemma wf_empty_numbering: - wf_numbering empty_numbering. +Lemma valnums_val_exten: + forall vl, + (forall v, In v vl -> Plt v upto) -> + map valu2 vl = map valu1 vl. Proof. - unfold empty_numbering; split; simpl; intros. - contradiction. - rewrite PTree.gempty in H. congruence. - rewrite PMap.gi in H. contradiction. + intros. apply list_map_exten. intros. symmetry. auto. Qed. -Lemma wf_rhs_increasing: - forall next1 next2 rh, - Ple next1 next2 -> - wf_rhs next1 rh -> wf_rhs next2 rh. +Lemma rhs_eval_to_exten: + forall r v, + rhs_eval_to valu1 ge sp m r v -> + (forall v, In v (valnums_rhs r) -> Plt v upto) -> + rhs_eval_to valu2 ge sp m r v. Proof. - intros; destruct rh; simpl; intros; apply Plt_Ple_trans with next1; auto. + intros. inv H; simpl in *. +- constructor. rewrite valnums_val_exten by assumption. auto. +- econstructor; eauto. rewrite valnums_val_exten by assumption. auto. Qed. -Lemma wf_equation_increasing: - forall next1 next2 vr rh, - Ple next1 next2 -> - wf_equation next1 vr rh -> wf_equation next2 vr rh. +Lemma equation_holds_exten: + forall e, + equation_holds valu1 ge sp m e -> + wf_equation upto e -> + equation_holds valu2 ge sp m e. Proof. - intros. destruct H0. split. - apply Plt_Ple_trans with next1; auto. - apply wf_rhs_increasing with next1; auto. + intros. destruct e. destruct H0. inv H. +- constructor. rewrite AGREE by auto. apply rhs_eval_to_exten; auto. +- econstructor. apply rhs_eval_to_exten; eauto. rewrite AGREE by auto. auto. Qed. -(** We now show that all operations over numberings - preserve well-formedness. *) - -Lemma wf_valnum_reg: - forall n r n' v, - wf_numbering n -> - valnum_reg n r = (n', v) -> - wf_numbering n' /\ Plt v n'.(num_next) /\ Ple n.(num_next) n'.(num_next). +Lemma numbering_holds_exten: + forall n, + numbering_holds valu1 ge sp rs m n -> + Ple n.(num_next) upto -> + numbering_holds valu2 ge sp rs m n. Proof. - intros until v. unfold valnum_reg. intros WF VREG. inversion WF. - destruct ((num_reg n)!r) as [v'|] eqn:?. -(* found *) - inv VREG. split. auto. split. eauto. apply Ple_refl. -(* not found *) - inv VREG. split. - split; simpl; intros. - apply wf_equation_increasing with (num_next n). apply Ple_succ. auto. - rewrite PTree.gsspec in H. destruct (peq r0 r). - inv H. apply Plt_succ. - apply Plt_trans_succ; eauto. - rewrite PMap.gsspec in H. destruct (peq v (num_next n)). - subst v. simpl in H. destruct H. subst r0. apply PTree.gss. contradiction. - rewrite PTree.gso. eauto. exploit VAL; eauto. congruence. - simpl. split. apply Plt_succ. apply Ple_succ. + intros. destruct H. constructor; intros. +- auto. +- apply equation_holds_exten. auto. + eapply wf_equation_incr; eauto with cse. +- rewrite AGREE. eauto. eapply Plt_le_trans; eauto. eapply wf_num_reg; eauto. Qed. -Lemma wf_valnum_regs: - forall rl n n' vl, - wf_numbering n -> - valnum_regs n rl = (n', vl) -> - wf_numbering n' /\ - (forall v, In v vl -> Plt v n'.(num_next)) /\ - Ple n.(num_next) n'.(num_next). -Proof. - induction rl; intros. - simpl in H0. inversion H0. subst n'; subst vl. - simpl. intuition. - simpl in H0. - caseEq (valnum_reg n a). intros n1 v1 EQ1. - caseEq (valnum_regs n1 rl). intros ns vs EQS. - rewrite EQ1 in H0; rewrite EQS in H0; simpl in H0. - inversion H0. subst n'; subst vl. - generalize (wf_valnum_reg _ _ _ _ H EQ1); intros [A1 [B1 C1]]. - generalize (IHrl _ _ _ A1 EQS); intros [As [Bs Cs]]. - split. auto. - split. simpl; intros. elim H1; intro. - subst v. eapply Plt_Ple_trans; eauto. - auto. - eapply Ple_trans; eauto. -Qed. +End EXTEN. -Lemma find_valnum_rhs_correct: - forall rh vn eqs, - find_valnum_rhs rh eqs = Some vn -> In (vn, rh) eqs. -Proof. - induction eqs; simpl. - congruence. - case a; intros v r'. case (eq_rhs rh r'); intro. - intro. left. congruence. - intro. right. auto. -Qed. +Ltac splitall := repeat (match goal with |- _ /\ _ => split end). -Lemma find_valnum_num_correct: - forall rh vn eqs, - find_valnum_num vn eqs = Some rh -> In (vn, rh) eqs. +Lemma valnum_reg_holds: + forall valu1 ge sp rs m n r n' v, + numbering_holds valu1 ge sp rs m n -> + valnum_reg n r = (n', v) -> + exists valu2, + numbering_holds valu2 ge sp rs m n' + /\ rs#r = valu2 v + /\ valu_agree valu1 valu2 n.(num_next) + /\ Plt v n'.(num_next) + /\ Ple n.(num_next) n'.(num_next). Proof. - induction eqs; simpl. - congruence. - destruct a as [v' r']. destruct (peq vn v'); intros. - left. congruence. - right. auto. + unfold valnum_reg; intros. + destruct (num_reg n)!r as [v'|] eqn:NR. +- inv H0. exists valu1; splitall. + + auto. + + eauto with cse. + + red; auto. + + eauto with cse. + + apply Ple_refl. +- inv H0; simpl. + set (valu2 := fun vn => if peq vn n.(num_next) then rs#r else valu1 vn). + assert (AG: valu_agree valu1 valu2 n.(num_next)). + { red; intros. unfold valu2. apply peq_false. apply Plt_ne; auto. } + exists valu2; splitall. ++ constructor; simpl; intros. +* constructor; simpl; intros. + apply wf_equation_incr with (num_next n). eauto with cse. xomega. + rewrite PTree.gsspec in H0. destruct (peq r0 r). + inv H0; xomega. + apply Plt_trans_succ; eauto with cse. + rewrite PMap.gsspec in H0. destruct (peq v (num_next n)). + replace r0 with r by (simpl in H0; intuition). rewrite PTree.gss. subst; auto. + exploit wf_num_val; eauto with cse. intro. + rewrite PTree.gso by congruence. auto. +* eapply equation_holds_exten; eauto with cse. +* unfold valu2. rewrite PTree.gsspec in H0. destruct (peq r0 r). + inv H0. rewrite peq_true; auto. + rewrite peq_false. eauto with cse. apply Plt_ne; eauto with cse. ++ unfold valu2. rewrite peq_true; auto. ++ auto. ++ xomega. ++ xomega. Qed. -Remark in_remove: - forall (A: Type) (eq: forall (x y: A), {x=y}+{x<>y}) x y l, - In y (List.remove eq x l) <-> x <> y /\ In y l. +Lemma valnum_regs_holds: + forall rl valu1 ge sp rs m n n' vl, + numbering_holds valu1 ge sp rs m n -> + valnum_regs n rl = (n', vl) -> + exists valu2, + numbering_holds valu2 ge sp rs m n' + /\ rs##rl = map valu2 vl + /\ valu_agree valu1 valu2 n.(num_next) + /\ (forall v, In v vl -> Plt v n'.(num_next)) + /\ Ple n.(num_next) n'.(num_next). Proof. - induction l; simpl. - tauto. - destruct (eq x a). - subst a. rewrite IHl. tauto. - simpl. rewrite IHl. intuition congruence. + induction rl; simpl; intros. +- inv H0. exists valu1; splitall; auto. red; auto. simpl; tauto. xomega. +- destruct (valnum_reg n a) as [n1 v1] eqn:V1. + destruct (valnum_regs n1 rl) as [n2 vs] eqn:V2. + inv H0. + exploit valnum_reg_holds; eauto. + intros (valu2 & A & B & C & D & E). + exploit (IHrl valu2); eauto. + intros (valu3 & P & Q & R & S & T). + exists valu3; splitall. + + auto. + + simpl; f_equal; auto. rewrite R; auto. + + red; intros. transitivity (valu2 v); auto. apply R. xomega. + + simpl; intros. destruct H0; auto. subst v1; xomega. + + xomega. Qed. -Lemma wf_forget_reg: - forall n rd r v, - wf_numbering n -> - In r (PMap.get v (forget_reg n rd)) -> r <> rd /\ PTree.get r n.(num_reg) = Some v. +Lemma find_valnum_rhs_charact: + forall rh v eqs, + find_valnum_rhs rh eqs = Some v -> In (Eq v true rh) eqs. Proof. - unfold forget_reg; intros. inversion H. - destruct ((num_reg n)!rd) as [vd|] eqn:?. - rewrite PMap.gsspec in H0. destruct (peq v vd). - subst vd. rewrite in_remove in H0. destruct H0. split. auto. eauto. - split; eauto. exploit VAL; eauto. congruence. - split; eauto. exploit VAL; eauto. congruence. + induction eqs; simpl; intros. +- inv H. +- destruct a. destruct (strict && eq_rhs rh r) eqn:T. + + InvBooleans. inv H. left; auto. + + right; eauto. Qed. -Lemma wf_update_reg: - forall n rd vd r v, - wf_numbering n -> - In r (PMap.get v (update_reg n rd vd)) -> PTree.get r (PTree.set rd vd n.(num_reg)) = Some v. +Lemma find_valnum_rhs'_charact: + forall rh v eqs, + find_valnum_rhs' rh eqs = Some v -> exists strict, In (Eq v strict rh) eqs. Proof. - unfold update_reg; intros. inversion H. rewrite PMap.gsspec in H0. - destruct (peq v vd). - subst v; simpl in H0. destruct H0. - subst r. apply PTree.gss. - exploit wf_forget_reg; eauto. intros [A B]. rewrite PTree.gso; eauto. - exploit wf_forget_reg; eauto. intros [A B]. rewrite PTree.gso; eauto. + induction eqs; simpl; intros. +- inv H. +- destruct a. destruct (eq_rhs rh r) eqn:T. + + inv H. exists strict; auto. + + exploit IHeqs; eauto. intros [s IN]. exists s; auto. Qed. -Lemma wf_add_rhs: - forall n rd rh, - wf_numbering n -> - wf_rhs n.(num_next) rh -> - wf_numbering (add_rhs n rd rh). +Lemma find_valnum_num_charact: + forall v r eqs, find_valnum_num v eqs = Some r -> In (Eq v true r) eqs. Proof. - intros. inversion H. unfold add_rhs. - destruct (find_valnum_rhs rh n.(num_eqs)) as [vres|] eqn:?. -(* found *) - exploit find_valnum_rhs_correct; eauto. intros IN. - split; simpl; intros. - auto. - rewrite PTree.gsspec in H1. destruct (peq r rd). - inv H1. exploit EQS; eauto. intros [A B]. auto. - eauto. - eapply wf_update_reg; eauto. -(* not found *) - split; simpl; intros. - destruct H1. - inv H1. split. apply Plt_succ. apply wf_rhs_increasing with n.(num_next). apply Ple_succ. auto. - apply wf_equation_increasing with n.(num_next). apply Ple_succ. auto. - rewrite PTree.gsspec in H1. destruct (peq r rd). - inv H1. apply Plt_succ. - apply Plt_trans_succ. eauto. - eapply wf_update_reg; eauto. + induction eqs; simpl; intros. +- inv H. +- destruct a. destruct (strict && peq v v0) eqn:T. + + InvBooleans. inv H. auto. + + eauto. Qed. -Lemma wf_add_op: - forall n rd op rs, - wf_numbering n -> - wf_numbering (add_op n rd op rs). +Lemma reg_valnum_sound: + forall n v r valu ge sp rs m, + reg_valnum n v = Some r -> + numbering_holds valu ge sp rs m n -> + rs#r = valu v. Proof. - intros. unfold add_op. destruct (is_move_operation op rs) as [r|] eqn:?. -(* move *) - destruct (valnum_reg n r) as [n' v] eqn:?. - exploit wf_valnum_reg; eauto. intros [A [B C]]. inversion A. - constructor; simpl; intros. - eauto. - rewrite PTree.gsspec in H0. destruct (peq r0 rd). inv H0. auto. eauto. - eapply wf_update_reg; eauto. -(* not a move *) - destruct (valnum_regs n rs) as [n' vs] eqn:?. - exploit wf_valnum_regs; eauto. intros [A [B C]]. - eapply wf_add_rhs; eauto. + unfold reg_valnum; intros. destruct (num_val n)#v as [ | r1 rl] eqn:E; inv H. + eapply num_holds_reg; eauto. eapply wf_num_val; eauto with cse. + rewrite E; auto with coqlib. Qed. -Lemma wf_add_load: - forall n rd chunk addr rs, - wf_numbering n -> - wf_numbering (add_load n rd chunk addr rs). +Lemma regs_valnums_sound: + forall n valu ge sp rs m, + numbering_holds valu ge sp rs m n -> + forall vl rl, + regs_valnums n vl = Some rl -> + rs##rl = map valu vl. Proof. - intros. unfold add_load. - caseEq (valnum_regs n rs). intros n' vl EQ. - generalize (wf_valnum_regs _ _ _ _ H EQ). intros [A [B C]]. - apply wf_add_rhs; auto. + induction vl; simpl; intros. +- inv H0; auto. +- destruct (reg_valnum n a) as [r1|] eqn:RV1; try discriminate. + destruct (regs_valnums n vl) as [rl1|] eqn:RVL; inv H0. + simpl; f_equal. eapply reg_valnum_sound; eauto. eauto. Qed. -Lemma wf_add_unknown: - forall n rd, - wf_numbering n -> - wf_numbering (add_unknown n rd). +Lemma find_rhs_sound: + forall n rh r valu ge sp rs m, + find_rhs n rh = Some r -> + numbering_holds valu ge sp rs m n -> + exists v, rhs_eval_to valu ge sp m rh v /\ Val.lessdef v rs#r. Proof. - intros. inversion H. unfold add_unknown. constructor; simpl; intros. - eapply wf_equation_increasing; eauto. auto with coqlib. - rewrite PTree.gsspec in H0. destruct (peq r rd). - inv H0. auto with coqlib. - apply Plt_trans_succ; eauto. - exploit wf_forget_reg; eauto. intros [A B]. rewrite PTree.gso; eauto. + unfold find_rhs; intros. destruct (find_valnum_rhs' rh (num_eqs n)) as [vres|] eqn:E; try discriminate. + exploit find_valnum_rhs'_charact; eauto. intros [strict IN]. + erewrite reg_valnum_sound by eauto. + exploit num_holds_eq; eauto. intros EH. inv EH. +- exists (valu vres); auto. +- exists v; auto. Qed. -Remark kill_eqs_in: - forall pred v rhs eqs, - In (v, rhs) (kill_eqs pred eqs) -> In (v, rhs) eqs /\ pred rhs = false. +Remark in_remove: + forall (A: Type) (eq: forall (x y: A), {x=y}+{x<>y}) x y l, + In y (List.remove eq x l) <-> x <> y /\ In y l. Proof. - induction eqs; simpl; intros. + induction l; simpl. tauto. - destruct (pred (snd a)) eqn:?. - exploit IHeqs; eauto. tauto. - simpl in H; destruct H. subst a. auto. exploit IHeqs; eauto. tauto. -Qed. - -Lemma wf_kill_equations: - forall pred n, wf_numbering n -> wf_numbering (kill_equations pred n). -Proof. - intros. inversion H. unfold kill_equations; split; simpl; intros. - exploit kill_eqs_in; eauto. intros [A B]. auto. - eauto. - eauto. -Qed. - -Lemma wf_add_store: - forall n chunk addr rargs rsrc, - wf_numbering n -> wf_numbering (add_store n chunk addr rargs rsrc). -Proof. - intros. unfold add_store. - destruct (valnum_regs n rargs) as [n1 vargs] eqn:?. - exploit wf_valnum_regs; eauto. intros [A [B C]]. - assert (wf_numbering (kill_equations (filter_after_store chunk addr vargs) n1)). - apply wf_kill_equations. auto. - destruct chunk; auto; apply wf_add_rhs; auto. + destruct (eq x a). + subst a. rewrite IHl. tauto. + simpl. rewrite IHl. intuition congruence. Qed. -Lemma wf_transfer: - forall f pc n, wf_numbering n -> wf_numbering (transfer f pc n). +Lemma forget_reg_charact: + forall n rd r v, + wf_numbering n -> + In r (PMap.get v (forget_reg n rd)) -> r <> rd /\ In r (PMap.get v n.(num_val)). Proof. - intros. unfold transfer. - destruct (f.(fn_code)!pc); auto. - destruct i; auto. - apply wf_add_op; auto. - apply wf_add_load; auto. - apply wf_add_store; auto. - apply wf_empty_numbering. - apply wf_empty_numbering. - destruct e; (apply wf_empty_numbering || - apply wf_add_unknown; auto; apply wf_kill_equations; auto). -Qed. - -(** As a consequence, the numberings computed by the static analysis - are well formed. *) - -Theorem wf_analyze: - forall f approx pc, analyze f = Some approx -> wf_numbering approx!!pc. + unfold forget_reg; intros. + destruct (PTree.get rd n.(num_reg)) as [vd|] eqn:GET. +- rewrite PMap.gsspec in H0. destruct (peq v vd). + + subst v. rewrite in_remove in H0. intuition. + + split; auto. exploit wf_num_val; eauto. congruence. +- split; auto. exploit wf_num_val; eauto. congruence. +Qed. + +Lemma update_reg_charact: + forall n rd vd r v, + wf_numbering n -> + In r (PMap.get v (update_reg n rd vd)) -> + PTree.get r (PTree.set rd vd n.(num_reg)) = Some v. Proof. - unfold analyze; intros. - eapply Solver.fixpoint_invariant with (P := wf_numbering); eauto. - exact wf_empty_numbering. - intros. eapply wf_transfer; eauto. + unfold update_reg; intros. + rewrite PMap.gsspec in H0. + destruct (peq v vd). +- subst v. destruct H0. ++ subst r. apply PTree.gss. ++ exploit forget_reg_charact; eauto. intros [A B]. + rewrite PTree.gso by auto. eapply wf_num_val; eauto. +- exploit forget_reg_charact; eauto. intros [A B]. + rewrite PTree.gso by auto. eapply wf_num_val; eauto. Qed. -(** ** Properties of satisfiability of numberings *) - -Module ValnumEq. - Definition t := valnum. - Definition eq := peq. -End ValnumEq. - -Module VMap := EMap(ValnumEq). - -Section SATISFIABILITY. - -Variable ge: genv. -Variable sp: val. - -(** Agremment between two mappings from value numbers to values - up to a given value number. *) - -Definition valu_agree (valu1 valu2: valnum -> val) (upto: valnum) : Prop := - forall v, Plt v upto -> valu2 v = valu1 v. - -Lemma valu_agree_refl: - forall valu upto, valu_agree valu valu upto. +Lemma rhs_eval_to_inj: + forall valu ge sp m rh v1 v2, + rhs_eval_to valu ge sp m rh v1 -> rhs_eval_to valu ge sp m rh v2 -> v1 = v2. Proof. - intros; red; auto. + intros. inv H; inv H0; congruence. Qed. -Lemma valu_agree_trans: - forall valu1 valu2 valu3 upto12 upto23, - valu_agree valu1 valu2 upto12 -> - valu_agree valu2 valu3 upto23 -> - Ple upto12 upto23 -> - valu_agree valu1 valu3 upto12. +Lemma add_rhs_holds: + forall valu1 ge sp rs m n rd rh rs', + numbering_holds valu1 ge sp rs m n -> + rhs_eval_to valu1 ge sp m rh (rs'#rd) -> + wf_rhs n.(num_next) rh -> + (forall r, r <> rd -> rs'#r = rs#r) -> + exists valu2, numbering_holds valu2 ge sp rs' m (add_rhs n rd rh). Proof. - intros; red; intros. transitivity (valu2 v). - apply H0. apply Plt_Ple_trans with upto12; auto. - apply H; auto. -Qed. + unfold add_rhs; intros. + destruct (find_valnum_rhs rh n.(num_eqs)) as [vres|] eqn:FIND. + +- (* A value number exists already *) + exploit find_valnum_rhs_charact; eauto. intros IN. + exploit wf_num_eqs; eauto with cse. intros [A B]. + exploit num_holds_eq; eauto. intros EH. inv EH. + assert (rs'#rd = valu1 vres) by (eapply rhs_eval_to_inj; eauto). + exists valu1; constructor; simpl; intros. ++ constructor; simpl; intros. + * eauto with cse. + * rewrite PTree.gsspec in H5. destruct (peq r rd). + inv H5. auto. + eauto with cse. + * eapply update_reg_charact; eauto with cse. ++ eauto with cse. ++ rewrite PTree.gsspec in H5. destruct (peq r rd). + congruence. + rewrite H2 by auto. eauto with cse. -Lemma valu_agree_list: - forall valu1 valu2 upto vl, - valu_agree valu1 valu2 upto -> - (forall v, In v vl -> Plt v upto) -> - map valu2 vl = map valu1 vl. -Proof. - intros. apply list_map_exten. intros. symmetry. apply H. auto. +- (* Assigning a new value number *) + set (valu2 := fun v => if peq v n.(num_next) then rs'#rd else valu1 v). + assert (AG: valu_agree valu1 valu2 n.(num_next)). + { red; intros. unfold valu2. apply peq_false. apply Plt_ne; auto. } + exists valu2; constructor; simpl; intros. ++ constructor; simpl; intros. + * destruct H3. inv H3. simpl; split. xomega. + red; intros. apply Plt_trans_succ; eauto. + apply wf_equation_incr with (num_next n). eauto with cse. xomega. + * rewrite PTree.gsspec in H3. destruct (peq r rd). + inv H3. xomega. + apply Plt_trans_succ; eauto with cse. + * apply update_reg_charact; eauto with cse. ++ destruct H3. inv H3. + constructor. unfold valu2 at 2; rewrite peq_true. + eapply rhs_eval_to_exten; eauto. + eapply equation_holds_exten; eauto with cse. ++ rewrite PTree.gsspec in H3. unfold valu2. destruct (peq r rd). + inv H3. rewrite peq_true; auto. + rewrite peq_false. rewrite H2 by auto. eauto with cse. + apply Plt_ne; eauto with cse. Qed. -(** The [numbering_holds] predicate (defined in file [CSE]) is - extensional with respect to [valu_agree]. *) - -Lemma numbering_holds_exten: - forall valu1 valu2 n rs m, - valu_agree valu1 valu2 n.(num_next) -> - wf_numbering n -> +Lemma add_op_holds: + forall valu1 ge sp rs m n op (args: list reg) v dst, numbering_holds valu1 ge sp rs m n -> - numbering_holds valu2 ge sp rs m n. + eval_operation ge sp op rs##args m = Some v -> + exists valu2, numbering_holds valu2 ge sp (rs#dst <- v) m (add_op n dst op args). Proof. - intros. inversion H0. inversion H1. split; intros. - exploit EQS; eauto. intros [A B]. red in B. - generalize (H2 _ _ H4). - unfold equation_holds; destruct rh. - rewrite (valu_agree_list valu1 valu2 n.(num_next)). - rewrite H. auto. auto. auto. auto. - rewrite (valu_agree_list valu1 valu2 n.(num_next)). - rewrite H. auto. auto. auto. auto. - rewrite H. auto. eauto. + unfold add_op; intros. + destruct (is_move_operation op args) as [src|] eqn:ISMOVE. +- (* special case for moves *) + exploit is_move_operation_correct; eauto. intros [A B]; subst op args. + simpl in H0. inv H0. + destruct (valnum_reg n src) as [n1 vsrc] eqn:VN. + exploit valnum_reg_holds; eauto. + intros (valu2 & A & B & C & D & E). + exists valu2; constructor; simpl; intros. ++ constructor; simpl; intros; eauto with cse. + * rewrite PTree.gsspec in H0. destruct (peq r dst). + inv H0. auto. + eauto with cse. + * eapply update_reg_charact; eauto with cse. ++ eauto with cse. ++ rewrite PTree.gsspec in H0. rewrite Regmap.gsspec. + destruct (peq r dst). congruence. eauto with cse. + +- (* general case *) + destruct (valnum_regs n args) as [n1 vl] eqn:VN. + exploit valnum_regs_holds; eauto. + intros (valu2 & A & B & C & D & E). + eapply add_rhs_holds; eauto. ++ constructor. rewrite Regmap.gss. congruence. ++ intros. apply Regmap.gso; auto. Qed. -(** [valnum_reg] and [valnum_regs] preserve the [numbering_holds] predicate. - Moreover, it is always the case that the returned value number has - the value of the given register in the final assignment of values to - value numbers. *) - -Lemma valnum_reg_holds: - forall valu1 rs n r n' v m, - wf_numbering n -> +Lemma add_load_holds: + forall valu1 ge sp rs m n addr (args: list reg) a chunk v dst, numbering_holds valu1 ge sp rs m n -> - valnum_reg n r = (n', v) -> - exists valu2, - numbering_holds valu2 ge sp rs m n' /\ - valu2 v = rs#r /\ - valu_agree valu1 valu2 n.(num_next). + eval_addressing ge sp addr rs##args = Some a -> + Mem.loadv chunk m a = Some v -> + exists valu2, numbering_holds valu2 ge sp (rs#dst <- v) m (add_load n dst chunk addr args). Proof. - intros until v. unfold valnum_reg. - caseEq (n.(num_reg)!r). - (* Register already has a value number *) - intros. inversion H2. subst n'; subst v0. - inversion H1. - exists valu1. split. auto. - split. symmetry. auto. - apply valu_agree_refl. - (* Register gets a fresh value number *) - intros. inversion H2. subst n'. subst v. inversion H1. - set (valu2 := VMap.set n.(num_next) rs#r valu1). - assert (AG: valu_agree valu1 valu2 n.(num_next)). - red; intros. unfold valu2. apply VMap.gso. - auto with coqlib. - destruct (numbering_holds_exten _ _ _ _ _ AG H0 H1) as [A B]. - exists valu2. - split. split; simpl; intros. auto. - unfold valu2, VMap.set, ValnumEq.eq. - rewrite PTree.gsspec in H5. destruct (peq r0 r). - inv H5. rewrite peq_true. auto. - rewrite peq_false. auto. - assert (Plt vn (num_next n)). inv H0. eauto. - red; intros; subst; eapply Plt_strict; eauto. - split. unfold valu2. rewrite VMap.gss. auto. - auto. + unfold add_load; intros. + destruct (valnum_regs n args) as [n1 vl] eqn:VN. + exploit valnum_regs_holds; eauto. + intros (valu2 & A & B & C & D & E). + eapply add_rhs_holds; eauto. ++ econstructor. rewrite <- B; eauto. rewrite Regmap.gss; auto. ++ intros. apply Regmap.gso; auto. Qed. -Lemma valnum_regs_holds: - forall rs rl valu1 n n' vl m, - wf_numbering n -> - numbering_holds valu1 ge sp rs m n -> - valnum_regs n rl = (n', vl) -> - exists valu2, - numbering_holds valu2 ge sp rs m n' /\ - List.map valu2 vl = rs##rl /\ - valu_agree valu1 valu2 n.(num_next). +Lemma set_unknown_holds: + forall valu ge sp rs m n r v, + numbering_holds valu ge sp rs m n -> + numbering_holds valu ge sp (rs#r <- v) m (set_unknown n r). Proof. - induction rl; simpl; intros. - (* base case *) - inversion H1; subst n'; subst vl. - exists valu1. split. auto. split. auto. apply valu_agree_refl. - (* inductive case *) - caseEq (valnum_reg n a); intros n1 v1 EQ1. - caseEq (valnum_regs n1 rl); intros ns vs EQs. - rewrite EQ1 in H1; rewrite EQs in H1. inversion H1. subst vl; subst n'. - generalize (valnum_reg_holds _ _ _ _ _ _ _ H H0 EQ1). - intros [valu2 [A [B C]]]. - generalize (wf_valnum_reg _ _ _ _ H EQ1). intros [D [E F]]. - generalize (IHrl _ _ _ _ _ D A EQs). - intros [valu3 [P [Q R]]]. - exists valu3. - split. auto. - split. simpl. rewrite R. congruence. auto. - eapply valu_agree_trans; eauto. + intros; constructor; simpl; intros. +- constructor; simpl; intros. + + eauto with cse. + + rewrite PTree.grspec in H0. destruct (PTree.elt_eq r0 r). + discriminate. + eauto with cse. + + exploit forget_reg_charact; eauto with cse. intros [A B]. + rewrite PTree.gro; eauto with cse. +- eauto with cse. +- rewrite PTree.grspec in H0. destruct (PTree.elt_eq r0 r). + discriminate. + rewrite Regmap.gso; eauto with cse. Qed. -(** A reformulation of the [equation_holds] predicate in terms - of the value to which a right-hand side of an equation evaluates. *) - -Definition rhs_evals_to - (valu: valnum -> val) (rh: rhs) (m: mem) (v: val) : Prop := - match rh with - | Op op vl => - eval_operation ge sp op (List.map valu vl) m = Some v - | Load chunk addr vl => - exists a, - eval_addressing ge sp addr (List.map valu vl) = Some a /\ - Mem.loadv chunk m a = Some v - end. - -Lemma equation_evals_to_holds_1: - forall valu rh v vres m, - rhs_evals_to valu rh m v -> - equation_holds valu ge sp m vres rh -> - valu vres = v. +Lemma kill_eqs_charact: + forall pred l strict r eqs, + In (Eq l strict r) (kill_eqs pred eqs) -> + pred r = false /\ In (Eq l strict r) eqs. Proof. - intros until m. unfold rhs_evals_to, equation_holds. - destruct rh. congruence. - intros [a1 [A1 B1]] [a2 [A2 B2]]. congruence. + induction eqs; simpl; intros. +- tauto. +- destruct a. destruct (pred r0) eqn:PRED. + tauto. + inv H. inv H0. auto. tauto. Qed. -Lemma equation_evals_to_holds_2: - forall valu rh v vres m, - wf_rhs vres rh -> - rhs_evals_to valu rh m v -> - equation_holds (VMap.set vres v valu) ge sp m vres rh. +Lemma kill_equations_hold: + forall valu ge sp rs m n pred m', + numbering_holds valu ge sp rs m n -> + (forall r v, + pred r = false -> + rhs_eval_to valu ge sp m r v -> + rhs_eval_to valu ge sp m' r v) -> + numbering_holds valu ge sp rs m' (kill_equations pred n). Proof. - intros until m. unfold wf_rhs, rhs_evals_to, equation_holds. - rewrite VMap.gss. - assert (forall vl: list valnum, - (forall v, In v vl -> Plt v vres) -> - map (VMap.set vres v valu) vl = map valu vl). - intros. apply list_map_exten. intros. - symmetry. apply VMap.gso. apply Plt_ne. auto. - destruct rh; intros; rewrite H; auto. + intros; constructor; simpl; intros. +- constructor; simpl; intros; eauto with cse. + destruct e. exploit kill_eqs_charact; eauto. intros [A B]. eauto with cse. +- destruct eq. exploit kill_eqs_charact; eauto. intros [A B]. + exploit num_holds_eq; eauto. intro EH; inv EH; econstructor; eauto. +- eauto with cse. Qed. -(** The numbering obtained by adding an equation [rd = rhs] is satisfiable - in a concrete register set where [rd] is set to the value of [rhs]. *) - -Lemma add_rhs_satisfiable_gen: - forall n rh valu1 rs rd rs' m, - wf_numbering n -> - wf_rhs n.(num_next) rh -> - numbering_holds valu1 ge sp rs m n -> - rhs_evals_to valu1 rh m (rs'#rd) -> - (forall r, r <> rd -> rs'#r = rs#r) -> - numbering_satisfiable ge sp rs' m (add_rhs n rd rh). +Lemma kill_all_loads_hold: + forall valu ge sp rs m n m', + numbering_holds valu ge sp rs m n -> + numbering_holds valu ge sp rs m' (kill_all_loads n). Proof. - intros. unfold add_rhs. - caseEq (find_valnum_rhs rh n.(num_eqs)). - (* RHS found *) - intros vres FINDVN. inversion H1. - exists valu1. split; simpl; intros. - auto. - rewrite PTree.gsspec in H6. - destruct (peq r rd). - inv H6. - symmetry. eapply equation_evals_to_holds_1; eauto. - apply H4. apply find_valnum_rhs_correct. congruence. - rewrite H3; auto. - (* RHS not found *) - intro FINDVN. - set (valu2 := VMap.set n.(num_next) (rs'#rd) valu1). - assert (AG: valu_agree valu1 valu2 n.(num_next)). - red; intros. unfold valu2. apply VMap.gso. - auto with coqlib. - elim (numbering_holds_exten _ _ _ _ _ AG H H1); intros. - exists valu2. split; simpl; intros. - destruct H6. - inv H6. unfold valu2. eapply equation_evals_to_holds_2; eauto. auto. - rewrite PTree.gsspec in H6. destruct (peq r rd). - unfold valu2. inv H6. rewrite VMap.gss. auto. - rewrite H3; auto. + intros. eapply kill_equations_hold; eauto. + unfold filter_loads; intros. inv H1. + constructor. rewrite <- H2. apply op_depends_on_memory_correct; auto. + discriminate. Qed. -Lemma add_rhs_satisfiable: - forall n rh valu1 rs rd v m, - wf_numbering n -> - wf_rhs n.(num_next) rh -> - numbering_holds valu1 ge sp rs m n -> - rhs_evals_to valu1 rh m v -> - numbering_satisfiable ge sp (rs#rd <- v) m (add_rhs n rd rh). +Lemma kill_loads_after_store_holds: + forall valu ge sp rs m n addr args a chunk v m' bc approx ae am, + numbering_holds valu ge (Vptr sp Int.zero) rs m n -> + eval_addressing ge (Vptr sp Int.zero) addr rs##args = Some a -> + Mem.storev chunk m a v = Some m' -> + genv_match bc ge -> + bc sp = BCstack -> + ematch bc rs ae -> + approx = VA.State ae am -> + numbering_holds valu ge (Vptr sp Int.zero) rs m' + (kill_loads_after_store approx n chunk addr args). Proof. - intros. eapply add_rhs_satisfiable_gen; eauto. - rewrite Regmap.gss; auto. - intros. apply Regmap.gso; auto. + intros. apply kill_equations_hold with m; auto. + intros. unfold filter_after_store in H6; inv H7. +- constructor. rewrite <- H8. apply op_depends_on_memory_correct; auto. +- destruct (regs_valnums n vl) as [rl|] eqn:RV; try discriminate. + econstructor; eauto. rewrite <- H9. + destruct a; simpl in H1; try discriminate. + destruct a0; simpl in H9; try discriminate. + simpl. + rewrite negb_false_iff in H6. unfold aaddressing in H6. + eapply Mem.load_store_other. eauto. + eapply pdisjoint_sound. eauto. + apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto. + erewrite <- regs_valnums_sound by eauto. eauto with va. + apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto with va. Qed. -(** [add_op] returns a numbering that is satisfiable in the register - set after execution of the corresponding [Iop] instruction. *) - -Lemma add_op_satisfiable: - forall n rs op args dst v m, - wf_numbering n -> - numbering_satisfiable ge sp rs m n -> - eval_operation ge sp op rs##args m = Some v -> - numbering_satisfiable ge sp (rs#dst <- v) m (add_op n dst op args). +Lemma store_normalized_range_sound: + forall bc chunk v, + vmatch bc v (store_normalized_range chunk) -> + Val.lessdef (Val.load_result chunk v) v. Proof. - intros. inversion H0. - unfold add_op. - caseEq (@is_move_operation reg op args). - intros arg EQ. - destruct (is_move_operation_correct _ _ EQ) as [A B]. subst op args. - caseEq (valnum_reg n arg). intros n1 v1 VL. - generalize (valnum_reg_holds _ _ _ _ _ _ _ H H2 VL). intros [valu2 [A [B C]]]. - generalize (wf_valnum_reg _ _ _ _ H VL). intros [D [E F]]. - elim A; intros. exists valu2; split; simpl; intros. - auto. rewrite Regmap.gsspec. rewrite PTree.gsspec in H5. - destruct (peq r dst). simpl in H1. congruence. auto. - intro NEQ. caseEq (valnum_regs n args). intros n1 vl VRL. - generalize (valnum_regs_holds _ _ _ _ _ _ _ H H2 VRL). intros [valu2 [A [B C]]]. - generalize (wf_valnum_regs _ _ _ _ H VRL). intros [D [E F]]. - apply add_rhs_satisfiable with valu2; auto. - simpl. congruence. + intros. destruct chunk; simpl in *; destruct v; auto. +- inv H. rewrite is_sgn_sign_ext in H3 by omega. rewrite H3; auto. +- inv H. rewrite is_uns_zero_ext in H3 by omega. rewrite H3; auto. +- inv H. rewrite is_sgn_sign_ext in H3 by omega. rewrite H3; auto. +- inv H. rewrite is_uns_zero_ext in H3 by omega. rewrite H3; auto. +- inv H. rewrite Float.singleoffloat_of_single by auto. auto. Qed. -(** [add_load] returns a numbering that is satisfiable in the register - set after execution of the corresponding [Iload] instruction. *) - -Lemma add_load_satisfiable: - forall n rs chunk addr args dst a v m, - wf_numbering n -> - numbering_satisfiable ge sp rs m n -> +Lemma add_store_result_hold: + forall valu1 ge sp rs m' n addr args a chunk m src bc ae approx am, + numbering_holds valu1 ge sp rs m' n -> eval_addressing ge sp addr rs##args = Some a -> - Mem.loadv chunk m a = Some v -> - numbering_satisfiable ge sp (rs#dst <- v) m (add_load n dst chunk addr args). + Mem.storev chunk m a rs#src = Some m' -> + ematch bc rs ae -> + approx = VA.State ae am -> + exists valu2, numbering_holds valu2 ge sp rs m' (add_store_result approx n chunk addr args src). Proof. - intros. inversion H0. - unfold add_load. - caseEq (valnum_regs n args). intros n1 vl VRL. - generalize (valnum_regs_holds _ _ _ _ _ _ _ H H3 VRL). intros [valu2 [A [B C]]]. - generalize (wf_valnum_regs _ _ _ _ H VRL). intros [D [E F]]. - apply add_rhs_satisfiable with valu2; auto. - simpl. exists a; split; congruence. + unfold add_store_result; intros. + unfold avalue; rewrite H3. + destruct (vincl (AE.get src ae) (store_normalized_range chunk)) eqn:INCL. +- destruct (valnum_reg n src) as [n1 vsrc] eqn:VR1. + destruct (valnum_regs n1 args) as [n2 vargs] eqn:VR2. + exploit valnum_reg_holds; eauto. intros (valu2 & A & B & C & D & E). + exploit valnum_regs_holds; eauto. intros (valu3 & P & Q & R & S & T). + exists valu3. constructor; simpl; intros. ++ constructor; simpl; intros; eauto with cse. + destruct H4; eauto with cse. subst e. split. + eapply Plt_le_trans; eauto. + red; simpl; intros. auto. ++ destruct H4; eauto with cse. subst eq. apply eq_holds_lessdef with (Val.load_result chunk rs#src). + apply load_eval_to with a. rewrite <- Q; auto. + destruct a; try discriminate. simpl. eapply Mem.load_store_same; eauto. + rewrite B. rewrite R by auto. apply store_normalized_range_sound with bc. + rewrite <- B. eapply vmatch_ge. apply vincl_ge; eauto. apply H2. ++ eauto with cse. + +- exists valu1; auto. Qed. -(** [add_unknown] returns a numbering that is satisfiable in the - register set after setting the target register to any value. *) - -Lemma add_unknown_satisfiable: - forall n rs dst v m, - wf_numbering n -> - numbering_satisfiable ge sp rs m n -> - numbering_satisfiable ge sp (rs#dst <- v) m (add_unknown n dst). +Lemma kill_loads_after_storebytes_holds: + forall valu ge sp rs m n dst b ofs bytes m' bc approx ae am sz, + numbering_holds valu ge (Vptr sp Int.zero) rs m n -> + rs#dst = Vptr b ofs -> + Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' -> + genv_match bc ge -> + bc sp = BCstack -> + ematch bc rs ae -> + approx = VA.State ae am -> + length bytes = nat_of_Z sz -> sz >= 0 -> + numbering_holds valu ge (Vptr sp Int.zero) rs m' + (kill_loads_after_storebytes approx n dst sz). Proof. - intros. destruct H0 as [valu A]. - set (valu' := VMap.set n.(num_next) v valu). - assert (numbering_holds valu' ge sp rs m n). - eapply numbering_holds_exten; eauto. - unfold valu'; red; intros. apply VMap.gso. auto with coqlib. - destruct H0 as [B C]. - exists valu'; split; simpl; intros. - eauto. - rewrite PTree.gsspec in H0. rewrite Regmap.gsspec. - destruct (peq r dst). inversion H0. unfold valu'. rewrite VMap.gss; auto. - eauto. + intros. apply kill_equations_hold with m; auto. + intros. unfold filter_after_store in H8; inv H9. +- constructor. rewrite <- H10. apply op_depends_on_memory_correct; auto. +- destruct (regs_valnums n vl) as [rl|] eqn:RV; try discriminate. + econstructor; eauto. rewrite <- H11. + destruct a; simpl in H10; try discriminate. + simpl. + rewrite negb_false_iff in H8. + eapply Mem.load_storebytes_other. eauto. + rewrite H6. rewrite nat_of_Z_eq by auto. + eapply pdisjoint_sound. eauto. + unfold aaddressing. apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto. + erewrite <- regs_valnums_sound by eauto. eauto with va. + unfold aaddr. apply match_aptr_of_aval. rewrite <- H0. apply H4. Qed. -(** Satisfiability of [kill_equations]. *) - -Lemma kill_equations_holds: - forall pred valu n rs m m', - (forall v r, - equation_holds valu ge sp m v r -> pred r = false -> equation_holds valu ge sp m' v r) -> - numbering_holds valu ge sp rs m n -> - numbering_holds valu ge sp rs m' (kill_equations pred n). +Lemma load_memcpy: + forall m b1 ofs1 sz bytes b2 ofs2 m' chunk i v, + Mem.loadbytes m b1 ofs1 sz = Some bytes -> + Mem.storebytes m b2 ofs2 bytes = Some m' -> + Mem.load chunk m b1 i = Some v -> + ofs1 <= i -> i + size_chunk chunk <= ofs1 + sz -> + (align_chunk chunk | ofs2 - ofs1) -> + Mem.load chunk m' b2 (i + (ofs2 - ofs1)) = Some v. Proof. - intros. destruct H0 as [A B]. red; simpl. split; intros. - exploit kill_eqs_in; eauto. intros [C D]. eauto. - auto. + intros. + generalize (size_chunk_pos chunk); intros SPOS. + set (n1 := i - ofs1). + set (n2 := size_chunk chunk). + set (n3 := sz - (n1 + n2)). + replace sz with (n1 + (n2 + n3)) in H by (unfold n3, n2, n1; omega). + exploit Mem.loadbytes_split; eauto. + unfold n1; omega. + unfold n3, n2, n1; omega. + intros (bytes1 & bytes23 & LB1 & LB23 & EQ). + clear H. + exploit Mem.loadbytes_split; eauto. + unfold n2; omega. + unfold n3, n2, n1; omega. + intros (bytes2 & bytes3 & LB2 & LB3 & EQ'). + subst bytes23; subst bytes. + exploit Mem.load_loadbytes; eauto. intros (bytes2' & A & B). + assert (bytes2' = bytes2). + { replace (ofs1 + n1) with i in LB2 by (unfold n1; omega). unfold n2 in LB2. congruence. } + subst bytes2'. + exploit Mem.storebytes_split; eauto. intros (m1 & SB1 & SB23). + clear H0. + exploit Mem.storebytes_split; eauto. intros (m2 & SB2 & SB3). + clear SB23. + assert (L1: Z.of_nat (length bytes1) = n1). + { erewrite Mem.loadbytes_length by eauto. apply nat_of_Z_eq. unfold n1; omega. } + assert (L2: Z.of_nat (length bytes2) = n2). + { erewrite Mem.loadbytes_length by eauto. apply nat_of_Z_eq. unfold n2; omega. } + rewrite L1 in *. rewrite L2 in *. + assert (LB': Mem.loadbytes m2 b2 (ofs2 + n1) n2 = Some bytes2). + { rewrite <- L2. eapply Mem.loadbytes_storebytes_same; eauto. } + assert (LB'': Mem.loadbytes m' b2 (ofs2 + n1) n2 = Some bytes2). + { rewrite <- LB'. eapply Mem.loadbytes_storebytes_other; eauto. + unfold n2; omega. + right; left; omega. } + exploit Mem.load_valid_access; eauto. intros [P Q]. + rewrite B. apply Mem.loadbytes_load. + replace (i + (ofs2 - ofs1)) with (ofs2 + n1) by (unfold n1; omega). + exact LB''. + apply Z.divide_add_r; auto. Qed. -(** [kill_loads] preserves satisfiability. Moreover, the resulting numbering - is satisfiable in any concrete memory state. *) - -Lemma kill_loads_satisfiable: - forall n rs m m', - numbering_satisfiable ge sp rs m n -> - numbering_satisfiable ge sp rs m' (kill_loads n). -Proof. - intros. destruct H as [valu A]. exists valu. eapply kill_equations_holds with (m := m); eauto. - intros. destruct r; simpl in *. rewrite <- H. apply op_depends_on_memory_correct; auto. - congruence. +Lemma shift_memcpy_eq_wf: + forall src sz delta e e' next, + shift_memcpy_eq src sz delta e = Some e' -> + wf_equation next e -> + wf_equation next e'. +Proof with (try discriminate). + unfold shift_memcpy_eq; intros. + destruct e. destruct r... destruct a... + destruct (zle src (Int.unsigned i) && + zle (Int.unsigned i + size_chunk m) (src + sz) && + zeq (delta mod align_chunk m) 0 && zle 0 (Int.unsigned i + delta) && + zle (Int.unsigned i + delta) Int.max_unsigned)... + inv H. destruct H0. split. auto. red; simpl; tauto. Qed. -(** [add_store] returns a numbering that is satisfiable in the memory state - after execution of the corresponding [Istore] instruction. *) - -Lemma add_store_satisfiable: - forall n rs chunk addr args src a m m', - wf_numbering n -> - numbering_satisfiable ge sp rs m n -> - eval_addressing ge sp addr rs##args = Some a -> - Mem.storev chunk m a (rs#src) = Some m' -> - Val.has_type (rs#src) (type_of_chunk_use chunk) -> - numbering_satisfiable ge sp rs m' (add_store n chunk addr args src). -Proof. - intros. unfold add_store. destruct H0 as [valu A]. - destruct (valnum_regs n args) as [n1 vargs] eqn:?. - exploit valnum_regs_holds; eauto. intros [valu' [B [C D]]]. - exploit wf_valnum_regs; eauto. intros [U [V W]]. - set (n2 := kill_equations (filter_after_store chunk addr vargs) n1). - assert (numbering_holds valu' ge sp rs m' n2). - apply kill_equations_holds with (m := m); auto. - intros. destruct r; simpl in *. - rewrite <- H0. apply op_depends_on_memory_correct; auto. - destruct H0 as [a' [P Q]]. - destruct (eq_list_valnum vargs l); simpl in H4; try congruence. subst l. - rewrite negb_false_iff in H4. - exists a'; split; auto. - destruct a; simpl in H2; try congruence. - destruct a'; simpl in Q; try congruence. - simpl. rewrite <- Q. - rewrite C in P. eapply Mem.load_store_other; eauto. - exploit addressing_separated_sound; eauto. intuition congruence. - assert (N2: numbering_satisfiable ge sp rs m' n2). - exists valu'; auto. - set (n3 := add_rhs n2 src (Load chunk addr vargs)). - assert (N3: Val.load_result chunk (rs#src) = rs#src -> numbering_satisfiable ge sp rs m' n3). - intro EQ. unfold n3. apply add_rhs_satisfiable_gen with valu' rs. - apply wf_kill_equations; auto. - red. auto. auto. - red. exists a; split. congruence. - rewrite <- EQ. destruct a; simpl in H2; try discriminate. simpl. - eapply Mem.load_store_same; eauto. - auto. - destruct chunk; auto; apply N3. - simpl in H3. destruct (rs#src); auto || contradiction. - simpl in H3. destruct (rs#src); auto || contradiction. - simpl in H3. destruct (rs#src); auto || contradiction. - simpl in H3. destruct (rs#src); auto || contradiction. +Lemma shift_memcpy_eq_holds: + forall src dst sz e e' m sp bytes m' valu ge, + shift_memcpy_eq src sz (dst - src) e = Some e' -> + Mem.loadbytes m sp src sz = Some bytes -> + Mem.storebytes m sp dst bytes = Some m' -> + equation_holds valu ge (Vptr sp Int.zero) m e -> + equation_holds valu ge (Vptr sp Int.zero) m' e'. +Proof with (try discriminate). + intros. set (delta := dst - src) in *. unfold shift_memcpy_eq in H. + destruct e as [l strict rhs] eqn:E. + destruct rhs as [op vl | chunk addr vl]... + destruct addr... + set (i1 := Int.unsigned i) in *. set (j := i1 + delta) in *. + destruct (zle src i1)... + destruct (zle (i1 + size_chunk chunk) (src + sz))... + destruct (zeq (delta mod align_chunk chunk) 0)... + destruct (zle 0 j)... + destruct (zle j Int.max_unsigned)... + simpl in H; inv H. + assert (LD: forall v, + Mem.loadv chunk m (Vptr sp i) = Some v -> + Mem.loadv chunk m' (Vptr sp (Int.repr j)) = Some v). + { + simpl; intros. rewrite Int.unsigned_repr by omega. + unfold j, delta. eapply load_memcpy; eauto. + apply Zmod_divide; auto. generalize (align_chunk_pos chunk); omega. + } + inv H2. ++ inv H3. destruct vl... simpl in H6. rewrite Int.add_zero_l in H6. inv H6. + apply eq_holds_strict. econstructor. simpl. rewrite Int.add_zero_l. eauto. + apply LD; auto. ++ inv H4. destruct vl... simpl in H7. rewrite Int.add_zero_l in H7. inv H7. + apply eq_holds_lessdef with v; auto. + econstructor. simpl. rewrite Int.add_zero_l. eauto. apply LD; auto. Qed. -(** Correctness of [reg_valnum]: if it returns a register [r], - that register correctly maps back to the given value number. *) - -Lemma reg_valnum_correct: - forall n v r, wf_numbering n -> reg_valnum n v = Some r -> n.(num_reg)!r = Some v. +Lemma add_memcpy_eqs_charact: + forall e' src sz delta eqs2 eqs1, + In e' (add_memcpy_eqs src sz delta eqs1 eqs2) -> + In e' eqs2 \/ exists e, In e eqs1 /\ shift_memcpy_eq src sz delta e = Some e'. Proof. - unfold reg_valnum; intros. inv H. - destruct ((num_val n)#v) as [| r1 rl] eqn:?; inv H0. - eapply VAL. rewrite Heql. auto with coqlib. + induction eqs1; simpl; intros. +- auto. +- destruct (shift_memcpy_eq src sz delta a) as [e''|] eqn:SHIFT. + + destruct H. subst e''. right; exists a; auto. + destruct IHeqs1 as [A | [e [A B]]]; auto. right; exists e; auto. + + destruct IHeqs1 as [A | [e [A B]]]; auto. right; exists e; auto. Qed. -(** Correctness of [find_rhs]: if successful and in a - satisfiable numbering, the returned register does contain the - result value of the operation or memory load. *) - -Lemma find_rhs_correct: - forall valu rs m n rh r, - wf_numbering n -> - numbering_holds valu ge sp rs m n -> - find_rhs n rh = Some r -> - rhs_evals_to valu rh m rs#r. +Lemma add_memcpy_holds: + forall m bsrc osrc sz bytes bdst odst m' valu ge sp rs n1 n2 bc approx ae am rsrc rdst, + Mem.loadbytes m bsrc (Int.unsigned osrc) sz = Some bytes -> + Mem.storebytes m bdst (Int.unsigned odst) bytes = Some m' -> + numbering_holds valu ge (Vptr sp Int.zero) rs m n1 -> + numbering_holds valu ge (Vptr sp Int.zero) rs m' n2 -> + genv_match bc ge -> + bc sp = BCstack -> + ematch bc rs ae -> + approx = VA.State ae am -> + rs#rsrc = Vptr bsrc osrc -> + rs#rdst = Vptr bdst odst -> + Ple (num_next n1) (num_next n2) -> + numbering_holds valu ge (Vptr sp Int.zero) rs m' (add_memcpy approx n1 n2 rsrc rdst sz). Proof. - intros until r. intros WF NH. - unfold find_rhs. - caseEq (find_valnum_rhs rh n.(num_eqs)); intros. - exploit find_valnum_rhs_correct; eauto. intros. - exploit reg_valnum_correct; eauto. intros. - inversion NH. - generalize (H3 _ _ H1). rewrite (H4 _ _ H2). - destruct rh; simpl; auto. - discriminate. + intros. unfold add_memcpy. + destruct (aaddr approx rsrc) eqn:ASRC; auto. + destruct (aaddr approx rdst) eqn:ADST; auto. + assert (A: forall r b o i, + rs#r = Vptr b o -> aaddr approx r = Stk i -> b = sp /\ i = o). + { + intros until i. unfold aaddr; subst approx. intros. + specialize (H5 r). rewrite H6 in H5. rewrite match_aptr_of_aval in H5. + rewrite H10 in H5. inv H5. split; auto. eapply bc_stack; eauto. + } + exploit (A rsrc); eauto. intros [P Q]. + exploit (A rdst); eauto. intros [U V]. + subst bsrc ofs bdst ofs0. + constructor; simpl; intros; eauto with cse. +- constructor; simpl; eauto with cse. + intros. exploit add_memcpy_eqs_charact; eauto. intros [X | (e0 & X & Y)]. + eauto with cse. + apply wf_equation_incr with (num_next n1); auto. + eapply shift_memcpy_eq_wf; eauto with cse. +- exploit add_memcpy_eqs_charact; eauto. intros [X | (e0 & X & Y)]. + eauto with cse. + eapply shift_memcpy_eq_holds; eauto with cse. Qed. (** Correctness of operator reduction *) @@ -740,29 +721,19 @@ Section REDUCE. Variable A: Type. Variable f: (valnum -> option rhs) -> A -> list valnum -> option (A * list valnum). Variable V: Type. +Variable ge: genv. +Variable sp: val. Variable rs: regset. Variable m: mem. Variable sem: A -> list val -> option V. Hypothesis f_sound: forall eqs valu op args op' args', - (forall v rhs, eqs v = Some rhs -> equation_holds valu ge sp m v rhs) -> + (forall v rhs, eqs v = Some rhs -> rhs_eval_to valu ge sp m rhs (valu v)) -> f eqs op args = Some(op', args') -> sem op' (map valu args') = sem op (map valu args). Variable n: numbering. Variable valu: valnum -> val. Hypothesis n_holds: numbering_holds valu ge sp rs m n. -Hypothesis n_wf: wf_numbering n. - -Lemma regs_valnums_correct: - forall vl rl, regs_valnums n vl = Some rl -> rs##rl = map valu vl. -Proof. - induction vl; simpl; intros. - inv H. auto. - destruct (reg_valnum n a) as [r1|] eqn:?; try discriminate. - destruct (regs_valnums n vl) as [rx|] eqn:?; try discriminate. - inv H. simpl; decEq; auto. - eapply (proj2 n_holds); eauto. eapply reg_valnum_correct; eauto. -Qed. Lemma reduce_rec_sound: forall niter op args op' rl' res, @@ -776,11 +747,14 @@ Proof. as [[op1 args1] | ] eqn:?. assert (sem op1 (map valu args1) = Some res). rewrite <- H0. eapply f_sound; eauto. - simpl; intros. apply (proj1 n_holds). eapply find_valnum_num_correct; eauto. + simpl; intros. + exploit num_holds_eq; eauto. + eapply find_valnum_num_charact; eauto with cse. + intros EH; inv EH; auto. destruct (reduce_rec A f n niter op1 args1) as [[op2 rl2] | ] eqn:?. inv H. eapply IHniter; eauto. destruct (regs_valnums n args1) as [rl|] eqn:?. - inv H. erewrite regs_valnums_correct; eauto. + inv H. erewrite regs_valnums_sound; eauto. discriminate. discriminate. Qed. @@ -800,8 +774,6 @@ Qed. End REDUCE. -End SATISFIABILITY. - (** The numberings associated to each instruction by the static analysis are inductively satisfiable, in the following sense: the numbering at the function entry point is satisfiable, and for any RTL execution @@ -809,26 +781,26 @@ End SATISFIABILITY. satisfiability at [pc']. *) Theorem analysis_correct_1: - forall ge sp rs m f approx pc pc' i, - analyze f = Some approx -> + forall ge sp rs m f vapprox approx pc pc' i, + analyze f vapprox = Some approx -> f.(fn_code)!pc = Some i -> In pc' (successors_instr i) -> - numbering_satisfiable ge sp rs m (transfer f pc approx!!pc) -> - numbering_satisfiable ge sp rs m approx!!pc'. + (exists valu, numbering_holds valu ge sp rs m (transfer f vapprox pc approx!!pc)) -> + (exists valu, numbering_holds valu ge sp rs m approx!!pc'). Proof. intros. - assert (Numbering.ge approx!!pc' (transfer f pc approx!!pc)). + assert (Numbering.ge approx!!pc' (transfer f vapprox pc approx!!pc)). eapply Solver.fixpoint_solution; eauto. - apply H3. auto. + destruct H2 as [valu NH]. exists valu; apply H3. auto. Qed. Theorem analysis_correct_entry: - forall ge sp rs m f approx, - analyze f = Some approx -> - numbering_satisfiable ge sp rs m approx!!(f.(fn_entrypoint)). + forall ge sp rs m f vapprox approx, + analyze f vapprox = Some approx -> + exists valu, numbering_holds valu ge sp rs m approx!!(f.(fn_entrypoint)). Proof. intros. replace (approx!!(f.(fn_entrypoint))) with Solver.L.top. - apply empty_numbering_satisfiable. + exists (fun v => Vundef). apply empty_numbering_holds. symmetry. eapply Solver.fixpoint_entry; eauto. Qed. @@ -841,45 +813,34 @@ Variable tprog : program. Hypothesis TRANSF: transf_program prog = OK tprog. Let ge := Genv.globalenv prog. Let tge := Genv.globalenv tprog. +Let rm := romem_for_program prog. Lemma symbols_preserved: forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s. -Proof (Genv.find_symbol_transf_partial transf_fundef prog TRANSF). +Proof (Genv.find_symbol_transf_partial (transf_fundef rm) prog TRANSF). Lemma varinfo_preserved: forall b, Genv.find_var_info tge b = Genv.find_var_info ge b. -Proof (Genv.find_var_info_transf_partial transf_fundef prog TRANSF). +Proof (Genv.find_var_info_transf_partial (transf_fundef rm) prog TRANSF). Lemma functions_translated: forall (v: val) (f: RTL.fundef), Genv.find_funct ge v = Some f -> - exists tf, Genv.find_funct tge v = Some tf /\ transf_fundef f = OK tf. -Proof (Genv.find_funct_transf_partial transf_fundef prog TRANSF). + exists tf, Genv.find_funct tge v = Some tf /\ transf_fundef rm f = OK tf. +Proof (Genv.find_funct_transf_partial (transf_fundef rm) prog TRANSF). Lemma funct_ptr_translated: forall (b: block) (f: RTL.fundef), Genv.find_funct_ptr ge b = Some f -> - exists tf, Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf. -Proof (Genv.find_funct_ptr_transf_partial transf_fundef prog TRANSF). + exists tf, Genv.find_funct_ptr tge b = Some tf /\ transf_fundef rm f = OK tf. +Proof (Genv.find_funct_ptr_transf_partial (transf_fundef rm) prog TRANSF). Lemma sig_preserved: - forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f. + forall f tf, transf_fundef rm f = OK tf -> funsig tf = funsig f. Proof. unfold transf_fundef; intros. destruct f; monadInv H; auto. - unfold transf_function in EQ. destruct (type_function f); try discriminate. - destruct (analyze f); try discriminate. inv EQ; auto. -Qed. - -Lemma find_function_translated: - forall ros rs f, - find_function ge ros rs = Some f -> - exists tf, find_function tge ros rs = Some tf /\ transf_fundef f = OK tf. -Proof. - intros until f; destruct ros; simpl. - intro. apply functions_translated; auto. - rewrite symbols_preserved. destruct (Genv.find_symbol ge i); intro. - apply funct_ptr_translated; auto. - discriminate. + unfold transf_function in EQ. + destruct (analyze f (vanalyze rm f)); try discriminate. inv EQ; auto. Qed. Definition transf_function' (f: function) (approxs: PMap.t numbering) : function := @@ -890,6 +851,50 @@ Definition transf_function' (f: function) (approxs: PMap.t numbering) : function (transf_code approxs f.(fn_code)) f.(fn_entrypoint). +Definition regs_lessdef (rs1 rs2: regset) : Prop := + forall r, Val.lessdef (rs1#r) (rs2#r). + +Lemma regs_lessdef_regs: + forall rs1 rs2, regs_lessdef rs1 rs2 -> + forall rl, Val.lessdef_list rs1##rl rs2##rl. +Proof. + induction rl; constructor; auto. +Qed. + +Lemma set_reg_lessdef: + forall r v1 v2 rs1 rs2, + Val.lessdef v1 v2 -> regs_lessdef rs1 rs2 -> regs_lessdef (rs1#r <- v1) (rs2#r <- v2). +Proof. + intros; red; intros. repeat rewrite Regmap.gsspec. + destruct (peq r0 r); auto. +Qed. + +Lemma init_regs_lessdef: + forall rl vl1 vl2, + Val.lessdef_list vl1 vl2 -> + regs_lessdef (init_regs vl1 rl) (init_regs vl2 rl). +Proof. + induction rl; simpl; intros. + red; intros. rewrite Regmap.gi. auto. + inv H. red; intros. rewrite Regmap.gi. auto. + apply set_reg_lessdef; auto. +Qed. + +Lemma find_function_translated: + forall ros rs fd rs', + find_function ge ros rs = Some fd -> + regs_lessdef rs rs' -> + exists tfd, find_function tge ros rs' = Some tfd /\ transf_fundef rm fd = OK tfd. +Proof. + unfold find_function; intros; destruct ros. +- specialize (H0 r). inv H0. + apply functions_translated; auto. + rewrite <- H2 in H; discriminate. +- rewrite symbols_preserved. destruct (Genv.find_symbol ge i). + apply funct_ptr_translated; auto. + discriminate. +Qed. + (** The proof of semantic preservation is a simulation argument using diagrams of the following form: << @@ -906,45 +911,44 @@ Definition transf_function' (f: function) (approxs: PMap.t numbering) : function the numbering at [pc] (returned by the static analysis) is satisfiable. *) -Inductive match_stackframes: list stackframe -> list stackframe -> typ -> Prop := +Inductive match_stackframes: list stackframe -> list stackframe -> Prop := | match_stackframes_nil: - match_stackframes nil nil Tint + match_stackframes nil nil | match_stackframes_cons: - forall res sp pc rs f approx tyenv s s' ty - (ANALYZE: analyze f = Some approx) - (WTF: wt_function f tyenv) - (WTREGS: wt_regset tyenv rs) - (TYRES: subtype ty (tyenv res) = true) - (SAT: forall v m, numbering_satisfiable ge sp (rs#res <- v) m approx!!pc) - (STACKS: match_stackframes s s' (proj_sig_res (fn_sig f))), + forall res sp pc rs f approx s rs' s' + (ANALYZE: analyze f (vanalyze rm f) = Some approx) + (SAT: forall v m, exists valu, numbering_holds valu ge sp (rs#res <- v) m approx!!pc) + (RLD: regs_lessdef rs rs') + (STACKS: match_stackframes s s'), match_stackframes (Stackframe res f sp pc rs :: s) - (Stackframe res (transf_function' f approx) sp pc rs :: s') - ty. + (Stackframe res (transf_function' f approx) sp pc rs' :: s'). Inductive match_states: state -> state -> Prop := | match_states_intro: - forall s sp pc rs m s' f approx tyenv - (ANALYZE: analyze f = Some approx) - (WTF: wt_function f tyenv) - (WTREGS: wt_regset tyenv rs) - (SAT: numbering_satisfiable ge sp rs m approx!!pc) - (STACKS: match_stackframes s s' (proj_sig_res (fn_sig f))), + forall s sp pc rs m s' rs' m' f approx + (ANALYZE: analyze f (vanalyze rm f) = Some approx) + (SAT: exists valu, numbering_holds valu ge sp rs m approx!!pc) + (RLD: regs_lessdef rs rs') + (MEXT: Mem.extends m m') + (STACKS: match_stackframes s s'), match_states (State s f sp pc rs m) - (State s' (transf_function' f approx) sp pc rs m) + (State s' (transf_function' f approx) sp pc rs' m') | match_states_call: - forall s f tf args m s', - match_stackframes s s' (proj_sig_res (funsig f)) -> - Val.has_type_list args (sig_args (funsig f)) -> - transf_fundef f = OK tf -> + forall s f tf args m s' args' m', + match_stackframes s s' -> + transf_fundef rm f = OK tf -> + Val.lessdef_list args args' -> + Mem.extends m m' -> match_states (Callstate s f args m) - (Callstate s' tf args m) + (Callstate s' tf args' m') | match_states_return: - forall s s' ty v m, - match_stackframes s s' ty -> - Val.has_type v ty -> + forall s s' v v' m m', + match_stackframes s s' -> + Val.lessdef v v' -> + Mem.extends m m' -> match_states (Returnstate s v m) - (Returnstate s' v m). + (Returnstate s' v' m'). Ltac TransfInstr := match goal with @@ -960,196 +964,241 @@ Ltac TransfInstr := Lemma transf_step_correct: forall s1 t s2, step ge s1 t s2 -> - forall s1' (MS: match_states s1 s1'), + forall s1' (MS: match_states s1 s1') (SOUND: sound_state prog s1), exists s2', step tge s1' t s2' /\ match_states s2 s2'. Proof. induction 1; intros; inv MS; try (TransfInstr; intro C). (* Inop *) - exists (State s' (transf_function' f approx) sp pc' rs m); split. - apply exec_Inop; auto. +- econstructor; split. + eapply exec_Inop; eauto. econstructor; eauto. eapply analysis_correct_1; eauto. simpl; auto. unfold transfer; rewrite H; auto. (* Iop *) - exists (State s' (transf_function' f approx) sp pc' (rs#res <- v) m); split. - destruct (is_trivial_op op) eqn:?. - eapply exec_Iop'; eauto. - rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved. +- destruct (is_trivial_op op) eqn:TRIV. ++ (* unchanged *) + exploit eval_operation_lessdef. eapply regs_lessdef_regs; eauto. eauto. eauto. + intros [v' [A B]]. + econstructor; split. + eapply exec_Iop with (v := v'); eauto. + rewrite <- A. apply eval_operation_preserved. exact symbols_preserved. + econstructor; eauto. + eapply analysis_correct_1; eauto. simpl; auto. + unfold transfer; rewrite H. + destruct SAT as [valu NH]. eapply add_op_holds; eauto. + apply set_reg_lessdef; auto. ++ (* possibly optimized *) destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?. - assert (wf_numbering approx!!pc). eapply wf_analyze; eauto. destruct SAT as [valu1 NH1]. - exploit valnum_regs_holds; eauto. intros [valu2 [NH2 [EQ AG]]]. - assert (wf_numbering n1). eapply wf_valnum_regs; eauto. + exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q). destruct (find_rhs n1 (Op op vl)) as [r|] eqn:?. - (* replaced by move *) - assert (EV: rhs_evals_to ge sp valu2 (Op op vl) m rs#r). eapply find_rhs_correct; eauto. - assert (RES: rs#r = v). red in EV. congruence. - eapply exec_Iop'; eauto. simpl. congruence. - (* possibly simplified *) +* (* replaced by move *) + exploit find_rhs_sound; eauto. intros (v' & EV & LD). + assert (v' = v) by (inv EV; congruence). subst v'. + econstructor; split. + eapply exec_Iop; eauto. simpl; eauto. + econstructor; eauto. + eapply analysis_correct_1; eauto. simpl; auto. + unfold transfer; rewrite H. + eapply add_op_holds; eauto. + apply set_reg_lessdef; auto. + eapply Val.lessdef_trans; eauto. +* (* possibly simplified *) destruct (reduce operation combine_op n1 op args vl) as [op' args'] eqn:?. assert (RES: eval_operation ge sp op' rs##args' m = Some v). eapply reduce_sound with (sem := fun op vl => eval_operation ge sp op vl m); eauto. intros; eapply combine_op_sound; eauto. - eapply exec_Iop'; eauto. - rewrite <- RES. apply eval_operation_preserved. exact symbols_preserved. - (* state matching *) + exploit eval_operation_lessdef. eapply regs_lessdef_regs; eauto. eauto. eauto. + intros [v' [A B]]. + econstructor; split. + eapply exec_Iop with (v := v'); eauto. + rewrite <- A. apply eval_operation_preserved. exact symbols_preserved. econstructor; eauto. - eapply wt_exec_Iop; eauto. eapply wt_instr_at; eauto. eapply analysis_correct_1; eauto. simpl; auto. unfold transfer; rewrite H. - eapply add_op_satisfiable; eauto. eapply wf_analyze; eauto. + eapply add_op_holds; eauto. + apply set_reg_lessdef; auto. - (* Iload *) - exists (State s' (transf_function' f approx) sp pc' (rs#dst <- v) m); split. +- (* Iload *) destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?. - assert (wf_numbering approx!!pc). eapply wf_analyze; eauto. destruct SAT as [valu1 NH1]. - exploit valnum_regs_holds; eauto. intros [valu2 [NH2 [EQ AG]]]. - assert (wf_numbering n1). eapply wf_valnum_regs; eauto. + exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q). destruct (find_rhs n1 (Load chunk addr vl)) as [r|] eqn:?. - (* replaced by move *) - assert (EV: rhs_evals_to ge sp valu2 (Load chunk addr vl) m rs#r). eapply find_rhs_correct; eauto. - assert (RES: rs#r = v). red in EV. destruct EV as [a' [EQ1 EQ2]]. congruence. - eapply exec_Iop'; eauto. simpl. congruence. - (* possibly simplified *) ++ (* replaced by move *) + exploit find_rhs_sound; eauto. intros (v' & EV & LD). + assert (v' = v) by (inv EV; congruence). subst v'. + econstructor; split. + eapply exec_Iop; eauto. simpl; eauto. + econstructor; eauto. + eapply analysis_correct_1; eauto. simpl; auto. + unfold transfer; rewrite H. + eapply add_load_holds; eauto. + apply set_reg_lessdef; auto. eapply Val.lessdef_trans; eauto. ++ (* load is preserved, but addressing is possibly simplified *) destruct (reduce addressing combine_addr n1 addr args vl) as [addr' args'] eqn:?. assert (ADDR: eval_addressing ge sp addr' rs##args' = Some a). - eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto. - intros; eapply combine_addr_sound; eauto. - assert (ADDR': eval_addressing tge sp addr' rs##args' = Some a). - rewrite <- ADDR. apply eval_addressing_preserved. exact symbols_preserved. + { eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto. + intros; eapply combine_addr_sound; eauto. } + exploit eval_addressing_lessdef. apply regs_lessdef_regs; eauto. eexact ADDR. + intros [a' [A B]]. + assert (ADDR': eval_addressing tge sp addr' rs'##args' = Some a'). + { rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved. } + exploit Mem.loadv_extends; eauto. + intros [v' [X Y]]. + econstructor; split. eapply exec_Iload; eauto. - (* state matching *) econstructor; eauto. - eapply wt_exec_Iload; eauto. eapply wt_instr_at; eauto. eapply analysis_correct_1; eauto. simpl; auto. unfold transfer; rewrite H. - eapply add_load_satisfiable; eauto. eapply wf_analyze; eauto. + eapply add_load_holds; eauto. + apply set_reg_lessdef; auto. - (* Istore *) - exists (State s' (transf_function' f approx) sp pc' rs m'); split. +- (* Istore *) destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?. - assert (wf_numbering approx!!pc). eapply wf_analyze; eauto. destruct SAT as [valu1 NH1]. - exploit valnum_regs_holds; eauto. intros [valu2 [NH2 [EQ AG]]]. - assert (wf_numbering n1). eapply wf_valnum_regs; eauto. + exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q). destruct (reduce addressing combine_addr n1 addr args vl) as [addr' args'] eqn:?. assert (ADDR: eval_addressing ge sp addr' rs##args' = Some a). - eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto. - intros; eapply combine_addr_sound; eauto. - assert (ADDR': eval_addressing tge sp addr' rs##args' = Some a). - rewrite <- ADDR. apply eval_addressing_preserved. exact symbols_preserved. + { eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto. + intros; eapply combine_addr_sound; eauto. } + exploit eval_addressing_lessdef. apply regs_lessdef_regs; eauto. eexact ADDR. + intros [a' [A B]]. + assert (ADDR': eval_addressing tge sp addr' rs'##args' = Some a'). + { rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved. } + exploit Mem.storev_extends; eauto. intros [m'' [X Y]]. + econstructor; split. eapply exec_Istore; eauto. econstructor; eauto. eapply analysis_correct_1; eauto. simpl; auto. unfold transfer; rewrite H. - eapply add_store_satisfiable; eauto. eapply wf_analyze; eauto. - exploit wt_instr_at; eauto. intros WTI; inv WTI. - eapply Val.has_subtype; eauto. + inv SOUND. + eapply add_store_result_hold; eauto. + eapply kill_loads_after_store_holds; eauto. - (* Icall *) +- (* Icall *) exploit find_function_translated; eauto. intros [tf [FIND' TRANSF']]. econstructor; split. eapply exec_Icall; eauto. apply sig_preserved; auto. - exploit wt_instr_at; eauto. intros WTI; inv WTI. econstructor; eauto. econstructor; eauto. intros. eapply analysis_correct_1; eauto. simpl; auto. unfold transfer; rewrite H. - apply empty_numbering_satisfiable. - eapply Val.has_subtype_list; eauto. apply wt_regset_list; auto. + exists (fun _ => Vundef); apply empty_numbering_holds. + apply regs_lessdef_regs; auto. - (* Itailcall *) +- (* Itailcall *) exploit find_function_translated; eauto. intros [tf [FIND' TRANSF']]. + exploit Mem.free_parallel_extends; eauto. intros [m'' [A B]]. econstructor; split. eapply exec_Itailcall; eauto. apply sig_preserved; auto. - exploit wt_instr_at; eauto. intros WTI; inv WTI. econstructor; eauto. - replace (proj_sig_res (funsig fd)) with (proj_sig_res (fn_sig f)). auto. - unfold proj_sig_res. rewrite H6; auto. - eapply Val.has_subtype_list; eauto. apply wt_regset_list; auto. + apply regs_lessdef_regs; auto. - (* Ibuiltin *) +- (* Ibuiltin *) + exploit external_call_mem_extends; eauto. + instantiate (1 := rs'##args). apply regs_lessdef_regs; auto. + intros (v' & m1' & P & Q & R & S). econstructor; split. eapply exec_Ibuiltin; eauto. eapply external_call_symbols_preserved; eauto. exact symbols_preserved. exact varinfo_preserved. econstructor; eauto. - eapply wt_exec_Ibuiltin; eauto. eapply wt_instr_at; eauto. eapply analysis_correct_1; eauto. simpl; auto. - unfold transfer; rewrite H. - assert (CASE1: numbering_satisfiable ge sp (rs#res <- v) m' empty_numbering). - { apply empty_numbering_satisfiable. } - assert (CASE2: m' = m -> numbering_satisfiable ge sp (rs#res <- v) m' (add_unknown approx#pc res)). - { intros. rewrite H1. apply add_unknown_satisfiable. - eapply wf_analyze; eauto. auto. } - assert (CASE3: numbering_satisfiable ge sp (rs#res <- v) m' - (add_unknown (kill_loads approx#pc) res)). - { apply add_unknown_satisfiable. apply wf_kill_equations. eapply wf_analyze; eauto. - eapply kill_loads_satisfiable; eauto. } - destruct ef; auto; apply CASE2; inv H0; auto. - - (* Icond *) +* unfold transfer; rewrite H. + destruct SAT as [valu NH]. + assert (CASE1: exists valu, numbering_holds valu ge sp (rs#res <- v) m' empty_numbering). + { exists valu; apply empty_numbering_holds. } + assert (CASE2: m' = m -> exists valu, numbering_holds valu ge sp (rs#res <- v) m' (set_unknown approx#pc res)). + { intros. rewrite H1. exists valu. apply set_unknown_holds; auto. } + assert (CASE3: exists valu, numbering_holds valu ge sp (rs#res <- v) m' + (set_unknown (kill_all_loads approx#pc) res)). + { exists valu. apply set_unknown_holds. eapply kill_all_loads_hold; eauto. } + destruct ef. + + apply CASE1. + + apply CASE3. + + apply CASE2; inv H0; auto. + + apply CASE3. + + apply CASE2; inv H0; auto. + + apply CASE3; auto. + + apply CASE1. + + apply CASE1. + + destruct args as [ | rdst args]; auto. + destruct args as [ | rsrc args]; auto. + destruct args; auto. + simpl in H0. inv H0. + exists valu. + apply set_unknown_holds. + inv SOUND. eapply add_memcpy_holds; eauto. + eapply kill_loads_after_storebytes_holds; eauto. + eapply Mem.loadbytes_length; eauto. + omega. + simpl. apply Ple_refl. + + apply CASE2; inv H0; auto. + + apply CASE2; inv H0; auto. + + apply CASE1. +* apply set_reg_lessdef; auto. + +- (* Icond *) destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?. - assert (wf_numbering approx!!pc). eapply wf_analyze; eauto. elim SAT; intros valu1 NH1. - exploit valnum_regs_holds; eauto. intros [valu2 [NH2 [EQ AG]]]. - assert (wf_numbering n1). eapply wf_valnum_regs; eauto. + exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q). destruct (reduce condition combine_cond n1 cond args vl) as [cond' args'] eqn:?. assert (RES: eval_condition cond' rs##args' m = Some b). - eapply reduce_sound with (sem := fun cond vl => eval_condition cond vl m); eauto. - intros; eapply combine_cond_sound; eauto. + { eapply reduce_sound with (sem := fun cond vl => eval_condition cond vl m); eauto. + intros; eapply combine_cond_sound; eauto. } econstructor; split. eapply exec_Icond; eauto. + eapply eval_condition_lessdef; eauto. apply regs_lessdef_regs; auto. econstructor; eauto. destruct b; eapply analysis_correct_1; eauto; simpl; auto; unfold transfer; rewrite H; auto. - (* Ijumptable *) +- (* Ijumptable *) + generalize (RLD arg); rewrite H0; intro LD; inv LD. econstructor; split. eapply exec_Ijumptable; eauto. econstructor; eauto. eapply analysis_correct_1; eauto. simpl. eapply list_nth_z_in; eauto. unfold transfer; rewrite H; auto. - (* Ireturn *) +- (* Ireturn *) + exploit Mem.free_parallel_extends; eauto. intros [m'' [A B]]. econstructor; split. eapply exec_Ireturn; eauto. econstructor; eauto. - exploit wt_instr_at; eauto. intros WTI; inv WTI; simpl. - auto. - unfold proj_sig_res; rewrite H2. eapply Val.has_subtype; eauto. + destruct or; simpl; auto. - (* internal function *) - monadInv H7. unfold transf_function in EQ. - destruct (type_function f) as [tyenv|] eqn:?; try discriminate. - destruct (analyze f) as [approx|] eqn:?; inv EQ. - assert (WTF: wt_function f tyenv). apply type_function_correct; auto. +- (* internal function *) + monadInv H6. unfold transf_function in EQ. + destruct (analyze f (vanalyze rm f)) as [approx|] eqn:?; inv EQ. + exploit Mem.alloc_extends; eauto. apply Zle_refl. apply Zle_refl. + intros (m'' & A & B). econstructor; split. - eapply exec_function_internal; eauto. + eapply exec_function_internal; simpl; eauto. simpl. econstructor; eauto. - apply wt_init_regs. inv WTF. eapply Val.has_subtype_list; eauto. - apply analysis_correct_entry; auto. + eapply analysis_correct_entry; eauto. + apply init_regs_lessdef; auto. - (* external function *) - monadInv H7. +- (* external function *) + monadInv H6. + exploit external_call_mem_extends; eauto. + intros (v' & m1' & P & Q & R & S). econstructor; split. eapply exec_function_external; eauto. eapply external_call_symbols_preserved; eauto. exact symbols_preserved. exact varinfo_preserved. econstructor; eauto. - simpl. eapply external_call_well_typed; eauto. - (* return *) - inv H3. +- (* return *) + inv H2. econstructor; split. eapply exec_return; eauto. econstructor; eauto. - apply wt_regset_assign; auto. eapply Val.has_subtype; eauto. + apply set_reg_lessdef; auto. Qed. Lemma transf_initial_states: @@ -1165,24 +1214,27 @@ Proof. rewrite symbols_preserved. eauto. symmetry. eapply transform_partial_program_main; eauto. rewrite <- H3. apply sig_preserved; auto. - constructor. rewrite H3. constructor. rewrite H3. constructor. auto. + constructor. constructor. auto. auto. apply Mem.extends_refl. Qed. Lemma transf_final_states: forall st1 st2 r, match_states st1 st2 -> final_state st1 r -> final_state st2 r. Proof. - intros. inv H0. inv H. inv H4. constructor. + intros. inv H0. inv H. inv H5. inv H3. constructor. Qed. Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (RTL.semantics tprog). Proof. - eapply forward_simulation_step. - eexact symbols_preserved. - eexact transf_initial_states. - eexact transf_final_states. - exact transf_step_correct. + eapply forward_simulation_step with + (match_states := fun s1 s2 => sound_state prog s1 /\ match_states s1 s2). +- eexact symbols_preserved. +- intros. exploit transf_initial_states; eauto. intros [s2 [A B]]. + exists s2. split. auto. split. apply sound_initial; auto. auto. +- intros. destruct H. eapply transf_final_states; eauto. +- intros. destruct H0. exploit transf_step_correct; eauto. + intros [s2' [A B]]. exists s2'; split. auto. split. eapply sound_step; eauto. auto. Qed. End PRESERVATION. diff --git a/backend/Constprop.v b/backend/Constprop.v index e5ea64d..76d8e6c 100644 --- a/backend/Constprop.v +++ b/backend/Constprop.v @@ -25,242 +25,30 @@ Require Import RTL. Require Import Lattice. Require Import Kildall. Require Import Liveness. +Require Import ValueDomain. +Require Import ValueAOp. +Require Import ValueAnalysis. Require Import ConstpropOp. -(** * Static analysis *) - -(** The type [approx] of compile-time approximations of values is - defined in the machine-dependent part [ConstpropOp]. *) - -(** We equip this type of approximations with a semi-lattice structure. - The ordering is inclusion between the sets of values denoted by - the approximations. *) - -Module Approx <: SEMILATTICE_WITH_TOP. - Definition t := approx. - Definition eq (x y: t) := (x = y). - Definition eq_refl: forall x, eq x x := (@refl_equal t). - Definition eq_sym: forall x y, eq x y -> eq y x := (@sym_equal t). - Definition eq_trans: forall x y z, eq x y -> eq y z -> eq x z := (@trans_equal t). - Lemma eq_dec: forall (x y: t), {x=y} + {x<>y}. - Proof. - decide equality. - apply Int.eq_dec. - apply Float.eq_dec. - apply Int64.eq_dec. - apply Int.eq_dec. - apply ident_eq. - apply Int.eq_dec. - Defined. - Definition beq (x y: t) := if eq_dec x y then true else false. - Lemma beq_correct: forall x y, beq x y = true -> x = y. - Proof. - unfold beq; intros. destruct (eq_dec x y). auto. congruence. - Qed. - - Definition ge (x y: t) : Prop := x = Unknown \/ y = Novalue \/ x = y. - - Lemma ge_refl: forall x y, eq x y -> ge x y. - Proof. - unfold eq, ge; tauto. - Qed. - Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z. - Proof. - unfold ge; intuition congruence. - Qed. - Lemma ge_compat: forall x x' y y', eq x x' -> eq y y' -> ge x y -> ge x' y'. - Proof. - unfold eq, ge; intros; congruence. - Qed. - Definition bot := Novalue. - Definition top := Unknown. - Lemma ge_bot: forall x, ge x bot. - Proof. - unfold ge, bot; tauto. - Qed. - Lemma ge_top: forall x, ge top x. - Proof. - unfold ge, bot; tauto. - Qed. - Definition lub (x y: t) : t := - if eq_dec x y then x else - match x, y with - | Novalue, _ => y - | _, Novalue => x - | _, _ => Unknown - end. - Lemma ge_lub_left: forall x y, ge (lub x y) x. - Proof. - unfold lub; intros. - case (eq_dec x y); intro. - apply ge_refl. apply eq_refl. - destruct x; destruct y; unfold ge; tauto. - Qed. - Lemma ge_lub_right: forall x y, ge (lub x y) y. - Proof. - unfold lub; intros. - case (eq_dec x y); intro. - apply ge_refl. subst. apply eq_refl. - destruct x; destruct y; unfold ge; tauto. - Qed. -End Approx. - -Module D := LPMap Approx. - -(** We keep track of read-only global variables (i.e. "const" global - variables in C) as a map from their names to their initialization - data. *) - -Definition global_approx : Type := PTree.t (list init_data). - -(** Given some initialization data and a byte offset, compute a static - approximation of the result of a memory load from a memory block - initialized with this data. *) - -Fixpoint eval_load_init (chunk: memory_chunk) (pos: Z) (il: list init_data): approx := - match il with - | nil => Unknown - | Init_int8 n :: il' => - if zeq pos 0 then - match chunk with - | Mint8unsigned => I (Int.zero_ext 8 n) - | Mint8signed => I (Int.sign_ext 8 n) - | _ => Unknown - end - else eval_load_init chunk (pos - 1) il' - | Init_int16 n :: il' => - if zeq pos 0 then - match chunk with - | Mint16unsigned => I (Int.zero_ext 16 n) - | Mint16signed => I (Int.sign_ext 16 n) - | _ => Unknown - end - else eval_load_init chunk (pos - 2) il' - | Init_int32 n :: il' => - if zeq pos 0 - then match chunk with Mint32 => I n | _ => Unknown end - else eval_load_init chunk (pos - 4) il' - | Init_int64 n :: il' => - if zeq pos 0 - then match chunk with Mint64 => L n | _ => Unknown end - else eval_load_init chunk (pos - 8) il' - | Init_float32 n :: il' => - if zeq pos 0 - then match chunk with - | Mfloat32 => if propagate_float_constants tt then F (Float.singleoffloat n) else Unknown - | _ => Unknown - end - else eval_load_init chunk (pos - 4) il' - | Init_float64 n :: il' => - if zeq pos 0 - then match chunk with - | Mfloat64 => if propagate_float_constants tt then F n else Unknown - | _ => Unknown - end - else eval_load_init chunk (pos - 8) il' - | Init_addrof symb ofs :: il' => - if zeq pos 0 - then match chunk with Mint32 => G symb ofs | _ => Unknown end - else eval_load_init chunk (pos - 4) il' - | Init_space n :: il' => - eval_load_init chunk (pos - Zmax n 0) il' - end. - -(** Compute a static approximation for the result of a load at an address whose - approximation is known. If the approximation points to a global variable, - and this global variable is read-only, we use its initialization data - to determine a static approximation. Otherwise, [Unknown] is returned. *) - -Definition eval_static_load (gapp: global_approx) (chunk: memory_chunk) (addr: approx) : approx := - match addr with - | G symb ofs => - match gapp!symb with - | None => Unknown - | Some il => eval_load_init chunk (Int.unsigned ofs) il - end - | _ => Unknown - end. - -(** The transfer function for the dataflow analysis is straightforward. - For [Iop] instructions, we set the approximation of the destination - register to the result of executing abstractly the operation. - For [Iload] instructions, we set the approximation of the destination - register to the result of [eval_static_load]. - For [Icall] and [Ibuiltin], the destination register becomes [Unknown]. - Other instructions keep the approximations unchanged, as they preserve - the values of all registers. *) - -Definition approx_reg (app: D.t) (r: reg) := - D.get r app. - -Definition approx_regs (app: D.t) (rl: list reg):= - List.map (approx_reg app) rl. - -Definition transfer (gapp: global_approx) (f: function) (pc: node) (before: D.t) := - match f.(fn_code)!pc with - | None => before - | Some i => - match i with - | Iop op args res s => - let a := eval_static_operation op (approx_regs before args) in - D.set res a before - | Iload chunk addr args dst s => - let a := eval_static_load gapp chunk - (eval_static_addressing addr (approx_regs before args)) in - D.set dst a before - | Icall sig ros args res s => - D.set res Unknown before - | Ibuiltin ef args res s => - D.set res Unknown before - | _ => - before - end - end. - -(** To reduce the size of approximations, we preventively set to [Top] - the approximations of registers used for the last time in the - current instruction. *) - -Definition transfer' (gapp: global_approx) (f: function) (lastuses: PTree.t (list reg)) - (pc: node) (before: D.t) := - let after := transfer gapp f pc before in - match lastuses!pc with - | None => after - | Some regs => List.fold_left (fun a r => D.set r Unknown a) regs after - end. - -(** The static analysis itself is then an instantiation of Kildall's - generic solver for forward dataflow inequations. [analyze f] - returns a mapping from program points to mappings of pseudo-registers - to approximations. It can fail to reach a fixpoint in a reasonable - number of iterations, in which case we use the trivial mapping - (program point -> [D.top]) instead. *) - -Module DS := Dataflow_Solver(D)(NodeSetForward). - -Definition analyze (gapp: global_approx) (f: RTL.function): PMap.t D.t := - let lu := Liveness.last_uses f in - match DS.fixpoint f.(fn_code) successors_instr (transfer' gapp f lu) - ((f.(fn_entrypoint), D.top) :: nil) with - | None => PMap.init D.top - | Some res => res - end. - -(** * Code transformation *) - -(** The code transformation proceeds instruction by instruction. - Operators whose arguments are all statically known are turned - into ``load integer constant'', ``load float constant'' or - ``load symbol address'' operations. Likewise for loads whose - result can be statically predicted. Operators for which some - but not all arguments are known are subject to strength reduction, - and similarly for the addressing modes of load and store instructions. - Conditional branches and multi-way branches are statically resolved - into [Inop] instructions if possible. Other instructions are unchanged. - - In addition, we try to jump over conditionals whose condition can - be statically resolved based on the abstract state "after" the - instruction that branches to the conditional. A typical example is: +(** The code transformation builds on the results of the static analysis + of values from module [ValueAnalysis]. It proceeds instruction by + instruction. +- Operators whose arguments are all statically known are turned into + ``load integer constant'', ``load float constant'' or ``load + symbol address'' operations. Likewise for loads whose result can + be statically predicted. +- Operators for which some but not all arguments are known are subject + to strength reduction (replacement by cheaper operators) and + similarly for the addressing modes of load and store instructions. +- Cast operators that have no effect (because their arguments are + already normalized to the destination type) are removed. +- Conditional branches and multi-way branches are statically resolved + into [Inop] instructions when possible. +- Other instructions are unchanged. + + In addition, we try to jump over conditionals whose condition can + be statically resolved based on the abstract state "after" the + instruction that branches to the conditional. A typical example is: << 1: x := 0 and goto 2 2: if (x == 0) goto 3 else goto 4 @@ -273,37 +61,35 @@ Definition analyze (gapp: global_approx) (f: RTL.function): PMap.t D.t := >> *) -Definition transf_ros (app: D.t) (ros: reg + ident) : reg + ident := +Definition transf_ros (ae: AE.t) (ros: reg + ident) : reg + ident := match ros with | inl r => - match D.get r app with - | G symb ofs => if Int.eq ofs Int.zero then inr _ symb else ros + match areg ae r with + | Ptr(Gl symb ofs) => if Int.eq ofs Int.zero then inr _ symb else ros | _ => ros end | inr s => ros end. -Parameter generate_float_constants : unit -> bool. - -Definition const_for_result (a: approx) : option operation := +Definition const_for_result (a: aval) : option operation := match a with | I n => Some(Ointconst n) | F n => if generate_float_constants tt then Some(Ofloatconst n) else None - | G symb ofs => Some(Oaddrsymbol symb ofs) - | S ofs => Some(Oaddrstack ofs) + | Ptr(Gl symb ofs) => Some(Oaddrsymbol symb ofs) + | Ptr(Stk ofs) => Some(Oaddrstack ofs) | _ => None end. -Fixpoint successor_rec (n: nat) (f: function) (app: D.t) (pc: node) : node := +Fixpoint successor_rec (n: nat) (f: function) (ae: AE.t) (pc: node) : node := match n with | O => pc - | Datatypes.S n' => + | S n' => match f.(fn_code)!pc with | Some (Inop s) => - successor_rec n' f app s + successor_rec n' f ae s | Some (Icond cond args s1 s2) => - match eval_static_condition cond (approx_regs app args) with - | Some b => if b then s1 else s2 + match resolve_branch (eval_static_condition cond (aregs ae args)) with + | Some b => successor_rec n' f ae (if b then s1 else s2) | None => pc end | _ => pc @@ -312,15 +98,15 @@ Fixpoint successor_rec (n: nat) (f: function) (app: D.t) (pc: node) : node := Definition num_iter := 10%nat. -Definition successor (f: function) (app: D.t) (pc: node) : node := - successor_rec num_iter f app pc. +Definition successor (f: function) (ae: AE.t) (pc: node) : node := + successor_rec num_iter f ae pc. -Function annot_strength_reduction - (app: D.t) (targs: list annot_arg) (args: list reg) := +Fixpoint annot_strength_reduction + (ae: AE.t) (targs: list annot_arg) (args: list reg) := match targs, args with | AA_arg ty :: targs', arg :: args' => - let (targs'', args'') := annot_strength_reduction app targs' args' in - match ty, approx_reg app arg with + let (targs'', args'') := annot_strength_reduction ae targs' args' in + match ty, areg ae arg with | Tint, I n => (AA_int n :: targs'', args'') | Tfloat, F n => if generate_float_constants tt then (AA_float n :: targs'', args'') @@ -328,117 +114,106 @@ Function annot_strength_reduction | _, _ => (AA_arg ty :: targs'', arg :: args'') end | targ :: targs', _ => - let (targs'', args'') := annot_strength_reduction app targs' args in + let (targs'', args'') := annot_strength_reduction ae targs' args in (targ :: targs'', args'') | _, _ => (targs, args) end. Function builtin_strength_reduction - (app: D.t) (ef: external_function) (args: list reg) := + (ae: AE.t) (ef: external_function) (args: list reg) := match ef, args with | EF_vload chunk, r1 :: nil => - match approx_reg app r1 with - | G symb n1 => (EF_vload_global chunk symb n1, nil) + match areg ae r1 with + | Ptr(Gl symb n1) => (EF_vload_global chunk symb n1, nil) | _ => (ef, args) end | EF_vstore chunk, r1 :: r2 :: nil => - match approx_reg app r1 with - | G symb n1 => (EF_vstore_global chunk symb n1, r2 :: nil) + match areg ae r1 with + | Ptr(Gl symb n1) => (EF_vstore_global chunk symb n1, r2 :: nil) | _ => (ef, args) end | EF_annot text targs, args => - let (targs', args') := annot_strength_reduction app targs args in + let (targs', args') := annot_strength_reduction ae targs args in (EF_annot text targs', args') | _, _ => (ef, args) end. -Definition transf_instr (gapp: global_approx) (f: function) (apps: PMap.t D.t) - (pc: node) (instr: instruction) := - let app := apps!!pc in - match instr with - | Iop op args res s => - let a := eval_static_operation op (approx_regs app args) in - let s' := successor f (D.set res a app) s in - match const_for_result a with - | Some cop => - Iop cop nil res s' - | None => - let (op', args') := op_strength_reduction op args (approx_regs app args) in - Iop op' args' res s' - end - | Iload chunk addr args dst s => - let a := eval_static_load gapp chunk - (eval_static_addressing addr (approx_regs app args)) in - match const_for_result a with - | Some cop => - Iop cop nil dst s - | None => - let (addr', args') := addr_strength_reduction addr args (approx_regs app args) in - Iload chunk addr' args' dst s - end - | Istore chunk addr args src s => - let (addr', args') := addr_strength_reduction addr args (approx_regs app args) in - Istore chunk addr' args' src s - | Icall sig ros args res s => - Icall sig (transf_ros app ros) args res s - | Itailcall sig ros args => - Itailcall sig (transf_ros app ros) args - | Ibuiltin ef args res s => - let (ef', args') := builtin_strength_reduction app ef args in - Ibuiltin ef' args' res s - | Icond cond args s1 s2 => - match eval_static_condition cond (approx_regs app args) with - | Some b => - if b then Inop s1 else Inop s2 - | None => - let (cond', args') := cond_strength_reduction cond args (approx_regs app args) in - Icond cond' args' s1 s2 - end - | Ijumptable arg tbl => - match approx_reg app arg with - | I n => - match list_nth_z tbl (Int.unsigned n) with - | Some s => Inop s - | None => instr +Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem) + (pc: node) (instr: instruction) := + match an!!pc with + | VA.Bot => + instr + | VA.State ae am => + match instr with + | Iop op args res s => + let aargs := aregs ae args in + let a := eval_static_operation op aargs in + let s' := successor f (AE.set res a ae) s in + match const_for_result a with + | Some cop => + Iop cop nil res s' + | None => + let (op', args') := op_strength_reduction op args aargs in + Iop op' args' res s' + end + | Iload chunk addr args dst s => + let aargs := aregs ae args in + let a := ValueDomain.loadv chunk rm am (eval_static_addressing addr aargs) in + match const_for_result a with + | Some cop => + Iop cop nil dst s + | None => + let (addr', args') := addr_strength_reduction addr args aargs in + Iload chunk addr' args' dst s + end + | Istore chunk addr args src s => + let aargs := aregs ae args in + let (addr', args') := addr_strength_reduction addr args aargs in + Istore chunk addr' args' src s + | Icall sig ros args res s => + Icall sig (transf_ros ae ros) args res s + | Itailcall sig ros args => + Itailcall sig (transf_ros ae ros) args + | Ibuiltin ef args res s => + let (ef', args') := builtin_strength_reduction ae ef args in + Ibuiltin ef' args' res s + | Icond cond args s1 s2 => + let aargs := aregs ae args in + match resolve_branch (eval_static_condition cond aargs) with + | Some b => + if b then Inop s1 else Inop s2 + | None => + let (cond', args') := cond_strength_reduction cond args aargs in + Icond cond' args' s1 s2 end - | _ => instr + | Ijumptable arg tbl => + match areg ae arg with + | I n => + match list_nth_z tbl (Int.unsigned n) with + | Some s => Inop s + | None => instr + end + | _ => instr + end + | _ => + instr end - | _ => - instr end. -Definition transf_code (gapp: global_approx) (f: function) (app: PMap.t D.t) (instrs: code) : code := - PTree.map (transf_instr gapp f app) instrs. - -Definition transf_function (gapp: global_approx) (f: function) : function := - let approxs := analyze gapp f in +Definition transf_function (rm: romem) (f: function) : function := + let an := ValueAnalysis.analyze rm f in mkfunction f.(fn_sig) f.(fn_params) f.(fn_stacksize) - (transf_code gapp f approxs f.(fn_code)) + (PTree.map (transf_instr f an rm) f.(fn_code)) f.(fn_entrypoint). -Definition transf_fundef (gapp: global_approx) (fd: fundef) : fundef := - AST.transf_fundef (transf_function gapp) fd. - -Fixpoint make_global_approx (gapp: global_approx) (gdl: list (ident * globdef fundef unit)): global_approx := - match gdl with - | nil => gapp - | (id, gl) :: gdl' => - let gapp1 := - match gl with - | Gfun f => PTree.remove id gapp - | Gvar gv => - if gv.(gvar_readonly) && negb gv.(gvar_volatile) - then PTree.set id gv.(gvar_init) gapp - else PTree.remove id gapp - end in - make_global_approx gapp1 gdl' - end. +Definition transf_fundef (rm: romem) (fd: fundef) : fundef := + AST.transf_fundef (transf_function rm) fd. Definition transf_program (p: program) : program := - let gapp := make_global_approx (PTree.empty _) p.(prog_defs) in - transform_program (transf_fundef gapp) p. + let rm := romem_for_program p in + transform_program (transf_fundef rm) p. diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v index 898c4df..41afe8c 100644 --- a/backend/Constpropproof.v +++ b/backend/Constpropproof.v @@ -26,7 +26,9 @@ Require Import Registers. Require Import RTL. Require Import Lattice. Require Import Kildall. -Require Import Liveness. +Require Import ValueDomain. +Require Import ValueAOp. +Require Import ValueAnalysis. Require Import ConstpropOp. Require Import Constprop. Require Import ConstpropOpproof. @@ -37,316 +39,7 @@ Variable prog: program. Let tprog := transf_program prog. Let ge := Genv.globalenv prog. Let tge := Genv.globalenv tprog. -Let gapp := make_global_approx (PTree.empty _) prog.(prog_defs). - -(** * Correctness of the static analysis *) - -Section ANALYSIS. - -Variable sp: val. - -Definition regs_match_approx (a: D.t) (rs: regset) : Prop := - forall r, val_match_approx ge sp (D.get r a) rs#r. - -Lemma regs_match_approx_top: - forall rs, regs_match_approx D.top rs. -Proof. - intros. red; intros. simpl. rewrite PTree.gempty. - unfold Approx.top, val_match_approx. auto. -Qed. - -Lemma val_match_approx_increasing: - forall a1 a2 v, - Approx.ge a1 a2 -> val_match_approx ge sp a2 v -> val_match_approx ge sp a1 v. -Proof. - intros until v. - intros [A|[B|C]]. - subst a1. simpl. auto. - subst a2. simpl. tauto. - subst a2. auto. -Qed. - -Lemma regs_match_approx_increasing: - forall a1 a2 rs, - D.ge a1 a2 -> regs_match_approx a2 rs -> regs_match_approx a1 rs. -Proof. - unfold D.ge, regs_match_approx. intros. - apply val_match_approx_increasing with (D.get r a2); auto. -Qed. - -Lemma regs_match_approx_update: - forall ra rs a v r, - val_match_approx ge sp a v -> - regs_match_approx ra rs -> - regs_match_approx (D.set r a ra) (rs#r <- v). -Proof. - intros; red; intros. - rewrite D.gsspec. rewrite Regmap.gsspec. destruct (peq r0 r); auto. - red; intros; subst ra. specialize (H0 xH). rewrite D.get_bot in H0. inv H0. - unfold Approx.eq. red; intros; subst a. inv H. -Qed. - -Lemma approx_regs_val_list: - forall ra rs rl, - regs_match_approx ra rs -> - val_list_match_approx ge sp (approx_regs ra rl) rs##rl. -Proof. - induction rl; simpl; intros. - constructor. - constructor. apply H. auto. -Qed. - -Lemma regs_match_approx_forget: - forall rs rl ra, - regs_match_approx ra rs -> - regs_match_approx (List.fold_left (fun a r => D.set r Unknown a) rl ra) rs. -Proof. - induction rl; simpl; intros. - auto. - apply IHrl. - red; intros. rewrite D.gsspec. destruct (peq r a). constructor. auto. - red; intros; subst ra. specialize (H xH). inv H. - unfold Approx.eq, Approx.bot. congruence. -Qed. - -(** The correctness of the static analysis follows from the results - of module [ConstpropOpproof] and the fact that the result of - the static analysis is a solution of the forward dataflow inequations. *) - -Lemma analyze_correct_1: - forall f pc rs pc' i, - f.(fn_code)!pc = Some i -> - In pc' (successors_instr i) -> - regs_match_approx (transfer gapp f pc (analyze gapp f)!!pc) rs -> - regs_match_approx (analyze gapp f)!!pc' rs. -Proof. - unfold analyze; intros. - set (lu := last_uses f) in *. - destruct (DS.fixpoint (fn_code f) successors_instr (transfer' gapp f lu) - ((fn_entrypoint f, D.top) :: nil)) as [approxs|] eqn:FIX. - apply regs_match_approx_increasing with (transfer' gapp f lu pc approxs!!pc). - eapply DS.fixpoint_solution; eauto. - unfold transfer'. destruct (lu!pc) as [regs|]. - apply regs_match_approx_forget; auto. - auto. - intros. rewrite PMap.gi. apply regs_match_approx_top. -Qed. - -Lemma analyze_correct_3: - forall f rs, - regs_match_approx (analyze gapp f)!!(f.(fn_entrypoint)) rs. -Proof. - intros. unfold analyze. - set (lu := last_uses f) in *. - destruct (DS.fixpoint (fn_code f) successors_instr (transfer' gapp f lu) - ((fn_entrypoint f, D.top) :: nil)) as [approxs|] eqn:FIX. - apply regs_match_approx_increasing with D.top. - eapply DS.fixpoint_entry; eauto. auto with coqlib. - apply regs_match_approx_top. - intros. rewrite PMap.gi. apply regs_match_approx_top. -Qed. - -(** eval_static_load *) - -Definition mem_match_approx (m: mem) : Prop := - forall id il b, - gapp!id = Some il -> Genv.find_symbol ge id = Some b -> - Genv.load_store_init_data ge m b 0 il /\ - Mem.valid_block m b /\ - (forall ofs, ~Mem.perm m b ofs Max Writable). - -Lemma eval_load_init_sound: - forall chunk m b il base ofs pos v, - Genv.load_store_init_data ge m b base il -> - Mem.load chunk m b ofs = Some v -> - ofs = base + pos -> - val_match_approx ge sp (eval_load_init chunk pos il) v. -Proof. - induction il; simpl; intros. -(* base case il = nil *) - auto. -(* inductive case *) - destruct a. - (* Init_int8 *) - destruct H. destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. - rewrite Mem.load_int8_signed_unsigned in H0. rewrite H in H0. simpl in H0. - inv H0. decEq. apply Int.sign_ext_zero_ext. compute; auto. - congruence. - eapply IHil; eauto. omega. - (* Init_int16 *) - destruct H. destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. - rewrite Mem.load_int16_signed_unsigned in H0. rewrite H in H0. simpl in H0. - inv H0. decEq. apply Int.sign_ext_zero_ext. compute; auto. - congruence. - eapply IHil; eauto. omega. - (* Init_int32 *) - destruct H. destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. - congruence. - eapply IHil; eauto. omega. - (* Init_int64 *) - destruct H. destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. - congruence. - eapply IHil; eauto. omega. - (* Init_float32 *) - destruct H. destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. destruct (propagate_float_constants tt); simpl; auto. - congruence. - eapply IHil; eauto. omega. - (* Init_float64 *) - destruct H. destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. destruct (propagate_float_constants tt); simpl; auto. - congruence. - eapply IHil; eauto. omega. - (* Init_space *) - eapply IHil; eauto. omega. - (* Init_symbol *) - destruct H as [[b' [A B]] C]. - destruct (zeq pos 0). subst. rewrite Zplus_0_r in H0. - destruct chunk; simpl; auto. - unfold symbol_address. rewrite A. congruence. - eapply IHil; eauto. omega. -Qed. - -Lemma eval_static_load_sound: - forall chunk m addr vaddr v, - Mem.loadv chunk m vaddr = Some v -> - mem_match_approx m -> - val_match_approx ge sp addr vaddr -> - val_match_approx ge sp (eval_static_load gapp chunk addr) v. -Proof. - intros. unfold eval_static_load. destruct addr; simpl; auto. - destruct (gapp!i) as [il|] eqn:?; auto. - red in H1. subst vaddr. unfold symbol_address in H. - destruct (Genv.find_symbol ge i) as [b'|] eqn:?; simpl in H; try discriminate. - exploit H0; eauto. intros [A [B C]]. - eapply eval_load_init_sound; eauto. - red; auto. -Qed. - -Lemma mem_match_approx_store: - forall chunk m addr v m', - mem_match_approx m -> - Mem.storev chunk m addr v = Some m' -> - mem_match_approx m'. -Proof. - intros; red; intros. exploit H; eauto. intros [A [B C]]. - destruct addr; simpl in H0; try discriminate. - exploit Mem.store_valid_access_3; eauto. intros [P Q]. - split. apply Genv.load_store_init_data_invariant with m; auto. - intros. eapply Mem.load_store_other; eauto. left; red; intro; subst b0. - eapply C. apply Mem.perm_cur_max. eapply P. instantiate (1 := Int.unsigned i). - generalize (size_chunk_pos chunk). omega. - split. eauto with mem. - intros; red; intros. eapply C. eapply Mem.perm_store_2; eauto. -Qed. - -Lemma mem_match_approx_alloc: - forall m lo hi b m', - mem_match_approx m -> - Mem.alloc m lo hi = (m', b) -> - mem_match_approx m'. -Proof. - intros; red; intros. exploit H; eauto. intros [A [B C]]. - split. apply Genv.load_store_init_data_invariant with m; auto. - intros. eapply Mem.load_alloc_unchanged; eauto. - split. eauto with mem. - intros; red; intros. exploit Mem.perm_alloc_inv; eauto. - rewrite dec_eq_false. apply C. eapply Mem.valid_not_valid_diff; eauto with mem. -Qed. - -Lemma mem_match_approx_free: - forall m lo hi b m', - mem_match_approx m -> - Mem.free m b lo hi = Some m' -> - mem_match_approx m'. -Proof. - intros; red; intros. exploit H; eauto. intros [A [B C]]. - split. apply Genv.load_store_init_data_invariant with m; auto. - intros. eapply Mem.load_free; eauto. - destruct (eq_block b0 b); auto. subst b0. - right. destruct (zlt lo hi); auto. - elim (C lo). apply Mem.perm_cur_max. - exploit Mem.free_range_perm; eauto. instantiate (1 := lo); omega. - intros; eapply Mem.perm_implies; eauto with mem. - split. eauto with mem. - intros; red; intros. eapply C. eauto with mem. -Qed. - -Lemma mem_match_approx_extcall: - forall ef vargs m t vres m', - mem_match_approx m -> - external_call ef ge vargs m t vres m' -> - mem_match_approx m'. -Proof. - intros; red; intros. exploit H; eauto. intros [A [B C]]. - split. apply Genv.load_store_init_data_invariant with m; auto. - intros. eapply external_call_readonly; eauto. - split. eapply external_call_valid_block; eauto. - intros; red; intros. elim (C ofs). eapply external_call_max_perm; eauto. -Qed. - -(* Show that mem_match_approx holds initially *) - -Definition global_approx_charact (g: genv) (ga: global_approx) : Prop := - forall id il b, - ga!id = Some il -> - Genv.find_symbol g id = Some b -> - Genv.find_var_info g b = Some (mkglobvar tt il true false). - -Lemma make_global_approx_correct: - forall gdl g ga, - global_approx_charact g ga -> - global_approx_charact (Genv.add_globals g gdl) (make_global_approx ga gdl). -Proof. - induction gdl; simpl; intros. - auto. - destruct a as [id gd]. apply IHgdl. - red; intros. - assert (EITHER: id0 = id /\ gd = Gvar(mkglobvar tt il true false) - \/ id0 <> id /\ ga!id0 = Some il). - destruct gd. - rewrite PTree.grspec in H0. destruct (PTree.elt_eq id0 id); [discriminate|auto]. - destruct (gvar_readonly v && negb (gvar_volatile v)) eqn:?. - rewrite PTree.gsspec in H0. destruct (peq id0 id). - inv H0. left. split; auto. - destruct v; simpl in *. - destruct gvar_readonly; try discriminate. - destruct gvar_volatile; try discriminate. - destruct gvar_info. auto. - auto. - rewrite PTree.grspec in H0. destruct (PTree.elt_eq id0 id); [discriminate|auto]. - - unfold Genv.add_global, Genv.find_symbol, Genv.find_var_info in *; - simpl in *. - destruct EITHER as [[A B] | [A B]]. - subst id0. rewrite PTree.gss in H1. inv H1. rewrite PTree.gss. auto. - rewrite PTree.gso in H1; auto. destruct gd. eapply H; eauto. - rewrite PTree.gso. eapply H; eauto. - red; intros; subst b. eelim Plt_strict; eapply Genv.genv_symb_range; eauto. -Qed. - -Theorem mem_match_approx_init: - forall m, Genv.init_mem prog = Some m -> mem_match_approx m. -Proof. - intros. - assert (global_approx_charact ge gapp). - unfold ge, gapp. unfold Genv.globalenv. - apply make_global_approx_correct. - red; intros. rewrite PTree.gempty in H0; discriminate. - red; intros. - exploit Genv.init_mem_characterization. - unfold ge in H0. eapply H0; eauto. eauto. - unfold Genv.perm_globvar; simpl. - intros [A [B C]]. - split. auto. split. eapply Genv.find_symbol_not_fresh; eauto. - intros; red; intros. exploit B; eauto. intros [P Q]. inv Q. -Qed. - -End ANALYSIS. +Let rm := romem_for_program prog. (** * Correctness of the code transformation *) @@ -370,24 +63,24 @@ Qed. Lemma functions_translated: forall (v: val) (f: fundef), Genv.find_funct ge v = Some f -> - Genv.find_funct tge v = Some (transf_fundef gapp f). + Genv.find_funct tge v = Some (transf_fundef rm f). Proof. intros. - exact (Genv.find_funct_transf (transf_fundef gapp) _ _ H). + exact (Genv.find_funct_transf (transf_fundef rm) _ _ H). Qed. Lemma function_ptr_translated: forall (b: block) (f: fundef), Genv.find_funct_ptr ge b = Some f -> - Genv.find_funct_ptr tge b = Some (transf_fundef gapp f). + Genv.find_funct_ptr tge b = Some (transf_fundef rm f). Proof. intros. - exact (Genv.find_funct_ptr_transf (transf_fundef gapp) _ _ H). + exact (Genv.find_funct_ptr_transf (transf_fundef rm) _ _ H). Qed. Lemma sig_function_translated: forall f, - funsig (transf_fundef gapp f) = funsig f. + funsig (transf_fundef rm f) = funsig f. Proof. intros. destruct f; reflexivity. Qed. @@ -422,82 +115,103 @@ Proof. Qed. Lemma transf_ros_correct: - forall sp ros rs rs' f approx, - regs_match_approx sp approx rs -> + forall bc rs ae ros f rs', + genv_match bc ge -> + ematch bc rs ae -> find_function ge ros rs = Some f -> regs_lessdef rs rs' -> - find_function tge (transf_ros approx ros) rs' = Some (transf_fundef gapp f). -Proof. - intros. destruct ros; simpl in *. - generalize (H r); intro MATCH. generalize (H1 r); intro LD. - destruct (rs#r); simpl in H0; try discriminate. - destruct (Int.eq_dec i Int.zero); try discriminate. - inv LD. - assert (find_function tge (inl _ r) rs' = Some (transf_fundef gapp f)). - simpl. rewrite <- H4. simpl. rewrite dec_eq_true. apply function_ptr_translated. auto. - destruct (D.get r approx); auto. - predSpec Int.eq Int.eq_spec i0 Int.zero; intros; auto. - simpl in *. unfold symbol_address in MATCH. rewrite symbols_preserved. - destruct (Genv.find_symbol ge i); try discriminate. - inv MATCH. apply function_ptr_translated; auto. - rewrite symbols_preserved. destruct (Genv.find_symbol ge i); try discriminate. + find_function tge (transf_ros ae ros) rs' = Some (transf_fundef rm f). +Proof. + intros until rs'; intros GE EM FF RLD. destruct ros; simpl in *. +- (* function pointer *) + generalize (EM r); fold (areg ae r); intro VM. generalize (RLD r); intro LD. + assert (DEFAULT: find_function tge (inl _ r) rs' = Some (transf_fundef rm f)). + { + simpl. inv LD. apply functions_translated; auto. rewrite <- H0 in FF; discriminate. + } + destruct (areg ae r); auto. destruct p; auto. + predSpec Int.eq Int.eq_spec ofs Int.zero; intros; auto. + subst ofs. exploit vmatch_ptr_gl; eauto. intros LD'. inv LD'; try discriminate. + rewrite H1 in FF. unfold symbol_address in FF. + simpl. rewrite symbols_preserved. + destruct (Genv.find_symbol ge id) as [b|]; try discriminate. + simpl in FF. rewrite dec_eq_true in FF. + apply function_ptr_translated; auto. + rewrite <- H0 in FF; discriminate. +- (* function symbol *) + rewrite symbols_preserved. + destruct (Genv.find_symbol ge i) as [b|]; try discriminate. apply function_ptr_translated; auto. Qed. Lemma const_for_result_correct: - forall a op sp v m, + forall a op bc v sp m, const_for_result a = Some op -> - val_match_approx ge sp a v -> - eval_operation tge sp op nil m = Some v. -Proof. - unfold const_for_result; intros. - destruct a; inv H; simpl in H0. - simpl. congruence. - destruct (generate_float_constants tt); inv H2. simpl. congruence. - simpl. subst v. unfold symbol_address. rewrite symbols_preserved. auto. - simpl. congruence. -Qed. - -Inductive match_pc (f: function) (app: D.t): nat -> node -> node -> Prop := + vmatch bc v a -> + bc sp = BCstack -> + genv_match bc ge -> + exists v', eval_operation tge (Vptr sp Int.zero) op nil m = Some v' /\ Val.lessdef v v'. +Proof. + unfold const_for_result; intros. + destruct a; try discriminate. +- (* integer *) + inv H. inv H0. exists (Vint n); auto. +- (* float *) + destruct (generate_float_constants tt); inv H. inv H0. exists (Vfloat f); auto. +- (* pointer *) + destruct p; try discriminate. + + (* global *) + inv H. exists (symbol_address ge id ofs); split. + unfold symbol_address. rewrite <- symbols_preserved. reflexivity. + eapply vmatch_ptr_gl; eauto. + + (* stack *) + inv H. exists (Vptr sp ofs); split. + simpl; rewrite Int.add_zero_l; auto. + eapply vmatch_ptr_stk; eauto. +Qed. + +Inductive match_pc (f: function) (ae: AE.t): nat -> node -> node -> Prop := | match_pc_base: forall n pc, - match_pc f app n pc pc + match_pc f ae n pc pc | match_pc_nop: forall n pc s pcx, f.(fn_code)!pc = Some (Inop s) -> - match_pc f app n s pcx -> - match_pc f app (Datatypes.S n) pc pcx - | match_pc_cond: forall n pc cond args s1 s2 b, + match_pc f ae n s pcx -> + match_pc f ae (S n) pc pcx + | match_pc_cond: forall n pc cond args s1 s2 b pcx, f.(fn_code)!pc = Some (Icond cond args s1 s2) -> - eval_static_condition cond (approx_regs app args) = Some b -> - match_pc f app (Datatypes.S n) pc (if b then s1 else s2). + resolve_branch (eval_static_condition cond (aregs ae args)) = Some b -> + match_pc f ae n (if b then s1 else s2) pcx -> + match_pc f ae (S n) pc pcx. Lemma match_successor_rec: - forall f app n pc, match_pc f app n pc (successor_rec n f app pc). + forall f ae n pc, match_pc f ae n pc (successor_rec n f ae pc). Proof. induction n; simpl; intros. - apply match_pc_base. - destruct (fn_code f)!pc as [i|] eqn:?; try apply match_pc_base. - destruct i; try apply match_pc_base. - eapply match_pc_nop; eauto. - destruct (eval_static_condition c (approx_regs app l)) as [b|] eqn:?. +- apply match_pc_base. +- destruct (fn_code f)!pc as [[]|] eqn:INSTR; try apply match_pc_base. + eapply match_pc_nop; eauto. + destruct (resolve_branch (eval_static_condition c (aregs ae l))) as [b|] eqn:COND. eapply match_pc_cond; eauto. apply match_pc_base. Qed. Lemma match_successor: - forall f app pc, match_pc f app num_iter pc (successor f app pc). + forall f ae pc, match_pc f ae num_iter pc (successor f ae pc). Proof. unfold successor; intros. apply match_successor_rec. Qed. Section BUILTIN_STRENGTH_REDUCTION. -Variable app: D.t. -Variable sp: val. + +Variable bc: block_classification. +Hypothesis GE: genv_match bc ge. +Variable ae: AE.t. Variable rs: regset. -Hypothesis MATCH: forall r, val_match_approx ge sp (approx_reg app r) rs#r. +Hypothesis MATCH: ematch bc rs ae. Lemma annot_strength_reduction_correct: forall targs args targs' args' eargs, - annot_strength_reduction app targs args = (targs', args') -> + annot_strength_reduction ae targs args = (targs', args') -> eventval_list_match ge eargs (annot_args_typ targs) rs##args -> exists eargs', eventval_list_match ge eargs' (annot_args_typ targs') rs##args' @@ -507,7 +221,7 @@ Proof. - inv H. simpl. exists eargs; auto. - destruct a. + destruct args as [ | arg args0]; simpl in H0; inv H0. - destruct (annot_strength_reduction app targs args0) as [targs'' args''] eqn:E. + destruct (annot_strength_reduction ae targs args0) as [targs'' args''] eqn:E. exploit IHtargs; eauto. intros [eargs'' [A B]]. assert (DFL: exists eargs', @@ -517,41 +231,48 @@ Proof. exists (ev1 :: eargs''); split. simpl; constructor; auto. simpl. congruence. } - destruct ty; destruct (approx_reg app arg) as [] eqn:E2; inv H; auto. - * exists eargs''; split; auto; simpl; f_equal; auto. - generalize (MATCH arg); rewrite E2; simpl; intros E3; - rewrite E3 in H5; inv H5; auto. + destruct ty; destruct (areg ae arg) as [] eqn:E2; inv H; auto. + * exists eargs''; split; auto; simpl; f_equal; auto. + generalize (MATCH arg); fold (areg ae arg); rewrite E2; intros VM. + inv VM. rewrite <- H0 in *. inv H5; auto. * destruct (generate_float_constants tt); inv H1; auto. - exists eargs''; split; auto; simpl; f_equal; auto. - generalize (MATCH arg); rewrite E2; simpl; intros E3; - rewrite E3 in H5; inv H5; auto. - + destruct (annot_strength_reduction app targs args) as [targs'' args''] eqn:E. + exists eargs''; split; auto; simpl; f_equal; auto. + generalize (MATCH arg); fold (areg ae arg); rewrite E2; intros VM. + inv VM. rewrite <- H0 in *. inv H5; auto. + + destruct (annot_strength_reduction ae targs args) as [targs'' args''] eqn:E. inv H. exploit IHtargs; eauto. intros [eargs'' [A B]]. exists eargs''; simpl; split; auto. congruence. - + destruct (annot_strength_reduction app targs args) as [targs'' args''] eqn:E. + + destruct (annot_strength_reduction ae targs args) as [targs'' args''] eqn:E. inv H. exploit IHtargs; eauto. intros [eargs'' [A B]]. exists eargs''; simpl; split; auto. congruence. Qed. +Lemma vmatch_ptr_gl': + forall v id ofs, + vmatch bc v (Ptr (Gl id ofs)) -> + v = Vundef \/ exists b, Genv.find_symbol ge id = Some b /\ v = Vptr b ofs. +Proof. + intros. inv H; auto. inv H2. right; exists b; split; auto. eapply GE; eauto. +Qed. + Lemma builtin_strength_reduction_correct: forall ef args m t vres m', external_call ef ge rs##args m t vres m' -> - let (ef', args') := builtin_strength_reduction app ef args in + let (ef', args') := builtin_strength_reduction ae ef args in external_call ef' ge rs##args' m t vres m'. Proof. - intros until m'. functional induction (builtin_strength_reduction app ef args); intros; auto. -+ generalize (MATCH r1); rewrite e1; simpl; intros E. simpl in H. - unfold symbol_address in E. destruct (Genv.find_symbol ge symb) as [b|] eqn:?; rewrite E in H. - rewrite volatile_load_global_charact. exists b; auto. - inv H. -+ generalize (MATCH r1); rewrite e1; simpl; intros E. simpl in H. - unfold symbol_address in E. destruct (Genv.find_symbol ge symb) as [b|] eqn:?; rewrite E in H. - rewrite volatile_store_global_charact. exists b; auto. - inv H. -+ inv H. exploit annot_strength_reduction_correct; eauto. - intros [eargs' [A B]]. + intros until m'. functional induction (builtin_strength_reduction ae ef args); intros; auto. ++ simpl in H. assert (V: vmatch bc (rs#r1) (Ptr (Gl symb n1))) by (rewrite <- e1; apply MATCH). + exploit vmatch_ptr_gl'; eauto. intros [A | [b [A B]]]. + * simpl in H; rewrite A in H; inv H. + * simpl; rewrite volatile_load_global_charact. exists b; split; congruence. ++ simpl in H. assert (V: vmatch bc (rs#r1) (Ptr (Gl symb n1))) by (rewrite <- e1; apply MATCH). + exploit vmatch_ptr_gl'; eauto. intros [A | [b [A B]]]. + * simpl in H; rewrite A in H; inv H. + * simpl; rewrite volatile_store_global_charact. exists b; split; congruence. ++ inv H. exploit annot_strength_reduction_correct; eauto. intros [eargs' [A B]]. rewrite <- B. econstructor; eauto. Qed. @@ -575,9 +296,8 @@ End BUILTIN_STRENGTH_REDUCTION. and an initial state [st2] in the transformed code. This invariant expresses that all code fragments appearing in [st2] are obtained by [transf_code] transformation of the corresponding - fragments in [st1]. Moreover, the values of registers in [st1] - must match their compile-time approximations at the current program - point. + fragments in [st1]. Moreover, the state [st1] must match its compile-time + approximations at the current program point. These two parts of the diagram are the hypotheses. In conclusions, we want to prove the other two parts: the right vertical arrow, which is a transition in the transformed RTL code, and the bottom @@ -588,72 +308,63 @@ Inductive match_stackframes: stackframe -> stackframe -> Prop := match_stackframe_intro: forall res sp pc rs f rs', regs_lessdef rs rs' -> - (forall v, regs_match_approx sp (analyze gapp f)!!pc (rs#res <- v)) -> match_stackframes (Stackframe res f sp pc rs) - (Stackframe res (transf_function gapp f) sp pc rs'). + (Stackframe res (transf_function rm f) sp pc rs'). Inductive match_states: nat -> state -> state -> Prop := | match_states_intro: - forall s sp pc rs m f s' pc' rs' m' app n - (MATCH1: regs_match_approx sp app rs) - (MATCH2: regs_match_approx sp (analyze gapp f)!!pc rs) - (GMATCH: mem_match_approx m) + forall s sp pc rs m f s' pc' rs' m' bc ae n + (MATCH: ematch bc rs ae) (STACKS: list_forall2 match_stackframes s s') - (PC: match_pc f app n pc pc') + (PC: match_pc f ae n pc pc') (REGS: regs_lessdef rs rs') (MEM: Mem.extends m m'), match_states n (State s f sp pc rs m) - (State s' (transf_function gapp f) sp pc' rs' m') + (State s' (transf_function rm f) sp pc' rs' m') | match_states_call: forall s f args m s' args' m' - (GMATCH: mem_match_approx m) (STACKS: list_forall2 match_stackframes s s') (ARGS: Val.lessdef_list args args') (MEM: Mem.extends m m'), match_states O (Callstate s f args m) - (Callstate s' (transf_fundef gapp f) args' m') + (Callstate s' (transf_fundef rm f) args' m') | match_states_return: forall s v m s' v' m' - (GMATCH: mem_match_approx m) (STACKS: list_forall2 match_stackframes s s') (RES: Val.lessdef v v') (MEM: Mem.extends m m'), list_forall2 match_stackframes s s' -> match_states O (Returnstate s v m) - (Returnstate s' v' m'). + (Returnstate s' v' m'). Lemma match_states_succ: - forall s f sp pc2 rs m s' rs' m' pc1 i, - f.(fn_code)!pc1 = Some i -> - In pc2 (successors_instr i) -> - regs_match_approx sp (transfer gapp f pc1 (analyze gapp f)!!pc1) rs -> - mem_match_approx m -> + forall s f sp pc rs m s' rs' m', + sound_state prog (State s f sp pc rs m) -> list_forall2 match_stackframes s s' -> regs_lessdef rs rs' -> Mem.extends m m' -> - match_states O (State s f sp pc2 rs m) - (State s' (transf_function gapp f) sp pc2 rs' m'). + match_states O (State s f sp pc rs m) + (State s' (transf_function rm f) sp pc rs' m'). Proof. - intros. - assert (regs_match_approx sp (analyze gapp f)!!pc2 rs). - eapply analyze_correct_1; eauto. - apply match_states_intro with (app := (analyze gapp f)!!pc2); auto. + intros. inv H. + apply match_states_intro with (bc := bc) (ae := ae); auto. constructor. Qed. Lemma transf_instr_at: forall f pc i, f.(fn_code)!pc = Some i -> - (transf_function gapp f).(fn_code)!pc = Some(transf_instr gapp f (analyze gapp f) pc i). + (transf_function rm f).(fn_code)!pc = Some(transf_instr f (analyze rm f) rm pc i). Proof. - intros. simpl. unfold transf_code. rewrite PTree.gmap. rewrite H. auto. + intros. simpl. rewrite PTree.gmap. rewrite H. auto. Qed. Ltac TransfInstr := match goal with - | H: (PTree.get ?pc (fn_code ?f) = Some ?instr) |- _ => - generalize (transf_instr_at _ _ _ H); simpl + | H1: (PTree.get ?pc (fn_code ?f) = Some ?instr), + H2: (analyze (romem_for_program prog) ?f)#?pc = VA.State ?ae ?am |- _ => + fold rm in H2; generalize (transf_instr_at _ _ _ H1); unfold transf_instr; rewrite H2 end. (** The proof of simulation proceeds by case analysis on the transition @@ -662,113 +373,107 @@ Ltac TransfInstr := Lemma transf_step_correct: forall s1 t s2, step ge s1 t s2 -> - forall n1 s1' (MS: match_states n1 s1 s1'), + forall n1 s1' (SS1: sound_state prog s1) (SS2: sound_state prog s2) (MS: match_states n1 s1 s1'), (exists n2, exists s2', step tge s1' t s2' /\ match_states n2 s2 s2') \/ (exists n2, n2 < n1 /\ t = E0 /\ match_states n2 s2 s1')%nat. Proof. - induction 1; intros; inv MS; try (inv PC; try congruence). + induction 1; intros; inv SS1; inv MS; try (inv PC; try congruence). (* Inop, preserved *) - rename pc'0 into pc. TransfInstr; intro. + rename pc'0 into pc. TransfInstr; intros. left; econstructor; econstructor; split. eapply exec_Inop; eauto. - eapply match_states_succ; eauto. simpl; auto. - unfold transfer; rewrite H. auto. + eapply match_states_succ; eauto. (* Inop, skipped over *) - rewrite H0 in H; inv H. + assert (s0 = pc') by congruence. subst s0. right; exists n; split. omega. split. auto. - apply match_states_intro with app; auto. - eapply analyze_correct_1; eauto. simpl; auto. - unfold transfer; rewrite H0. auto. + apply match_states_intro with bc0 ae0; auto. (* Iop *) rename pc'0 into pc. TransfInstr. - set (app_before := (analyze gapp f)#pc). - set (a := eval_static_operation op (approx_regs app_before args)). - set (app_after := D.set res a app_before). - assert (VMATCH: val_match_approx ge sp a v). - eapply eval_static_operation_correct; eauto. - apply approx_regs_val_list; auto. - assert (MATCH': regs_match_approx sp app_after rs#res <- v). - apply regs_match_approx_update; auto. - assert (MATCH'': regs_match_approx sp (analyze gapp f) # pc' rs # res <- v). - eapply analyze_correct_1 with (pc := pc); eauto. simpl; auto. - unfold transfer; rewrite H. auto. + set (a := eval_static_operation op (aregs ae args)). + set (ae' := AE.set res a ae). + assert (VMATCH: vmatch bc v a) by (eapply eval_static_operation_sound; eauto with va). + assert (MATCH': ematch bc (rs#res <- v) ae') by (eapply ematch_update; eauto). destruct (const_for_result a) as [cop|] eqn:?; intros. (* constant is propagated *) + exploit const_for_result_correct; eauto. intros (v' & A & B). left; econstructor; econstructor; split. eapply exec_Iop; eauto. - eapply const_for_result_correct; eauto. - apply match_states_intro with app_after; auto. + apply match_states_intro with bc ae'; auto. apply match_successor. apply set_reg_lessdef; auto. (* operator is strength-reduced *) - exploit op_strength_reduction_correct. eexact MATCH2. reflexivity. eauto. - fold app_before. - destruct (op_strength_reduction op args (approx_regs app_before args)) as [op' args']. - intros [v' [EV' LD']]. - assert (EV'': exists v'', eval_operation ge sp op' rs'##args' m' = Some v'' /\ Val.lessdef v' v''). - eapply eval_operation_lessdef; eauto. eapply regs_lessdef_regs; eauto. + assert(OP: + let (op', args') := op_strength_reduction op args (aregs ae args) in + exists v', + eval_operation ge (Vptr sp0 Int.zero) op' rs ## args' m = Some v' /\ + Val.lessdef v v'). + { eapply op_strength_reduction_correct with (ae := ae); eauto with va. } + destruct (op_strength_reduction op args (aregs ae args)) as [op' args']. + destruct OP as [v' [EV' LD']]. + assert (EV'': exists v'', eval_operation ge (Vptr sp0 Int.zero) op' rs'##args' m' = Some v'' /\ Val.lessdef v' v''). + { eapply eval_operation_lessdef; eauto. eapply regs_lessdef_regs; eauto. } destruct EV'' as [v'' [EV'' LD'']]. left; econstructor; econstructor; split. eapply exec_Iop; eauto. erewrite eval_operation_preserved. eexact EV''. exact symbols_preserved. - apply match_states_intro with app_after; auto. + apply match_states_intro with bc ae'; auto. apply match_successor. apply set_reg_lessdef; auto. eapply Val.lessdef_trans; eauto. (* Iload *) - rename pc'0 into pc. TransfInstr. - set (ap1 := eval_static_addressing addr - (approx_regs (analyze gapp f) # pc args)). - set (ap2 := eval_static_load gapp chunk ap1). - assert (VM1: val_match_approx ge sp ap1 a). - eapply eval_static_addressing_correct; eauto. - eapply approx_regs_val_list; eauto. - assert (VM2: val_match_approx ge sp ap2 v). - eapply eval_static_load_sound; eauto. - destruct (const_for_result ap2) as [cop|] eqn:?; intros. + rename pc'0 into pc. TransfInstr. + set (aa := eval_static_addressing addr (aregs ae args)). + assert (VM1: vmatch bc a aa) by (eapply eval_static_addressing_sound; eauto with va). + set (av := loadv chunk rm am aa). + assert (VM2: vmatch bc v av) by (eapply loadv_sound; eauto). + destruct (const_for_result av) as [cop|] eqn:?; intros. (* constant-propagated *) + exploit const_for_result_correct; eauto. intros (v' & A & B). left; econstructor; econstructor; split. - eapply exec_Iop; eauto. eapply const_for_result_correct; eauto. - eapply match_states_succ; eauto. simpl; auto. - unfold transfer; rewrite H. apply regs_match_approx_update; auto. + eapply exec_Iop; eauto. + eapply match_states_succ; eauto. apply set_reg_lessdef; auto. (* strength-reduced *) - generalize (addr_strength_reduction_correct ge sp (analyze gapp f)!!pc rs - MATCH2 addr args (approx_regs (analyze gapp f) # pc args) (refl_equal _)). - destruct (addr_strength_reduction addr args (approx_regs (analyze gapp f) # pc args)) as [addr' args']. - rewrite H0. intros P. - assert (ADDR': exists a', eval_addressing ge sp addr' rs'##args' = Some a' /\ Val.lessdef a a'). - eapply eval_addressing_lessdef; eauto. eapply regs_lessdef_regs; eauto. - destruct ADDR' as [a' [A B]]. - assert (C: eval_addressing tge sp addr' rs'##args' = Some a'). - rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved. - exploit Mem.loadv_extends; eauto. intros [v' [D E]]. + assert (ADDR: + let (addr', args') := addr_strength_reduction addr args (aregs ae args) in + exists a', + eval_addressing ge (Vptr sp0 Int.zero) addr' rs ## args' = Some a' /\ + Val.lessdef a a'). + { eapply addr_strength_reduction_correct with (ae := ae); eauto with va. } + destruct (addr_strength_reduction addr args (aregs ae args)) as [addr' args']. + destruct ADDR as (a' & P & Q). + exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P. + intros (a'' & U & V). + assert (W: eval_addressing tge (Vptr sp0 Int.zero) addr' rs'##args' = Some a''). + { rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. } + exploit Mem.loadv_extends. eauto. eauto. apply Val.lessdef_trans with a'; eauto. + intros (v' & X & Y). left; econstructor; econstructor; split. eapply exec_Iload; eauto. - eapply match_states_succ; eauto. simpl; auto. - unfold transfer; rewrite H. apply regs_match_approx_update; auto. - apply set_reg_lessdef; auto. + eapply match_states_succ; eauto. apply set_reg_lessdef; auto. (* Istore *) rename pc'0 into pc. TransfInstr. - generalize (addr_strength_reduction_correct ge sp (analyze gapp f)!!pc rs - MATCH2 addr args (approx_regs (analyze gapp f) # pc args) (refl_equal _)). - destruct (addr_strength_reduction addr args (approx_regs (analyze gapp f) # pc args)) as [addr' args']. - intros P Q. rewrite H0 in P. - assert (ADDR': exists a', eval_addressing ge sp addr' rs'##args' = Some a' /\ Val.lessdef a a'). - eapply eval_addressing_lessdef; eauto. eapply regs_lessdef_regs; eauto. - destruct ADDR' as [a' [A B]]. - assert (C: eval_addressing tge sp addr' rs'##args' = Some a'). - rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved. - exploit Mem.storev_extends; eauto. intros [m2' [D E]]. + assert (ADDR: + let (addr', args') := addr_strength_reduction addr args (aregs ae args) in + exists a', + eval_addressing ge (Vptr sp0 Int.zero) addr' rs ## args' = Some a' /\ + Val.lessdef a a'). + { eapply addr_strength_reduction_correct with (ae := ae); eauto with va. } + destruct (addr_strength_reduction addr args (aregs ae args)) as [addr' args']. + destruct ADDR as (a' & P & Q). + exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P. + intros (a'' & U & V). + assert (W: eval_addressing tge (Vptr sp0 Int.zero) addr' rs'##args' = Some a''). + { rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. } + exploit Mem.storev_extends. eauto. eauto. apply Val.lessdef_trans with a'; eauto. apply REGS. + intros (m2' & X & Y). left; econstructor; econstructor; split. eapply exec_Istore; eauto. - eapply match_states_succ; eauto. simpl; auto. - unfold transfer; rewrite H. auto. - eapply mem_match_approx_store; eauto. + eapply match_states_succ; eauto. (* Icall *) rename pc'0 into pc. @@ -777,10 +482,7 @@ Proof. left; econstructor; econstructor; split. eapply exec_Icall; eauto. apply sig_function_translated; auto. constructor; auto. constructor; auto. - econstructor; eauto. - intros. eapply analyze_correct_1; eauto. simpl; auto. - unfold transfer; rewrite H. - apply regs_match_approx_update; auto. simpl. auto. + econstructor; eauto. apply regs_lessdef_regs; auto. (* Itailcall *) @@ -790,7 +492,6 @@ Proof. left; econstructor; econstructor; split. eapply exec_Itailcall; eauto. apply sig_function_translated; auto. constructor; auto. - eapply mem_match_approx_free; eauto. apply regs_lessdef_regs; auto. (* Ibuiltin *) @@ -798,7 +499,7 @@ Proof. Opaque builtin_strength_reduction. exploit builtin_strength_reduction_correct; eauto. TransfInstr. - destruct (builtin_strength_reduction (analyze gapp f)#pc ef args) as [ef' args']. + destruct (builtin_strength_reduction ae ef args) as [ef' args']. intros P Q. exploit external_call_mem_extends; eauto. instantiate (1 := rs'##args'). apply regs_lessdef_regs; auto. @@ -808,74 +509,68 @@ Opaque builtin_strength_reduction. eapply external_call_symbols_preserved; eauto. exact symbols_preserved. exact varinfo_preserved. eapply match_states_succ; eauto. simpl; auto. - unfold transfer; rewrite H. - apply regs_match_approx_update; auto. simpl; auto. - eapply mem_match_approx_extcall; eauto. apply set_reg_lessdef; auto. (* Icond, preserved *) - rename pc'0 into pc. TransfInstr. - generalize (cond_strength_reduction_correct ge sp (analyze gapp f)#pc rs m - MATCH2 cond args (approx_regs (analyze gapp f) # pc args) (refl_equal _)). - destruct (cond_strength_reduction cond args (approx_regs (analyze gapp f) # pc args)) as [cond' args']. + rename pc' into pc. TransfInstr. + set (ac := eval_static_condition cond (aregs ae args)). + assert (C: cmatch (eval_condition cond rs ## args m) ac) + by (eapply eval_static_condition_sound; eauto with va). + rewrite H0 in C. + generalize (cond_strength_reduction_correct bc ae rs m EM cond args (aregs ae args) (refl_equal _)). + destruct (cond_strength_reduction cond args (aregs ae args)) as [cond' args']. intros EV1 TCODE. - left; exists O; exists (State s' (transf_function gapp f) sp (if b then ifso else ifnot) rs' m'); split. - destruct (eval_static_condition cond (approx_regs (analyze gapp f) # pc args)) eqn:?. - assert (eval_condition cond rs ## args m = Some b0). - eapply eval_static_condition_correct; eauto. eapply approx_regs_val_list; eauto. - assert (b = b0) by congruence. subst b0. + left; exists O; exists (State s' (transf_function rm f) (Vptr sp0 Int.zero) (if b then ifso else ifnot) rs' m'); split. + destruct (resolve_branch ac) eqn: RB. + assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0. destruct b; eapply exec_Inop; eauto. eapply exec_Icond; eauto. eapply eval_condition_lessdef with (vl1 := rs##args'); eauto. eapply regs_lessdef_regs; eauto. congruence. eapply match_states_succ; eauto. - destruct b; simpl; auto. - unfold transfer; rewrite H. auto. (* Icond, skipped over *) rewrite H1 in H; inv H. - assert (eval_condition cond rs ## args m = Some b0). - eapply eval_static_condition_correct; eauto. eapply approx_regs_val_list; eauto. - assert (b = b0) by congruence. subst b0. - right; exists n; split. omega. split. auto. - assert (MATCH': regs_match_approx sp (analyze gapp f) # (if b then ifso else ifnot) rs). - eapply analyze_correct_1; eauto. destruct b; simpl; auto. - unfold transfer; rewrite H1; auto. - econstructor; eauto. constructor. + set (ac := eval_static_condition cond (aregs ae0 args)) in *. + assert (C: cmatch (eval_condition cond rs ## args m) ac) + by (eapply eval_static_condition_sound; eauto with va). + rewrite H0 in C. + assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0. + right; exists n; split. omega. split. auto. + econstructor; eauto. (* Ijumptable *) rename pc'0 into pc. - assert (A: (fn_code (transf_function gapp f))!pc = Some(Ijumptable arg tbl) - \/ (fn_code (transf_function gapp f))!pc = Some(Inop pc')). - TransfInstr. destruct (approx_reg (analyze gapp f) # pc arg) eqn:?; auto. - generalize (MATCH2 arg). unfold approx_reg in Heqt. rewrite Heqt. rewrite H0. - simpl. intro EQ; inv EQ. rewrite H1. auto. - assert (B: rs'#arg = Vint n). - generalize (REGS arg); intro LD; inv LD; congruence. - left; exists O; exists (State s' (transf_function gapp f) sp pc' rs' m'); split. + assert (A: (fn_code (transf_function rm f))!pc = Some(Ijumptable arg tbl) + \/ (fn_code (transf_function rm f))!pc = Some(Inop pc')). + { TransfInstr. + destruct (areg ae arg) eqn:A; auto. + generalize (EM arg). fold (areg ae arg); rewrite A. + intros V; inv V. replace n0 with n by congruence. + rewrite H1. auto. } + assert (rs'#arg = Vint n). + { generalize (REGS arg). rewrite H0. intros LD; inv LD; auto. } + left; exists O; exists (State s' (transf_function rm f) (Vptr sp0 Int.zero) pc' rs' m'); split. destruct A. eapply exec_Ijumptable; eauto. eapply exec_Inop; eauto. eapply match_states_succ; eauto. - simpl. eapply list_nth_z_in; eauto. - unfold transfer; rewrite H; auto. (* Ireturn *) exploit Mem.free_parallel_extends; eauto. intros [m2' [A B]]. left; exists O; exists (Returnstate s' (regmap_optget or Vundef rs') m2'); split. eapply exec_Ireturn; eauto. TransfInstr; auto. constructor; auto. - eapply mem_match_approx_free; eauto. destruct or; simpl; auto. (* internal function *) exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl. intros [m2' [A B]]. + assert (X: exists bc ae, ematch bc (init_regs args (fn_params f)) ae). + { inv SS2. exists bc0; exists ae; auto. } + destruct X as (bc1 & ae1 & MATCH). simpl. unfold transf_function. left; exists O; econstructor; split. eapply exec_function_internal; simpl; eauto. simpl. econstructor; eauto. - apply analyze_correct_3; auto. - apply analyze_correct_3; auto. - eapply mem_match_approx_alloc; eauto. - instantiate (1 := f). constructor. + constructor. apply init_regs_lessdef; auto. (* external function *) @@ -886,9 +581,11 @@ Opaque builtin_strength_reduction. eapply external_call_symbols_preserved; eauto. exact symbols_preserved. exact varinfo_preserved. constructor; auto. - eapply mem_match_approx_extcall; eauto. (* return *) + assert (X: exists bc ae, ematch bc (rs#res <- vres) ae). + { inv SS2. exists bc0; exists ae; auto. } + destruct X as (bc1 & ae1 & MATCH). inv H4. inv H1. left; exists O; econstructor; split. eapply exec_return; eauto. @@ -901,16 +598,14 @@ Lemma transf_initial_states: Proof. intros. inversion H. exploit function_ptr_translated; eauto. intro FIND. - exists O; exists (Callstate nil (transf_fundef gapp f) nil m0); split. + exists O; exists (Callstate nil (transf_fundef rm f) nil m0); split. econstructor; eauto. apply Genv.init_mem_transf; auto. replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto. reflexivity. rewrite <- H3. apply sig_function_translated. - constructor. - eapply mem_match_approx_init; eauto. - constructor. constructor. apply Mem.extends_refl. + constructor. constructor. constructor. apply Mem.extends_refl. Qed. Lemma transf_final_states: @@ -926,16 +621,21 @@ Qed. Theorem transf_program_correct: forward_simulation (RTL.semantics prog) (RTL.semantics tprog). Proof. - eapply Forward_simulation with (fsim_order := lt); simpl. - apply lt_wf. - eexact transf_initial_states. - eexact transf_final_states. - fold ge; fold tge. intros. - exploit transf_step_correct; eauto. - intros [ [n2 [s2' [A B]]] | [n2 [A [B C]]]]. - exists n2; exists s2'; split; auto. left; apply plus_one; auto. - exists n2; exists s2; split; auto. right; split; auto. subst t; apply star_refl. - eexact symbols_preserved. + apply Forward_simulation with + (fsim_order := lt) + (fsim_match_states := fun n s1 s2 => sound_state prog s1 /\ match_states n s1 s2). +- apply lt_wf. +- simpl; intros. exploit transf_initial_states; eauto. intros (n & st2 & A & B). + exists n, st2; intuition. eapply sound_initial; eauto. +- simpl; intros. destruct H. eapply transf_final_states; eauto. +- simpl; intros. destruct H0. + assert (sound_state prog s1') by (eapply sound_step; eauto). + fold ge; fold tge. + exploit transf_step_correct; eauto. + intros [ [n2 [s2' [A B]]] | [n2 [A [B C]]]]. + exists n2; exists s2'; split; auto. left; apply plus_one; auto. + exists n2; exists s2; split; auto. right; split; auto. subst t; apply star_refl. +- eexact symbols_preserved. Qed. End PRESERVATION. diff --git a/backend/Deadcode.v b/backend/Deadcode.v new file mode 100644 index 0000000..9efeca1 --- /dev/null +++ b/backend/Deadcode.v @@ -0,0 +1,192 @@ +(* *********************************************************************) +(* *) +(* 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. *) +(* *) +(* *********************************************************************) + +(** Elimination of unneeded computations over RTL. *) + +Require Import Coqlib. +Require Import Errors. +Require Import Maps. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Memory. +Require Import Registers. +Require Import Op. +Require Import RTL. +Require Import Lattice. +Require Import Kildall. +Require Import ValueDomain. +Require Import ValueAnalysis. +Require Import NeedDomain. +Require Import NeedOp. + +(** * Part 1: the static analysis *) + +Definition add_need (r: reg) (nv: nval) (ne: nenv) : nenv := + NE.set r (nlub nv (NE.get r ne)) ne. + +Fixpoint add_needs (rl: list reg) (nv: nval) (ne: nenv) : nenv := + match rl with + | nil => ne + | r1 :: rs => add_need r1 nv (add_needs rs nv ne) + end. + +Definition add_ros_need (ros: reg + ident) (ne: nenv) : nenv := + match ros with + | inl r => add_need r All ne + | inr s => ne + end. + +Definition add_opt_need (or: option reg) (ne: nenv) : nenv := + match or with + | Some r => add_need r All ne + | None => ne + end. + +Definition kill (r: reg) (ne: nenv) : nenv := NE.set r Nothing ne. + +Definition is_dead (v: nval) := + match v with Nothing => true | _ => false end. + +Definition is_int_zero (v: nval) := + match v with I n => Int.eq n Int.zero | _ => false end. + +Function transfer_builtin (app: VA.t) (ef: external_function) (args: list reg) (res: reg) + (ne: NE.t) (nm: nmem) : NA.t := + match ef, args with + | EF_vload chunk, a1::nil => + (add_needs args All (kill res ne), + nmem_add nm (aaddr app a1) (size_chunk chunk)) + | EF_vload_global chunk id ofs, nil => + (add_needs args All (kill res ne), + nmem_add nm (Gl id ofs) (size_chunk chunk)) + | EF_vstore chunk, a1::a2::nil => + (add_need a1 All (add_need a2 (store_argument chunk) (kill res ne)), nm) + | EF_vstore_global chunk id ofs, a1::nil => + (add_need a1 (store_argument chunk) (kill res ne), nm) + | EF_memcpy sz al, dst::src::nil => + if nmem_contains nm (aaddr app dst) sz then + (add_needs args All (kill res ne), + nmem_add (nmem_remove nm (aaddr app dst) sz) (aaddr app src) sz) + else (ne, nm) + | EF_annot txt targs, _ => + (add_needs args All (kill res ne), nm) + | EF_annot_val txt targ, _ => + (add_needs args All (kill res ne), nm) + | _, _ => + (add_needs args All (kill res ne), nmem_all) + end. + +Definition transfer (f: function) (approx: PMap.t VA.t) + (pc: node) (after: NA.t) : NA.t := + let (ne, nm) := after in + match f.(fn_code)!pc with + | None => + NA.bot + | Some (Inop s) => + after + | Some (Iop op args res s) => + let nres := nreg ne res in + if is_dead nres then after + else if is_int_zero nres then (kill res ne, nm) + else (add_needs args (needs_of_operation op nres) (kill res ne), nm) + | Some (Iload chunk addr args dst s) => + let ndst := nreg ne dst in + if is_dead ndst then after + else if is_int_zero ndst then (kill dst ne, nm) + else (add_needs args All (kill dst ne), + nmem_add nm (aaddressing approx!!pc addr args) (size_chunk chunk)) + | Some (Istore chunk addr args src s) => + let p := aaddressing approx!!pc addr args in + if nmem_contains nm p (size_chunk chunk) + then (add_needs args All (add_need src (store_argument chunk) ne), + nmem_remove nm p (size_chunk chunk)) + else after + | Some(Icall sig ros args res s) => + (add_needs args All (add_ros_need ros (kill res ne)), nmem_all) + | Some(Itailcall sig ros args) => + (add_needs args All (add_ros_need ros NE.bot), + nmem_dead_stack f.(fn_stacksize)) + | Some(Ibuiltin ef args res s) => + transfer_builtin approx!!pc ef args res ne nm + | Some(Icond cond args s1 s2) => + (add_needs args (needs_of_condition cond) ne, nm) + | Some(Ijumptable arg tbl) => + (add_need arg All ne, nm) + | Some(Ireturn optarg) => + (add_opt_need optarg ne, nmem_dead_stack f.(fn_stacksize)) + end. + +Module DS := Backward_Dataflow_Solver(NA)(NodeSetBackward). + +Definition analyze (approx: PMap.t VA.t) (f: function): option (PMap.t NA.t) := + DS.fixpoint f.(fn_code) successors_instr + (transfer f approx). + +(** * Part 2: the code transformation *) + +Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t) + (pc: node) (instr: instruction) := + match instr with + | Iop op args res s => + let nres := nreg (fst an!!pc) res in + if is_dead nres then + Inop s + else if is_int_zero nres then + Iop (Ointconst Int.zero) nil res s + else if operation_is_redundant op nres then + match args with arg :: _ => Iop Omove (arg :: nil) res s | nil => instr end + else + instr + | Iload chunk addr args dst s => + let ndst := nreg (fst an!!pc) dst in + if is_dead ndst then + Inop s + else if is_int_zero ndst then + Iop (Ointconst Int.zero) nil dst s + else + instr + | Istore chunk addr args src s => + let p := aaddressing approx!!pc addr args in + if nmem_contains (snd an!!pc) p (size_chunk chunk) + then instr + else Inop s + | Ibuiltin (EF_memcpy sz al) (dst :: src :: nil) res s => + if nmem_contains (snd an!!pc) (aaddr approx!!pc dst) sz + then instr + else Inop s + | _ => + instr + end. + +Definition vanalyze := ValueAnalysis.analyze. + +Definition transf_function (rm: romem) (f: function) : res function := + let approx := vanalyze rm f in + match analyze approx f with + | Some an => + OK {| fn_sig := f.(fn_sig); + fn_params := f.(fn_params); + fn_stacksize := f.(fn_stacksize); + fn_code := PTree.map (transf_instr approx an) f.(fn_code); + fn_entrypoint := f.(fn_entrypoint) |} + | None => + Error (msg "Neededness analysis failed") + end. + + +Definition transf_fundef (rm: romem) (fd: fundef) : res fundef := + AST.transf_partial_fundef (transf_function rm) fd. + +Definition transf_program (p: program) : res program := + transform_partial_program (transf_fundef (romem_for_program p)) p. + diff --git a/backend/Deadcodeproof.v b/backend/Deadcodeproof.v new file mode 100644 index 0000000..deb8628 --- /dev/null +++ b/backend/Deadcodeproof.v @@ -0,0 +1,1014 @@ +(* *********************************************************************) +(* *) +(* 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. *) +(* *) +(* *********************************************************************) + +(** Elimination of unneeded computations over RTL: correctness proof. *) + +Require Import Coqlib. +Require Import Errors. +Require Import Maps. +Require Import IntvSets. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Globalenvs. +Require Import Events. +Require Import Smallstep. +Require Import Op. +Require Import Registers. +Require Import RTL. +Require Import Lattice. +Require Import Kildall. +Require Import ValueDomain. +Require Import ValueAnalysis. +Require Import NeedDomain. +Require Import NeedOp. +Require Import Deadcode. + +(** * Relating the memory states *) + +(** The [magree] predicate is a variant of [Mem.extends] where we + allow the contents of the two memory states to differ arbitrarily + on some locations. The predicate [P] is true on the locations whose + contents must be in the [lessdef] relation. *) + +Definition locset := block -> Z -> Prop. + +Record magree (m1 m2: mem) (P: locset) : Prop := mk_magree { + ma_perm: + forall b ofs k p, + Mem.perm m1 b ofs k p -> + Mem.perm m2 b ofs k p; + ma_memval: + forall b ofs, + Mem.perm m1 b ofs Cur Readable -> + P b ofs -> + memval_lessdef (ZMap.get ofs (PMap.get b (Mem.mem_contents m1))) + (ZMap.get ofs (PMap.get b (Mem.mem_contents m2))); + ma_nextblock: + Mem.nextblock m2 = Mem.nextblock m1 +}. + +Lemma magree_monotone: + forall m1 m2 (P Q: locset), + magree m1 m2 P -> + (forall b ofs, Q b ofs -> P b ofs) -> + magree m1 m2 Q. +Proof. + intros. destruct H. constructor; auto. +Qed. + +Lemma mextends_agree: + forall m1 m2 P, Mem.extends m1 m2 -> magree m1 m2 P. +Proof. + intros. destruct H. destruct mext_inj. constructor; intros. +- replace ofs with (ofs + 0) by omega. eapply mi_perm; eauto. auto. +- exploit mi_memval; eauto. unfold inject_id; eauto. + rewrite Zplus_0_r. auto. +- auto. +Qed. + +Lemma magree_extends: + forall m1 m2 (P: locset), + (forall b ofs, P b ofs) -> + magree m1 m2 P -> Mem.extends m1 m2. +Proof. + intros. destruct H0. constructor; auto. constructor; unfold inject_id; intros. +- inv H0. rewrite Zplus_0_r. eauto. +- inv H0. apply Zdivide_0. +- inv H0. rewrite Zplus_0_r. eapply ma_memval0; eauto. +Qed. + +Lemma magree_loadbytes: + forall m1 m2 P b ofs n bytes, + magree m1 m2 P -> + Mem.loadbytes m1 b ofs n = Some bytes -> + (forall i, ofs <= i < ofs + n -> P b i) -> + exists bytes', Mem.loadbytes m2 b ofs n = Some bytes' /\ list_forall2 memval_lessdef bytes bytes'. +Proof. + assert (GETN: forall c1 c2 n ofs, + (forall i, ofs <= i < ofs + Z.of_nat n -> memval_lessdef (ZMap.get i c1) (ZMap.get i c2)) -> + list_forall2 memval_lessdef (Mem.getN n ofs c1) (Mem.getN n ofs c2)). + { + induction n; intros; simpl. + constructor. + rewrite inj_S in H. constructor. + apply H. omega. + apply IHn. intros; apply H; omega. + } +Local Transparent Mem.loadbytes. + unfold Mem.loadbytes; intros. destruct H. + destruct (Mem.range_perm_dec m1 b ofs (ofs + n) Cur Readable); inv H0. + rewrite pred_dec_true. econstructor; split; eauto. + apply GETN. intros. rewrite nat_of_Z_max in H. + assert (ofs <= i < ofs + n) by xomega. + apply ma_memval0; auto. + red; intros; eauto. +Qed. + +Lemma magree_load: + forall m1 m2 P chunk b ofs v, + magree m1 m2 P -> + Mem.load chunk m1 b ofs = Some v -> + (forall i, ofs <= i < ofs + size_chunk chunk -> P b i) -> + exists v', Mem.load chunk m2 b ofs = Some v' /\ Val.lessdef v v'. +Proof. + intros. exploit Mem.load_valid_access; eauto. intros [A B]. + exploit Mem.load_loadbytes; eauto. intros [bytes [C D]]. + exploit magree_loadbytes; eauto. intros [bytes' [E F]]. + exists (decode_val chunk bytes'); split. + apply Mem.loadbytes_load; auto. + apply val_inject_id. subst v. apply decode_val_inject; auto. +Qed. + +Lemma magree_storebytes_parallel: + forall m1 m2 (P Q: locset) b ofs bytes1 m1' bytes2, + magree m1 m2 P -> + Mem.storebytes m1 b ofs bytes1 = Some m1' -> + (forall b' i, Q b' i -> + b' <> b \/ i < ofs \/ ofs + Z_of_nat (length bytes1) <= i -> + P b' i) -> + list_forall2 memval_lessdef bytes1 bytes2 -> + exists m2', Mem.storebytes m2 b ofs bytes2 = Some m2' /\ magree m1' m2' Q. +Proof. + assert (SETN: forall (access: Z -> Prop) bytes1 bytes2, + list_forall2 memval_lessdef bytes1 bytes2 -> + forall p c1 c2, + (forall i, access i -> i < p \/ p + Z.of_nat (length bytes1) <= i -> memval_lessdef (ZMap.get i c1) (ZMap.get i c2)) -> + forall q, access q -> + memval_lessdef (ZMap.get q (Mem.setN bytes1 p c1)) + (ZMap.get q (Mem.setN bytes2 p c2))). + { + induction 1; intros; simpl. + - apply H; auto. simpl. omega. + - simpl length in H1; rewrite inj_S in H1. + apply IHlist_forall2; auto. + intros. rewrite ! ZMap.gsspec. destruct (ZIndexed.eq i p). auto. + apply H1; auto. unfold ZIndexed.t in *; omega. + } + intros. + destruct (Mem.range_perm_storebytes m2 b ofs bytes2) as [m2' ST2]. + { erewrite <- list_forall2_length by eauto. red; intros. + eapply ma_perm; eauto. + eapply Mem.storebytes_range_perm; eauto. } + exists m2'; split; auto. + constructor; intros. +- eapply Mem.perm_storebytes_1; eauto. eapply ma_perm; eauto. + eapply Mem.perm_storebytes_2; eauto. +- rewrite (Mem.storebytes_mem_contents _ _ _ _ _ H0). + rewrite (Mem.storebytes_mem_contents _ _ _ _ _ ST2). + rewrite ! PMap.gsspec. destruct (peq b0 b). ++ subst b0. apply SETN with (access := fun ofs => Mem.perm m1' b ofs Cur Readable /\ Q b ofs); auto. + intros. destruct H5. eapply ma_memval; eauto. + eapply Mem.perm_storebytes_2; eauto. + apply H1; auto. ++ eapply ma_memval; eauto. eapply Mem.perm_storebytes_2; eauto. apply H1; auto. +- rewrite (Mem.nextblock_storebytes _ _ _ _ _ H0). + rewrite (Mem.nextblock_storebytes _ _ _ _ _ ST2). + eapply ma_nextblock; eauto. +Qed. + +Lemma magree_store_parallel: + forall m1 m2 (P Q: locset) chunk b ofs v1 m1' v2, + magree m1 m2 P -> + Mem.store chunk m1 b ofs v1 = Some m1' -> + vagree v1 v2 (store_argument chunk) -> + (forall b' i, Q b' i -> + b' <> b \/ i < ofs \/ ofs + size_chunk chunk <= i -> + P b' i) -> + exists m2', Mem.store chunk m2 b ofs v2 = Some m2' /\ magree m1' m2' Q. +Proof. + intros. + exploit Mem.store_valid_access_3; eauto. intros [A B]. + exploit Mem.store_storebytes; eauto. intros SB1. + exploit magree_storebytes_parallel. eauto. eauto. + instantiate (1 := Q). intros. rewrite encode_val_length in H4. + rewrite <- size_chunk_conv in H4. apply H2; auto. + eapply store_argument_sound; eauto. + intros [m2' [SB2 AG]]. + exists m2'; split; auto. + apply Mem.storebytes_store; auto. +Qed. + +Lemma magree_storebytes_left: + forall m1 m2 P b ofs bytes1 m1', + magree m1 m2 P -> + Mem.storebytes m1 b ofs bytes1 = Some m1' -> + (forall i, ofs <= i < ofs + Z_of_nat (length bytes1) -> ~(P b i)) -> + magree m1' m2 P. +Proof. + intros. constructor; intros. +- eapply ma_perm; eauto. eapply Mem.perm_storebytes_2; eauto. +- rewrite (Mem.storebytes_mem_contents _ _ _ _ _ H0). + rewrite PMap.gsspec. destruct (peq b0 b). ++ subst b0. rewrite Mem.setN_outside. eapply ma_memval; eauto. eapply Mem.perm_storebytes_2; eauto. + destruct (zlt ofs0 ofs); auto. destruct (zle (ofs + Z.of_nat (length bytes1)) ofs0); try omega. + elim (H1 ofs0). omega. auto. ++ eapply ma_memval; eauto. eapply Mem.perm_storebytes_2; eauto. +- rewrite (Mem.nextblock_storebytes _ _ _ _ _ H0). + eapply ma_nextblock; eauto. +Qed. + +Lemma magree_store_left: + forall m1 m2 P chunk b ofs v1 m1', + magree m1 m2 P -> + Mem.store chunk m1 b ofs v1 = Some m1' -> + (forall i, ofs <= i < ofs + size_chunk chunk -> ~(P b i)) -> + magree m1' m2 P. +Proof. + intros. eapply magree_storebytes_left; eauto. + eapply Mem.store_storebytes; eauto. + intros. rewrite encode_val_length in H2. + rewrite <- size_chunk_conv in H2. apply H1; auto. +Qed. + +Lemma magree_free: + forall m1 m2 (P Q: locset) b lo hi m1', + magree m1 m2 P -> + Mem.free m1 b lo hi = Some m1' -> + (forall b' i, Q b' i -> + b' <> b \/ ~(lo <= i < hi) -> + P b' i) -> + exists m2', Mem.free m2 b lo hi = Some m2' /\ magree m1' m2' Q. +Proof. + intros. + destruct (Mem.range_perm_free m2 b lo hi) as [m2' FREE]. + red; intros. eapply ma_perm; eauto. eapply Mem.free_range_perm; eauto. + exists m2'; split; auto. + constructor; intros. +- (* permissions *) + assert (Mem.perm m2 b0 ofs k p). { eapply ma_perm; eauto. eapply Mem.perm_free_3; eauto. } + exploit Mem.perm_free_inv; eauto. intros [[A B] | A]; auto. + subst b0. eelim Mem.perm_free_2. eexact H0. eauto. eauto. +- (* contents *) + rewrite (Mem.free_result _ _ _ _ _ H0). + rewrite (Mem.free_result _ _ _ _ _ FREE). + simpl. eapply ma_memval; eauto. eapply Mem.perm_free_3; eauto. + apply H1; auto. destruct (eq_block b0 b); auto. + subst b0. right. red; intros. eelim Mem.perm_free_2. eexact H0. eauto. eauto. +- (* nextblock *) + rewrite (Mem.free_result _ _ _ _ _ H0). + rewrite (Mem.free_result _ _ _ _ _ FREE). + simpl. eapply ma_nextblock; eauto. +Qed. + +(** * Properties of the need environment *) + +Lemma add_need_ge: + forall r nv ne, + nge (NE.get r (add_need r nv ne)) nv /\ NE.ge (add_need r nv ne) ne. +Proof. + intros. unfold add_need. split. + rewrite NE.gsspec; rewrite peq_true. apply nge_lub_l. + red. intros. rewrite NE.gsspec. destruct (peq p r). + subst. apply NVal.ge_lub_right. + apply NVal.ge_refl. apply NVal.eq_refl. +Qed. + +Lemma add_needs_ge: + forall rl nv ne, + (forall r, In r rl -> nge (NE.get r (add_needs rl nv ne)) nv) + /\ NE.ge (add_needs rl nv ne) ne. +Proof. + induction rl; simpl; intros. + split. tauto. apply NE.ge_refl. apply NE.eq_refl. + destruct (IHrl nv ne) as [A B]. + split; intros. + destruct H. subst a. apply add_need_ge. + apply nge_trans with (NE.get r (add_needs rl nv ne)). + apply add_need_ge. apply A; auto. + eapply NE.ge_trans; eauto. apply add_need_ge. +Qed. + +Lemma add_need_eagree: + forall e e' r nv ne, eagree e e' (add_need r nv ne) -> eagree e e' ne. +Proof. + intros. eapply eagree_ge; eauto. apply add_need_ge. +Qed. + +Lemma add_need_vagree: + forall e e' r nv ne, eagree e e' (add_need r nv ne) -> vagree e#r e'#r nv. +Proof. + intros. eapply nge_agree. eapply add_need_ge. apply H. +Qed. + +Lemma add_needs_eagree: + forall nv rl e e' ne, eagree e e' (add_needs rl nv ne) -> eagree e e' ne. +Proof. + intros. eapply eagree_ge; eauto. apply add_needs_ge. +Qed. + +Lemma add_needs_vagree: + forall nv rl e e' ne, eagree e e' (add_needs rl nv ne) -> vagree_list e##rl e'##rl nv. +Proof. + intros. destruct (add_needs_ge rl nv ne) as [A B]. + set (ne' := add_needs rl nv ne) in *. + revert A; generalize rl. induction rl0; simpl; intros. + constructor. + constructor. eapply nge_agree; eauto. apply IHrl0. auto. +Qed. + +Lemma add_need_lessdef: + forall e e' r ne, eagree e e' (add_need r All ne) -> Val.lessdef e#r e'#r. +Proof. + intros. apply lessdef_vagree. eapply add_need_vagree; eauto. +Qed. + +Lemma add_needs_lessdef: + forall e e' rl ne, eagree e e' (add_needs rl All ne) -> Val.lessdef_list e##rl e'##rl. +Proof. + intros. exploit add_needs_vagree; eauto. + generalize rl. induction rl0; simpl; intros V; inv V. + constructor. + constructor; auto. +Qed. + +Lemma add_ros_need_eagree: + forall e e' ros ne, eagree e e' (add_ros_need ros ne) -> eagree e e' ne. +Proof. + intros. destruct ros; simpl in *. eapply add_need_eagree; eauto. auto. +Qed. + +Hint Resolve add_need_eagree add_need_vagree add_need_lessdef + add_needs_eagree add_needs_vagree add_needs_lessdef + add_ros_need_eagree: na. + +Lemma eagree_init_regs: + forall rl vl1 vl2 ne, + Val.lessdef_list vl1 vl2 -> + eagree (init_regs vl1 rl) (init_regs vl2 rl) ne. +Proof. + induction rl; intros until ne; intros LD; simpl. +- red; auto with na. +- inv LD. + + red; auto with na. + + apply eagree_update; auto with na. +Qed. + +(** * Basic properties of the translation *) + +Section PRESERVATION. + +Variable prog: program. +Variable tprog: program. +Hypothesis TRANSF: transf_program prog = OK tprog. +Let ge := Genv.globalenv prog. +Let tge := Genv.globalenv tprog. +Let rm := romem_for_program prog. + +Lemma symbols_preserved: + forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s. +Proof. + intro. unfold ge, tge. + apply Genv.find_symbol_transf_partial with (transf_fundef rm). + exact TRANSF. +Qed. + +Lemma varinfo_preserved: + forall b, Genv.find_var_info tge b = Genv.find_var_info ge b. +Proof. + intro. unfold ge, tge. + apply Genv.find_var_info_transf_partial with (transf_fundef rm). + exact TRANSF. +Qed. + +Lemma functions_translated: + forall (v: val) (f: RTL.fundef), + Genv.find_funct ge v = Some f -> + exists tf, + Genv.find_funct tge v = Some tf /\ transf_fundef rm f = OK tf. +Proof (Genv.find_funct_transf_partial (transf_fundef rm) _ TRANSF). + +Lemma function_ptr_translated: + forall (b: block) (f: RTL.fundef), + Genv.find_funct_ptr ge b = Some f -> + exists tf, + Genv.find_funct_ptr tge b = Some tf /\ transf_fundef rm f = OK tf. +Proof (Genv.find_funct_ptr_transf_partial (transf_fundef rm) _ TRANSF). + +Lemma sig_function_translated: + forall f tf, + transf_fundef rm f = OK tf -> + funsig tf = funsig f. +Proof. + intros; destruct f; monadInv H. + unfold transf_function in EQ. + destruct (analyze (vanalyze rm f) f); inv EQ; auto. + auto. +Qed. + +Lemma stacksize_translated: + forall f tf, + transf_function rm f = OK tf -> tf.(fn_stacksize) = f.(fn_stacksize). +Proof. + unfold transf_function; intros. destruct (analyze (vanalyze rm f) f); inv H; auto. +Qed. + +Lemma transf_function_at: + forall f tf an pc instr, + transf_function rm f = OK tf -> + analyze (vanalyze rm f) f = Some an -> + f.(fn_code)!pc = Some instr -> + tf.(fn_code)!pc = Some(transf_instr (vanalyze rm f) an pc instr). +Proof. + intros. unfold transf_function in H. rewrite H0 in H. inv H; simpl. + rewrite PTree.gmap. rewrite H1; auto. +Qed. + +Lemma is_dead_sound_1: + forall nv, is_dead nv = true -> nv = Nothing. +Proof. + destruct nv; simpl; congruence. +Qed. + +Lemma is_dead_sound_2: + forall nv, is_dead nv = false -> nv <> Nothing. +Proof. + intros; red; intros. subst nv; discriminate. +Qed. + +Hint Resolve is_dead_sound_1 is_dead_sound_2: na. + +Lemma is_int_zero_sound: + forall nv, is_int_zero nv = true -> nv = I Int.zero. +Proof. + unfold is_int_zero; destruct nv; try discriminate. + predSpec Int.eq Int.eq_spec m Int.zero; congruence. +Qed. + +Lemma find_function_translated: + forall ros rs fd trs ne, + find_function ge ros rs = Some fd -> + eagree rs trs (add_ros_need ros ne) -> + exists tfd, find_function tge ros trs = Some tfd /\ transf_fundef rm fd = OK tfd. +Proof. + intros. destruct ros as [r|id]; simpl in *. +- assert (LD: Val.lessdef rs#r trs#r) by eauto with na. inv LD. + apply functions_translated; auto. + rewrite <- H2 in H; discriminate. +- rewrite symbols_preserved. destruct (Genv.find_symbol ge id); try discriminate. + apply function_ptr_translated; auto. +Qed. + +(** * Semantic invariant *) + +Inductive match_stackframes: stackframe -> stackframe -> Prop := + | match_stackframes_intro: + forall res f sp pc e tf te an + (FUN: transf_function rm f = OK tf) + (ANL: analyze (vanalyze rm f) f = Some an) + (RES: forall v tv, + Val.lessdef v tv -> + eagree (e#res <- v) (te#res<- tv) + (fst (transfer f (vanalyze rm f) pc an!!pc))), + match_stackframes (Stackframe res f (Vptr sp Int.zero) pc e) + (Stackframe res tf (Vptr sp Int.zero) pc te). + +Inductive match_states: state -> state -> Prop := + | match_regular_states: + forall s f sp pc e m ts tf te tm an + (STACKS: list_forall2 match_stackframes s ts) + (FUN: transf_function rm f = OK tf) + (ANL: analyze (vanalyze rm f) f = Some an) + (ENV: eagree e te (fst (transfer f (vanalyze rm f) pc an!!pc))) + (MEM: magree m tm (nlive ge sp (snd (transfer f (vanalyze rm f) pc an!!pc)))), + match_states (State s f (Vptr sp Int.zero) pc e m) + (State ts tf (Vptr sp Int.zero) pc te tm) + | match_call_states: + forall s f args m ts tf targs tm + (STACKS: list_forall2 match_stackframes s ts) + (FUN: transf_fundef rm f = OK tf) + (ARGS: Val.lessdef_list args targs) + (MEM: Mem.extends m tm), + match_states (Callstate s f args m) + (Callstate ts tf targs tm) + | match_return_states: + forall s v m ts tv tm + (STACKS: list_forall2 match_stackframes s ts) + (RES: Val.lessdef v tv) + (MEM: Mem.extends m tm), + match_states (Returnstate s v m) + (Returnstate ts tv tm). + +(** [match_states] and CFG successors *) + +Lemma analyze_successors: + forall f an pc instr pc', + analyze (vanalyze rm f) f = Some an -> + f.(fn_code)!pc = Some instr -> + In pc' (successors_instr instr) -> + NA.ge an!!pc (transfer f (vanalyze rm f) pc' an!!pc'). +Proof. + intros. eapply DS.fixpoint_solution; eauto. + intros. unfold transfer; rewrite H2. destruct a. apply DS.L.eq_refl. +Qed. + +Lemma match_succ_states: + forall s f sp pc e m ts tf te tm an pc' instr ne nm + (STACKS: list_forall2 match_stackframes s ts) + (FUN: transf_function rm f = OK tf) + (ANL: analyze (vanalyze rm f) f = Some an) + (INSTR: f.(fn_code)!pc = Some instr) + (SUCC: In pc' (successors_instr instr)) + (ANPC: an!!pc = (ne, nm)) + (ENV: eagree e te ne) + (MEM: magree m tm (nlive ge sp nm)), + match_states (State s f (Vptr sp Int.zero) pc' e m) + (State ts tf (Vptr sp Int.zero) pc' te tm). +Proof. + intros. exploit analyze_successors; eauto. rewrite ANPC; simpl. intros [A B]. + econstructor; eauto. + eapply eagree_ge; eauto. + eapply magree_monotone; eauto. intros; apply B; auto. +Qed. + +(** Properties of volatile memory accesses *) + +(* +Lemma transf_volatile_load: + + forall s f sp pc e m te r tm nm chunk t v m', + + volatile_load_sem chunk ge (addr :: nil) m t v m' -> + Val.lessdef addr taddr -> + genv_match bc ge -> + bc sp = BCstack -> + pmatch + + sound_state prog (State s f (Vptr sp Int.zero) pc e m) -> + Val.lessdef e#r te#r -> + magree m tm + (nlive ge sp + (nmem_add nm (aaddr (vanalyze rm f) # pc r) (size_chunk chunk))) -> + m' = m /\ + exists tv, volatile_load_sem chunk ge (te#r :: nil) tm t tv tm /\ Val.lessdef v tv. +Proof. + intros. inv H2. split; auto. rewrite <- H3 in H0; inv H0. inv H4. +- (* volatile *) + exists (Val.load_result chunk v0); split; auto. + constructor. constructor; auto. +- (* not volatile *) + exploit magree_load; eauto. + exploit aaddr_sound; eauto. intros (bc & P & Q & R). + intros. eapply nlive_add; eauto. + intros (v' & P & Q). + exists v'; split. constructor. econstructor; eauto. auto. +Qed. +*) + +Lemma transf_volatile_store: + forall v1 v2 v1' v2' m tm chunk sp nm t v m', + volatile_store_sem chunk ge (v1::v2::nil) m t v m' -> + Val.lessdef v1 v1' -> + vagree v2 v2' (store_argument chunk) -> + magree m tm (nlive ge sp nm) -> + v = Vundef /\ + exists tm', volatile_store_sem chunk ge (v1'::v2'::nil) tm t Vundef tm' + /\ magree m' tm' (nlive ge sp nm). +Proof. + intros. inv H. split; auto. + inv H0. inv H9. +- (* volatile *) + exists tm; split; auto. econstructor. econstructor; eauto. + eapply eventval_match_lessdef; eauto. apply store_argument_load_result; auto. +- (* not volatile *) + exploit magree_store_parallel. eauto. eauto. eauto. + instantiate (1 := nlive ge sp nm). auto. + intros (tm' & P & Q). + exists tm'; split. econstructor. econstructor; eauto. auto. +Qed. + +Lemma eagree_set_undef: + forall e1 e2 ne r, eagree e1 e2 ne -> eagree (e1#r <- Vundef) e2 ne. +Proof. + intros; red; intros. rewrite PMap.gsspec. destruct (peq r0 r); auto with na. +Qed. + +(** * The simulation diagram *) + +Theorem step_simulation: + forall S1 t S2, step ge S1 t S2 -> + forall S1', match_states S1 S1' -> sound_state prog S1 -> + exists S2', step tge S1' t S2' /\ match_states S2 S2'. +Proof. + +Ltac TransfInstr := + match goal with + | [INSTR: (fn_code _)!_ = Some _, + FUN: transf_function _ _ = OK _, + ANL: analyze _ _ = Some _ |- _ ] => + generalize (transf_function_at _ _ _ _ _ FUN ANL INSTR); + intro TI; + unfold transf_instr in TI + end. + +Ltac UseTransfer := + match goal with + | [INSTR: (fn_code _)!?pc = Some _, + ANL: analyze _ _ = Some ?an |- _ ] => + destruct (an!!pc) as [ne nm] eqn:ANPC; + unfold transfer in *; + rewrite INSTR in *; + simpl in * + end. + + induction 1; intros S1' MS SS; inv MS. + +- (* nop *) + TransfInstr; UseTransfer. + econstructor; split. + eapply exec_Inop; eauto. + eapply match_succ_states; eauto. simpl; auto. + +- (* op *) + TransfInstr; UseTransfer. + destruct (is_dead (nreg ne res)) eqn:DEAD; + [idtac|destruct (is_int_zero (nreg ne res)) eqn:INTZERO; + [idtac|destruct (operation_is_redundant op (nreg ne res)) eqn:REDUNDANT]]. ++ (* dead instruction, turned into a nop *) + econstructor; split. + eapply exec_Inop; eauto. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update_dead; auto with na. ++ (* instruction with needs = [I Int.zero], turned into a load immediate of zero. *) + econstructor; split. + eapply exec_Iop with (v := Vint Int.zero); eauto. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; auto. + rewrite is_int_zero_sound by auto. + destruct v; simpl; auto. apply iagree_zero. ++ (* redundant operation *) + destruct args. + * (* kept as is because no arguments -- should never happen *) + simpl in *. + exploit needs_of_operation_sound. eapply ma_perm; eauto. + eauto. instantiate (1 := nreg ne res). eauto with na. eauto with na. intros [tv [A B]]. + econstructor; split. + eapply exec_Iop with (v := tv); eauto. + rewrite <- A. apply eval_operation_preserved. exact symbols_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; auto. + * (* turned into a move *) + simpl in *. + exploit operation_is_redundant_sound. eauto. eauto. eapply add_need_vagree. eauto. + intros VA. + econstructor; split. + eapply exec_Iop; eauto. simpl; reflexivity. + eapply match_succ_states; eauto. simpl; auto. + eapply eagree_update; eauto with na. ++ (* preserved operation *) + simpl in *. + exploit needs_of_operation_sound. eapply ma_perm; eauto. eauto. eauto with na. eauto with na. + intros [tv [A B]]. + econstructor; split. + eapply exec_Iop with (v := tv); eauto. + rewrite <- A. apply eval_operation_preserved. exact symbols_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. + +- (* load *) + TransfInstr; UseTransfer. + destruct (is_dead (nreg ne dst)) eqn:DEAD; + [idtac|destruct (is_int_zero (nreg ne dst)) eqn:INTZERO]; + simpl in *. ++ (* dead instruction, turned into a nop *) + econstructor; split. + eapply exec_Inop; eauto. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update_dead; auto with na. ++ (* instruction with needs = [I Int.zero], turned into a load immediate of zero. *) + econstructor; split. + eapply exec_Iop with (v := Vint Int.zero); eauto. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; auto. + rewrite is_int_zero_sound by auto. + destruct v; simpl; auto. apply iagree_zero. ++ (* preserved *) + exploit eval_addressing_lessdef. eapply add_needs_lessdef; eauto. eauto. + intros (ta & U & V). inv V; try discriminate. + destruct ta; simpl in H1; try discriminate. + exploit magree_load; eauto. + exploit aaddressing_sound; eauto. intros (bc & A & B & C). + intros. apply nlive_add with bc i; assumption. + intros (tv & P & Q). + econstructor; split. + eapply exec_Iload with (a := Vptr b i); eauto. + rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. + eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto. + +- (* store *) + TransfInstr; UseTransfer. + destruct (nmem_contains nm (aaddressing (vanalyze rm f) # pc addr args) + (size_chunk chunk)) eqn:CONTAINS. ++ (* preserved *) + simpl in *. + exploit eval_addressing_lessdef. eapply add_needs_lessdef; eauto. eauto. + intros (ta & U & V). inv V; try discriminate. + destruct ta; simpl in H1; try discriminate. + exploit magree_store_parallel. eauto. eauto. instantiate (1 := te#src). eauto with na. + instantiate (1 := nlive ge sp0 nm). + exploit aaddressing_sound; eauto. intros (bc & A & B & C). + intros. apply nlive_remove with bc b i; assumption. + intros (tm' & P & Q). + econstructor; split. + eapply exec_Istore with (a := Vptr b i); eauto. + rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. + eapply match_succ_states; eauto. simpl; auto. + eauto with na. ++ (* dead instruction, turned into a nop *) + destruct a; simpl in H1; try discriminate. + econstructor; split. + eapply exec_Inop; eauto. + eapply match_succ_states; eauto. simpl; auto. + eapply magree_store_left; eauto. + exploit aaddressing_sound; eauto. intros (bc & A & B & C). + intros. eapply nlive_contains; eauto. + +- (* call *) + TransfInstr; UseTransfer. + exploit find_function_translated; eauto with na. intros (tfd & A & B). + econstructor; split. + eapply exec_Icall; eauto. apply sig_function_translated; auto. + constructor. + constructor; auto. econstructor; eauto. + intros. + edestruct analyze_successors; eauto. simpl; eauto. + eapply eagree_ge; eauto. rewrite ANPC. simpl. + apply eagree_update; eauto with na. + auto. eauto with na. eapply magree_extends; eauto. apply nlive_all. + +- (* tailcall *) + TransfInstr; UseTransfer. + exploit find_function_translated; eauto with na. intros (tfd & A & B). + exploit magree_free. eauto. eauto. instantiate (1 := nlive ge stk nmem_all). + intros; eapply nlive_dead_stack; eauto. + intros (tm' & C & D). + econstructor; split. + eapply exec_Itailcall; eauto. apply sig_function_translated; auto. + erewrite stacksize_translated by eauto. eexact C. + constructor; eauto with na. eapply magree_extends; eauto. apply nlive_all. + +- (* builtin *) + TransfInstr; UseTransfer. revert ENV MEM TI. + functional induction (transfer_builtin (vanalyze rm f)#pc ef args res ne nm); + simpl in *; intros. ++ (* volatile load *) + assert (LD: Val.lessdef rs#a1 te#a1) by eauto with na. + inv H0. rewrite <- H1 in LD; inv LD. + assert (X: exists tv, volatile_load ge chunk tm b ofs t tv /\ Val.lessdef v tv). + { + inv H2. + * exists (Val.load_result chunk v0); split; auto. constructor; auto. + * exploit magree_load; eauto. + exploit aaddr_sound; eauto. intros (bc & A & B & C). + intros. eapply nlive_add; eassumption. + intros (tv & P & Q). + exists tv; split; auto. constructor; auto. + } + destruct X as (tv & A & B). + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. + simpl. rewrite <- H4. constructor. eauto. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. + eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto. ++ (* volatile global load *) + inv H0. + assert (X: exists tv, volatile_load ge chunk tm b ofs t tv /\ Val.lessdef v tv). + { + inv H2. + * exists (Val.load_result chunk v0); split; auto. constructor; auto. + * exploit magree_load; eauto. + inv SS. intros. eapply nlive_add; eauto. constructor. apply GE. auto. + intros (tv & P & Q). + exists tv; split; auto. constructor; auto. + } + destruct X as (tv & A & B). + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. + simpl. econstructor; eauto. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. + eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto. ++ (* volatile store *) + exploit transf_volatile_store. eauto. + instantiate (1 := te#a1). eauto with na. + instantiate (1 := te#a2). eauto with na. + eauto. + intros (EQ & tm' & A & B). subst v. + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. simpl; eauto. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. ++ (* volatile global store *) + rewrite volatile_store_global_charact in H0. destruct H0 as (b & P & Q). + exploit transf_volatile_store. eauto. eauto. + instantiate (1 := te#a1). eauto with na. + eauto. + intros (EQ & tm' & A & B). subst v. + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. simpl. + rewrite volatile_store_global_charact. exists b; split; eauto. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. ++ (* memcpy *) + rewrite e1 in TI. + inv H0. + set (adst := aaddr (vanalyze rm f) # pc dst) in *. + set (asrc := aaddr (vanalyze rm f) # pc src) in *. + exploit magree_loadbytes. eauto. eauto. + exploit aaddr_sound. eauto. symmetry; eexact H2. + intros (bc & A & B & C). + intros. eapply nlive_add; eassumption. + intros (tbytes & P & Q). + exploit magree_storebytes_parallel. + eapply magree_monotone. eexact MEM. + instantiate (1 := nlive ge sp0 (nmem_remove nm adst sz)). + intros. apply incl_nmem_add; auto. + eauto. + instantiate (1 := nlive ge sp0 nm). + exploit aaddr_sound. eauto. symmetry; eexact H1. + intros (bc & A & B & C). + intros. eapply nlive_remove; eauto. + erewrite Mem.loadbytes_length in H10 by eauto. + rewrite nat_of_Z_eq in H10 by omega. auto. + eauto. + intros (tm' & A & B). + assert (LD1: Val.lessdef rs#src te#src) by eauto with na. rewrite <- H2 in LD1. + assert (LD2: Val.lessdef rs#dst te#dst) by eauto with na. rewrite <- H1 in LD2. + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. simpl. + inv LD1. inv LD2. econstructor; eauto. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. ++ (* memcpy eliminated *) + rewrite e1 in TI. inv H0. + set (adst := aaddr (vanalyze rm f) # pc dst) in *. + set (asrc := aaddr (vanalyze rm f) # pc src) in *. + econstructor; split. + eapply exec_Inop; eauto. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_set_undef; auto. + eapply magree_storebytes_left; eauto. + exploit aaddr_sound. eauto. symmetry; eexact H1. + intros (bc & A & B & C). + intros. eapply nlive_contains; eauto. + erewrite Mem.loadbytes_length in H0 by eauto. + rewrite nat_of_Z_eq in H0 by omega. auto. ++ (* annot *) + inv H0. + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. simpl; constructor. + eapply eventval_list_match_lessdef; eauto with na. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. ++ (* annot val *) + inv H0. destruct _x; inv H1. destruct _x; inv H4. + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. simpl; constructor. + eapply eventval_match_lessdef; eauto with na. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. ++ (* all other builtins *) + assert ((fn_code tf)!pc = Some(Ibuiltin _x _x0 res pc')). + { + destruct _x; auto. destruct _x0; auto. destruct _x0; auto. destruct _x0; auto. contradiction. + } + clear y TI. + exploit external_call_mem_extends; eauto with na. + eapply magree_extends; eauto. intros. apply nlive_all. + intros (v' & tm' & A & B & C & D & E). + econstructor; split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved. eauto. + exact symbols_preserved. exact varinfo_preserved. + eapply match_succ_states; eauto. simpl; auto. + apply eagree_update; eauto with na. + eapply mextends_agree; eauto. + +- (* conditional *) + TransfInstr; UseTransfer. + econstructor; split. + eapply exec_Icond; eauto. + eapply needs_of_condition_sound. eapply ma_perm; eauto. eauto. eauto with na. + eapply match_succ_states; eauto with na. + simpl; destruct b; auto. + +- (* jumptable *) + TransfInstr; UseTransfer. + assert (LD: Val.lessdef rs#arg te#arg) by eauto with na. + rewrite H0 in LD. inv LD. + econstructor; split. + eapply exec_Ijumptable; eauto. + eapply match_succ_states; eauto with na. + simpl. eapply list_nth_z_in; eauto. + +- (* return *) + TransfInstr; UseTransfer. + exploit magree_free. eauto. eauto. instantiate (1 := nlive ge stk nmem_all). + intros; eapply nlive_dead_stack; eauto. + intros (tm' & A & B). + econstructor; split. + eapply exec_Ireturn; eauto. + erewrite stacksize_translated by eauto. eexact A. + constructor; auto. + destruct or; simpl; eauto with na. + eapply magree_extends; eauto. apply nlive_all. + +- (* internal function *) + monadInv FUN. generalize EQ. unfold transf_function. intros EQ'. + destruct (analyze (vanalyze rm f) f) as [an|] eqn:AN; inv EQ'. + exploit Mem.alloc_extends; eauto. apply Zle_refl. apply Zle_refl. + intros (tm' & A & B). + econstructor; split. + econstructor; simpl; eauto. + simpl. econstructor; eauto. + apply eagree_init_regs; auto. + apply mextends_agree; auto. + +- (* external function *) + exploit external_call_mem_extends; eauto. + intros (res' & tm' & A & B & C & D & E). + simpl in FUN. inv FUN. + econstructor; split. + econstructor; eauto. + eapply external_call_symbols_preserved; eauto. + exact symbols_preserved. exact varinfo_preserved. + econstructor; eauto. + +- (* return *) + inv STACKS. inv H1. + econstructor; split. + constructor. + econstructor; eauto. apply mextends_agree; auto. +Qed. + +Lemma transf_initial_states: + forall st1, initial_state prog st1 -> + exists st2, initial_state tprog st2 /\ match_states st1 st2. +Proof. + intros. inversion H. + exploit function_ptr_translated; eauto. intros (tf & A & B). + exists (Callstate nil tf nil m0); split. + econstructor; eauto. + eapply Genv.init_mem_transf_partial; eauto. + rewrite (transform_partial_program_main _ _ TRANSF). + rewrite symbols_preserved. eauto. + rewrite <- H3. apply sig_function_translated; auto. + constructor. constructor. auto. constructor. apply Mem.extends_refl. +Qed. + +Lemma transf_final_states: + forall st1 st2 r, + match_states st1 st2 -> final_state st1 r -> final_state st2 r. +Proof. + intros. inv H0. inv H. inv STACKS. inv RES. constructor. +Qed. + +(** * Semantic preservation *) + +Theorem transf_program_correct: + forward_simulation (RTL.semantics prog) (RTL.semantics tprog). +Proof. + intros. + apply forward_simulation_step with + (match_states := fun s1 s2 => sound_state prog s1 /\ match_states s1 s2). +- exact symbols_preserved. +- simpl; intros. exploit transf_initial_states; eauto. intros [st2 [A B]]. + exists st2; intuition. eapply sound_initial; eauto. +- simpl; intros. destruct H. eapply transf_final_states; eauto. +- simpl; intros. destruct H0. + assert (sound_state prog s1') by (eapply sound_step; eauto). + fold ge; fold tge. exploit step_simulation; eauto. intros [st2' [A B]]. + exists st2'; auto. +Qed. + +End PRESERVATION. + + diff --git a/backend/Kildall.v b/backend/Kildall.v index a071f30..0d414d2 100644 --- a/backend/Kildall.v +++ b/backend/Kildall.v @@ -25,7 +25,7 @@ Local Unset Case Analysis Schemes. - [X(s) >= transf n X(n)] if program point [s] is a successor of program point [n] - [X(n) >= a] - if [(n, a)] belongs to a given list of (program points, approximations). + if [n] is an entry point and [a] its minimal approximation. The unknowns are the [X(n)], indexed by program points (e.g. nodes in the CFG graph of a RTL function). They range over a given ordered set that @@ -39,7 +39,7 @@ Symmetrically, a backward dataflow problem is a set of inequations of the form - [X(n) >= transf s X(s)] if program point [s] is a successor of program point [n] - [X(n) >= a] - if [(n, a)] belongs to a given list of (program points, approximations). + if [n] is an entry point and [a] its minimal approximation. The only difference with a forward dataflow problem is that the transfer function [transf] now computes the approximation before a program point [s] @@ -59,11 +59,6 @@ approximations do not exist or are too expensive to compute. *) (** * Solving forward dataflow problems using Kildall's algorithm *) -Definition successors_list (successors: PTree.t (list positive)) (pc: positive) : list positive := - match successors!pc with None => nil | Some l => l end. - -Notation "a !!! b" := (successors_list a b) (at level 1). - (** A forward dataflow solver has the following generic interface. Unknowns range over the type [L.t], which is equipped with semi-lattice operations (see file [Lattice]). *) @@ -72,53 +67,52 @@ Module Type DATAFLOW_SOLVER. Declare Module L: SEMILATTICE. + (** [fixpoint successors transf ep ev] is the solver. + It returns either an error or a mapping from program points to + values of type [L.t] representing the solution. [successors] + is a finite map returning the list of successors of the given program + point. [transf] is the transfer function, [ep] the entry point, + and [ev] the minimal abstract value for [ep]. *) + Variable fixpoint: forall {A: Type} (code: PTree.t A) (successors: A -> list positive) (transf: positive -> L.t -> L.t) - (entrypoints: list (positive * L.t)), + (ep: positive) (ev: L.t), option (PMap.t L.t). - (** [fixpoint successors transf entrypoints] is the solver. - It returns either an error or a mapping from program points to - values of type [L.t] representing the solution. [successors] - is a finite map returning the list of successors of the given program - point. [transf] is the transfer function, and [entrypoints] the additional - constraints imposed on the solution. *) + (** The [fixpoint_solution] theorem shows that the returned solution, + if any, satisfies the dataflow inequations. *) Hypothesis fixpoint_solution: - forall A (code: PTree.t A) successors transf entrypoints res n instr s, - fixpoint code successors transf entrypoints = Some res -> + forall A (code: PTree.t A) successors transf ep ev res n instr s, + fixpoint code successors transf ep ev = Some res -> code!n = Some instr -> In s (successors instr) -> + (forall n, L.eq (transf n L.bot) L.bot) -> L.ge res!!s (transf n res!!n). - (** The [fixpoint_solution] theorem shows that the returned solution, - if any, satisfies the dataflow inequations. *) + (** The [fixpoint_entry] theorem shows that the returned solution, + if any, satisfies the additional constraint over the entry point. *) Hypothesis fixpoint_entry: - forall A (code: PTree.t A) successors transf entrypoints res n v, - fixpoint code successors transf entrypoints = Some res -> - In (n, v) entrypoints -> - L.ge res!!n v. + forall A (code: PTree.t A) successors transf ep ev res, + fixpoint code successors transf ep ev = Some res -> + L.ge res!!ep ev. - (** The [fixpoint_entry] theorem shows that the returned solution, - if any, satisfies the additional constraints expressed - by [entrypoints]. *) + (** Finally, any property that is preserved by [L.lub] and [transf] + and that holds for [L.bot] also holds for the results of + the analysis. *) Hypothesis fixpoint_invariant: - forall A (code: PTree.t A) successors transf entrypoints + forall A (code: PTree.t A) successors transf ep ev (P: L.t -> Prop), P L.bot -> (forall x y, P x -> P y -> P (L.lub x y)) -> (forall pc instr x, code!pc = Some instr -> P x -> P (transf pc x)) -> - (forall n v, In (n, v) entrypoints -> P v) -> + P ev -> forall res pc, - fixpoint code successors transf entrypoints = Some res -> + fixpoint code successors transf ep ev = Some res -> P res!!pc. - (** Finally, any property that is preserved by [L.lub] and [transf] - and that holds for [L.bot] also holds for the results of - the analysis. *) - End DATAFLOW_SOLVER. (** Kildall's algorithm manipulates worklists, which are sets of CFG nodes @@ -131,11 +125,14 @@ End DATAFLOW_SOLVER. Module Type NODE_SET. Variable t: Type. + Variable empty: t. Variable add: positive -> t -> t. Variable pick: t -> option (positive * t). - Variable initial: forall {A: Type}, PTree.t A -> t. + Variable all_nodes: forall {A: Type}, PTree.t A -> t. Variable In: positive -> t -> Prop. + Hypothesis empty_spec: + forall n, ~In n empty. Hypothesis add_spec: forall n n' s, In n' (add n s) <-> n = n' \/ In n' s. Hypothesis pick_none: @@ -143,16 +140,48 @@ Module Type NODE_SET. Hypothesis pick_some: forall s n s', pick s = Some(n, s') -> forall n', In n' s <-> n = n' \/ In n' s'. - Hypothesis initial_spec: + Hypothesis all_nodes_spec: forall A (code: PTree.t A) n instr, - code!n = Some instr -> In n (initial code). + code!n = Some instr -> In n (all_nodes code). End NODE_SET. -(** We now define a generic solver that works over - any semi-lattice structure. *) +(** Reachability in a control-flow graph. *) -Module Dataflow_Solver (LAT: SEMILATTICE) (NS: NODE_SET): +Section REACHABLE. + +Context {A: Type} (code: PTree.t A) (successors: A -> list positive). + +Inductive reachable: positive -> positive -> Prop := + | reachable_refl: forall n, reachable n n + | reachable_left: forall n1 n2 n3 i, + code!n1 = Some i -> In n2 (successors i) -> reachable n2 n3 -> + reachable n1 n3. + +Scheme reachable_ind := Induction for reachable Sort Prop. + +Lemma reachable_trans: + forall n1 n2, reachable n1 n2 -> forall n3, reachable n2 n3 -> reachable n1 n3. +Proof. + induction 1; intros. +- auto. +- econstructor; eauto. +Qed. + +Lemma reachable_right: + forall n1 n2 n3 i, + reachable n1 n2 -> code!n2 = Some i -> In n3 (successors i) -> + reachable n1 n3. +Proof. + intros. apply reachable_trans with n2; auto. econstructor; eauto. constructor. +Qed. + +End REACHABLE. + +(** We now define a generic solver for forward dataflow inequations + that works over any semi-lattice structure. *) + +Module Dataflow_Solver (LAT: SEMILATTICE) (NS: NODE_SET) <: DATAFLOW_SOLVER with Module L := LAT. Module L := LAT. @@ -163,61 +192,64 @@ Context {A: Type}. Variable code: PTree.t A. Variable successors: A -> list positive. Variable transf: positive -> L.t -> L.t. -Variable entrypoints: list (positive * L.t). -(** The state of the iteration has two components: -- A mapping from program points to values of type [L.t] representing +(** The state of the iteration has three components: +- [aval]: A mapping from program points to values of type [L.t] representing the candidate solution found so far. -- A worklist of program points that remain to be considered. +- [worklist]: A worklist of program points that remain to be considered. +- [visited]: A set of program points that were visited already + (i.e. put on the worklist at some point in the past). + +Only the first two components are computationally relevant. The third +is a ghost variable used only for stating and proving invariants. +For this reason, [visited] is defined at sort [Prop] so that it is +erased during program extraction. *) Record state : Type := - mkstate { st_in: PMap.t L.t; st_wrk: NS.t }. + mkstate { aval: PTree.t L.t; worklist: NS.t; visited: positive -> Prop }. + +Definition abstr_value (n: positive) (s: state) : L.t := + match s.(aval)!n with + | None => L.bot + | Some v => v + end. (** Kildall's algorithm, in pseudo-code, is as follows: << - while st_wrk is not empty, do - extract a node n from st_wrk - compute out = transf n st_in[n] + while worklist is not empty, do + extract a node n from worklist + compute out = transf n aval[n] for each successor s of n: - compute in = lub st_in[s] out - if in <> st_in[s]: - st_in[s] := in - st_wrk := st_wrk union {s} + compute in = lub aval[s] out + if in <> aval[s]: + aval[s] := in + worklist := worklist union {s} + visited := visited union {s} end if end for end while - return st_in + return aval >> - -The initial state is built as follows: -- The initial mapping sets all program points to [L.bot], except - those mentioned in the [entrypoints] list, for which we take - the associated approximation as initial value. Since a program - point can be mentioned several times in [entrypoints], with different - approximations, we actually take the l.u.b. of these approximations. -- The initial worklist contains all the program points. *) - -Fixpoint start_state_in (ep: list (positive * L.t)) : PMap.t L.t := - match ep with - | nil => - PMap.init L.bot - | (n, v) :: rem => - let m := start_state_in rem in PMap.set n (L.lub m!!n v) m - end. - -Definition start_state := - mkstate (start_state_in entrypoints) (NS.initial code). +*) (** [propagate_succ] corresponds, in the pseudocode, to the body of the [for] loop iterating over all successors. *) Definition propagate_succ (s: state) (out: L.t) (n: positive) := - let oldl := s.(st_in)!!n in - let newl := L.lub oldl out in - if L.beq oldl newl - then s - else mkstate (PMap.set n newl s.(st_in)) (NS.add n s.(st_wrk)). + match s.(aval)!n with + | None => + {| aval := PTree.set n out s.(aval); + worklist := NS.add n s.(worklist); + visited := fun p => p = n \/ s.(visited) p |} + | Some oldl => + let newl := L.lub oldl out in + if L.beq oldl newl + then s + else {| aval := PTree.set n newl s.(aval); + worklist := NS.add n s.(worklist); + visited := fun p => p = n \/ s.(visited) p |} + end. (** [propagate_succ_list] corresponds, in the pseudocode, to the [for] loop iterating over all successors. *) @@ -233,366 +265,530 @@ Fixpoint propagate_succ_list (s: state) (out: L.t) (succs: list positive) pseudocode. *) Definition step (s: state) : PMap.t L.t + state := - match NS.pick s.(st_wrk) with + match NS.pick s.(worklist) with | None => - inl _ s.(st_in) + inl _ (L.bot, s.(aval)) | Some(n, rem) => match code!n with - | None => inr _ (mkstate s.(st_in) rem) + | None => + inr _ {| aval := s.(aval); worklist := rem; visited := s.(visited) |} | Some instr => inr _ (propagate_succ_list - (mkstate s.(st_in) rem) - (transf n s.(st_in)!!n) + {| aval := s.(aval); worklist := rem; visited := s.(visited) |} + (transf n (abstr_value n s)) (successors instr)) end end. (** The whole fixpoint computation is the iteration of [step] from - the start state. *) + an initial state. *) -Definition fixpoint : option (PMap.t L.t) := - PrimIter.iterate _ _ step start_state. +Definition fixpoint_from (start: state) : option (PMap.t L.t) := + PrimIter.iterate _ _ step start. -(** ** Monotonicity properties *) +(** There are several ways to build the initial state. For forward + dataflow analyses, the initial worklist is the function entry point, + and the initial mapping sets the function entry point to the given + abstract value, and leaves unset all other program points, which + corresponds to [L.bot] initial abstract values. *) + +Definition start_state (enode: positive) (eval: L.t) := + {| aval := PTree.set enode eval (PTree.empty L.t); + worklist := NS.add enode NS.empty; + visited := fun n => n = enode |}. + +Definition fixpoint (enode: positive) (eval: L.t) := + fixpoint_from (start_state enode eval). -(** We first show that the [st_in] part of the state evolves monotonically: - at each step, the values of the [st_in[n]] either remain the same or - increase with respect to the [L.ge] ordering. *) +(** For backward analyses (viewed as forward analyses on the reversed CFG), + the following two variants are more useful. Both start with an + empty initial mapping, where all program points start at [L.bot]. + The first initializes the worklist with a given set of entry points + in the reversed CFG. (See the backward dataflow solver below for + how this list is computed.) The second start state construction + initializes the worklist with all program points of the given CFG. *) -Definition in_incr (in1 in2: PMap.t L.t) : Prop := - forall n, L.ge in2!!n in1!!n. +Definition start_state_nodeset (enodes: NS.t) := + {| aval := PTree.empty L.t; + worklist := enodes; + visited := fun n => NS.In n enodes |}. -Lemma in_incr_refl: - forall in1, in_incr in1 in1. +Definition fixpoint_nodeset (enodes: NS.t) := + fixpoint_from (start_state_nodeset enodes). + +Definition start_state_allnodes := + {| aval := PTree.empty L.t; + worklist := NS.all_nodes code; + visited := fun n => exists instr, code!n = Some instr |}. + +Definition fixpoint_allnodes := + fixpoint_from start_state_allnodes. + +(** ** Characterization of the propagation functions *) + +Inductive optge: option L.t -> option L.t -> Prop := + | optge_some: forall l l', + L.ge l l' -> optge (Some l) (Some l') + | optge_none: forall ol, + optge ol None. + +Remark optge_refl: forall ol, optge ol ol. +Proof. destruct ol; constructor. apply L.ge_refl; apply L.eq_refl. Qed. + +Remark optge_trans: forall ol1 ol2 ol3, optge ol1 ol2 -> optge ol2 ol3 -> optge ol1 ol3. Proof. - unfold in_incr; intros. apply L.ge_refl. apply L.eq_refl. + intros. inv H0. + inv H. constructor. eapply L.ge_trans; eauto. + constructor. Qed. -Lemma in_incr_trans: - forall in1 in2 in3, in_incr in1 in2 -> in_incr in2 in3 -> in_incr in1 in3. +Remark optge_abstr_value: + forall st st' n, + optge st.(aval)!n st'.(aval)!n -> + L.ge (abstr_value n st) (abstr_value n st'). Proof. - unfold in_incr; intros. apply L.ge_trans with in2!!n; auto. + intros. unfold abstr_value. inv H. auto. apply L.ge_bot. Qed. -Lemma propagate_succ_incr: +Lemma propagate_succ_charact: forall st out n, - in_incr st.(st_in) (propagate_succ st out n).(st_in). + let st' := propagate_succ st out n in + optge st'.(aval)!n (Some out) + /\ (forall s, n <> s -> st'.(aval)!s = st.(aval)!s) + /\ (forall s, optge st'.(aval)!s st.(aval)!s) + /\ (NS.In n st'.(worklist) \/ st'.(aval)!n = st.(aval)!n) + /\ (forall n', NS.In n' st.(worklist) -> NS.In n' st'.(worklist)) + /\ (forall n', NS.In n' st'.(worklist) -> n' = n \/ NS.In n' st.(worklist)) + /\ (forall n', st.(visited) n' -> st'.(visited) n') + /\ (forall n', st'.(visited) n' -> NS.In n' st'.(worklist) \/ st.(visited) n') + /\ (forall n', st.(aval)!n' = None -> st'.(aval)!n' <> None -> st'.(visited) n'). Proof. - unfold in_incr, propagate_succ; simpl; intros. - case (L.beq st.(st_in)!!n (L.lub st.(st_in)!!n out)). - apply L.ge_refl. apply L.eq_refl. - simpl. case (peq n n0); intro. - subst n0. rewrite PMap.gss. apply L.ge_lub_left. - rewrite PMap.gso; auto. apply L.ge_refl. apply L.eq_refl. + unfold propagate_succ; intros; simpl. + destruct st.(aval)!n as [v|] eqn:E; + [predSpec L.beq L.beq_correct v (L.lub v out) | idtac]. +- (* already there, unchanged *) + repeat split; intros. + + rewrite E. constructor. eapply L.ge_trans. apply L.ge_refl. apply H; auto. apply L.ge_lub_right. + + apply optge_refl. + + right; auto. + + auto. + + auto. + + auto. + + auto. + + congruence. +- (* already there, updated *) + simpl; repeat split; intros. + + rewrite PTree.gss. constructor. apply L.ge_lub_right. + + rewrite PTree.gso by auto. auto. + + rewrite PTree.gsspec. destruct (peq s n). + subst s. rewrite E. constructor. apply L.ge_lub_left. + apply optge_refl. + + rewrite NS.add_spec. auto. + + rewrite NS.add_spec. auto. + + rewrite NS.add_spec in H0. intuition. + + auto. + + destruct H0; auto. subst n'. rewrite NS.add_spec; auto. + + rewrite PTree.gsspec in H1. destruct (peq n' n). auto. congruence. +- (* not previously there, updated *) + simpl; repeat split; intros. + + rewrite PTree.gss. apply optge_refl. + + rewrite PTree.gso by auto. auto. + + rewrite PTree.gsspec. destruct (peq s n). + subst s. rewrite E. constructor. + apply optge_refl. + + rewrite NS.add_spec. auto. + + rewrite NS.add_spec. auto. + + rewrite NS.add_spec in H. intuition. + + auto. + + destruct H; auto. subst n'. rewrite NS.add_spec. auto. + + rewrite PTree.gsspec in H0. destruct (peq n' n). auto. congruence. Qed. -Lemma propagate_succ_list_incr: - forall out scs st, - in_incr st.(st_in) (propagate_succ_list st out scs).(st_in). +Lemma propagate_succ_list_charact: + forall out l st, + let st' := propagate_succ_list st out l in + (forall n, In n l -> optge st'.(aval)!n (Some out)) + /\ (forall n, ~In n l -> st'.(aval)!n = st.(aval)!n) + /\ (forall n, optge st'.(aval)!n st.(aval)!n) + /\ (forall n, NS.In n st'.(worklist) \/ st'.(aval)!n = st.(aval)!n) + /\ (forall n', NS.In n' st.(worklist) -> NS.In n' st'.(worklist)) + /\ (forall n', NS.In n' st'.(worklist) -> In n' l \/ NS.In n' st.(worklist)) + /\ (forall n', st.(visited) n' -> st'.(visited) n') + /\ (forall n', st'.(visited) n' -> NS.In n' st'.(worklist) \/ st.(visited) n') + /\ (forall n', st.(aval)!n' = None -> st'.(aval)!n' <> None -> st'.(visited) n'). Proof. - induction scs; simpl; intros. - apply in_incr_refl. - apply in_incr_trans with (propagate_succ st out a).(st_in). - apply propagate_succ_incr. auto. -Qed. + induction l; simpl; intros. +- repeat split; intros. + + contradiction. + + apply optge_refl. + + auto. + + auto. + + auto. + + auto. + + auto. + + congruence. +- generalize (propagate_succ_charact st out a). + set (st1 := propagate_succ st out a). + intros (A1 & A2 & A3 & A4 & A5 & A6 & A7 & A8 & A9). + generalize (IHl st1). + set (st2 := propagate_succ_list st1 out l). + intros (B1 & B2 & B3 & B4 & B5 & B6 & B7 & B8 & B9). clear IHl. + repeat split; intros. + + destruct H. + * subst n. eapply optge_trans; eauto. + * auto. + + rewrite B2 by tauto. apply A2; tauto. + + eapply optge_trans; eauto. + + destruct (B4 n). auto. + destruct (peq n a). + * subst n. destruct A4. left; auto. right; congruence. + * right. rewrite H. auto. + + eauto. + + exploit B6; eauto. intros [P|P]. auto. + exploit A6; eauto. intuition. + + eauto. + + specialize (B8 n'); specialize (A8 n'). intuition. + + destruct st1.(aval)!n' eqn:ST1. + apply B7. apply A9; auto. congruence. + apply B9; auto. +Qed. -Lemma fixpoint_incr: - forall res, - fixpoint = Some res -> in_incr (start_state_in entrypoints) res. +(** Characterization of [fixpoint_from]. *) + +Inductive steps: state -> state -> Prop := + | steps_base: forall s, steps s s + | steps_right: forall s1 s2 s3, steps s1 s2 -> step s2 = inr s3 -> steps s1 s3. + +Scheme steps_ind := Induction for steps Sort Prop. + +Lemma fixpoint_from_charact: + forall start res, + fixpoint_from start = Some res -> + exists st, steps start st /\ NS.pick st.(worklist) = None /\ res = (L.bot, st.(aval)). Proof. unfold fixpoint; intros. - change (start_state_in entrypoints) with start_state.(st_in). eapply (PrimIter.iterate_prop _ _ step - (fun st => in_incr start_state.(st_in) st.(st_in)) - (fun res => in_incr start_state.(st_in) res)). - - intros st INCR. unfold step. - destruct (NS.pick st.(st_wrk)) as [ [n rem] | ]. - destruct (code!n) as [instr | ]. - apply in_incr_trans with st.(st_in). auto. - change st.(st_in) with (mkstate st.(st_in) rem).(st_in). - apply propagate_succ_list_incr. - auto. - auto. - eauto. - apply in_incr_refl. + (fun st => steps start st) + (fun res => exists st, steps start st /\ NS.pick (worklist st) = None /\ res = (L.bot, aval st))); eauto. + intros. destruct (step a) eqn:E. + exists a; split; auto. + unfold step in E. destruct (NS.pick (worklist a)) as [[n rem]|]. + destruct (code!n); discriminate. + inv E. auto. + eapply steps_right; eauto. + constructor. Qed. -(** ** Correctness invariant *) - -(** The following invariant is preserved at each iteration of Kildall's - algorithm: for all program points [n], either - [n] is in the worklist, or the inequations associated with [n] - ([st_in[s] >= transf n st_in[n]] for all successors [s] of [n]) - hold. In other terms, the worklist contains all nodes that do not - yet satisfy their inequations. *) - -Definition good_state (st: state) : Prop := - forall n, - NS.In n st.(st_wrk) \/ - (forall instr s, - code!n = Some instr -> - In s (successors instr) -> - L.ge st.(st_in)!!s (transf n st.(st_in)!!n)). +(** ** Monotonicity properties *) -(** We show that the start state satisfies the invariant, and that - the [step] function preserves it. *) +(** We first show that the [aval] and [visited] parts of the state +evolve monotonically: +- at each step, the values of the [aval[n]] either remain the same or + increase with respect to the [optge] ordering; +- every node visited in the past remains visited in the future. +*) -Lemma start_state_good: - good_state start_state. +Lemma step_incr: + forall n s1 s2, step s1 = inr s2 -> + optge s2.(aval)!n s1.(aval)!n /\ (s1.(visited) n -> s2.(visited) n). Proof. - unfold good_state, start_state; intros. - destruct (code!n) as [instr|] eqn:INSTR. - left; simpl. eapply NS.initial_spec; eauto. - right; intros. discriminate. + unfold step; intros. + destruct (NS.pick (worklist s1)) as [[p rem] | ]; try discriminate. + destruct (code!p) as [instr|]; inv H. + + generalize (propagate_succ_list_charact + (transf p (abstr_value p s1)) + (successors instr) + {| aval := aval s1; worklist := rem; visited := visited s1 |}). + simpl. + set (s' := propagate_succ_list {| aval := aval s1; worklist := rem; visited := visited s1 |} + (transf p (abstr_value p s1)) (successors instr)). + intros (A1 & A2 & A3 & A4 & A5 & A6 & A7 & A8 & A9). + auto. + + split. apply optge_refl. auto. Qed. -Lemma propagate_succ_charact: - forall st out n, - let st' := propagate_succ st out n in - L.ge st'.(st_in)!!n out /\ - (forall s, n <> s -> st'.(st_in)!!s = st.(st_in)!!s). +Lemma steps_incr: + forall n s1 s2, steps s1 s2 -> + optge s2.(aval)!n s1.(aval)!n /\ (s1.(visited) n -> s2.(visited) n). Proof. - unfold propagate_succ; intros; simpl. - predSpec L.beq L.beq_correct - ((st_in st) !! n) (L.lub (st_in st) !! n out). - split. - eapply L.ge_trans. apply L.ge_refl. apply H; auto. - apply L.ge_lub_right. - auto. - - simpl. split. - rewrite PMap.gss. - apply L.ge_lub_right. - intros. rewrite PMap.gso; auto. + induction 1. +- split. apply optge_refl. auto. +- destruct IHsteps. exploit (step_incr n); eauto. intros [P Q]. + split. eapply optge_trans; eauto. eauto. Qed. -Lemma propagate_succ_list_charact: - forall out scs st, - let st' := propagate_succ_list st out scs in - forall s, - (In s scs -> L.ge st'.(st_in)!!s out) /\ - (~(In s scs) -> st'.(st_in)!!s = st.(st_in)!!s). +(** ** Correctness invariant *) + +(** The following invariant is preserved at each iteration of Kildall's + algorithm: for all visited program point [n], either + [n] is in the worklist, or the inequations associated with [n] + ([aval[s] >= transf n aval[n]] for all successors [s] of [n]) + hold. In other terms, the worklist contains all nodes that were + visited but do not yet satisfy their inequations. + + The second part of the invariant shows that nodes that already have + an abstract value associated with them have been visited. *) + +Record good_state (st: state) : Prop := { + gs_stable: forall n, + st.(visited) n -> + NS.In n st.(worklist) \/ + (forall i s, + code!n = Some i -> In s (successors i) -> + optge st.(aval)!s (Some (transf n (abstr_value n st)))); + gs_defined: forall n v, + st.(aval)!n = Some v -> st.(visited) n +}. + +(** We show that the [step] function preserves this invariant. *) + +Lemma step_state_good: + forall st pc rem instr, + NS.pick st.(worklist) = Some (pc, rem) -> + code!pc = Some instr -> + good_state st -> + good_state (propagate_succ_list (mkstate st.(aval) rem st.(visited)) + (transf pc (abstr_value pc st)) + (successors instr)). Proof. - induction scs; simpl; intros. - tauto. - generalize (IHscs (propagate_succ st out a) s). intros [P Q]. - generalize (propagate_succ_charact st out a). intros [U V]. - split; intros. - elim H; intro. - subst s. - apply L.ge_trans with (propagate_succ st out a).(st_in)!!a. - apply propagate_succ_list_incr. assumption. - apply P. auto. - transitivity (propagate_succ st out a).(st_in)!!s. - apply Q. tauto. - apply V. tauto. + intros until instr; intros PICK CODEAT [GOOD1 GOOD2]. + generalize (NS.pick_some _ _ _ PICK); intro PICK2. + set (out := transf pc (abstr_value pc st)). + generalize (propagate_succ_list_charact out (successors instr) {| aval := aval st; worklist := rem; visited := visited st |}). + set (st' := propagate_succ_list {| aval := aval st; worklist := rem; visited := visited st |} out + (successors instr)). + simpl; intros (A1 & A2 & A3 & A4 & A5 & A6 & A7 & A8 & A9). + constructor; intros. +- (* stable *) + destruct (A8 n H); auto. destruct (A4 n); auto. + replace (abstr_value n st') with (abstr_value n st) + by (unfold abstr_value; rewrite H1; auto). + exploit GOOD1; eauto. intros [P|P]. ++ (* n was on the worklist *) + rewrite PICK2 in P; destruct P. + * (* node n is our node pc *) + subst n. fold out. right; intros. + assert (i = instr) by congruence. subst i. + apply A1; auto. + * (* n was already on the worklist *) + left. apply A5; auto. ++ (* n was stable before, still is *) + right; intros. apply optge_trans with st.(aval)!s; eauto. +- (* defined *) + destruct st.(aval)!n as [v'|] eqn:ST. + + apply A7. eapply GOOD2; eauto. + + apply A9; auto. congruence. Qed. -Lemma propagate_succ_incr_worklist: - forall st out n x, - NS.In x st.(st_wrk) -> NS.In x (propagate_succ st out n).(st_wrk). +Lemma step_state_good_2: + forall st pc rem, + good_state st -> + NS.pick (worklist st) = Some (pc, rem) -> + code!pc = None -> + good_state (mkstate st.(aval) rem st.(visited)). Proof. - intros. unfold propagate_succ. - case (L.beq (st_in st) !! n (L.lub (st_in st) !! n out)). - auto. - simpl. rewrite NS.add_spec. auto. + intros until rem; intros [GOOD1 GOOD2] PICK CODE. + generalize (NS.pick_some _ _ _ PICK); intro PICK2. + constructor; simpl; intros. +- (* stable *) + exploit GOOD1; eauto. intros [P | P]. + + rewrite PICK2 in P. destruct P; auto. + subst n. right; intros. congruence. + + right; exact P. +- (* defined *) + eapply GOOD2; eauto. Qed. -Lemma propagate_succ_list_incr_worklist: - forall out scs st x, - NS.In x st.(st_wrk) -> NS.In x (propagate_succ_list st out scs).(st_wrk). +Lemma steps_state_good: + forall st1 st2, steps st1 st2 -> good_state st1 -> good_state st2. Proof. - induction scs; simpl; intros. - auto. - apply IHscs. apply propagate_succ_incr_worklist. auto. + induction 1; intros. +- auto. +- unfold step in e. + destruct (NS.pick (worklist s2)) as [[n rem] | ] eqn:PICK; try discriminate. + destruct (code!n) as [instr|] eqn:CODE; inv e. + eapply step_state_good; eauto. + eapply step_state_good_2; eauto. Qed. -Lemma propagate_succ_records_changes: - forall st out n s, - let st' := propagate_succ st out n in - NS.In s st'.(st_wrk) \/ st'.(st_in)!!s = st.(st_in)!!s. +(** We show that initial states satisfy the invariant. *) + +Lemma start_state_good: + forall enode eval, good_state (start_state enode eval). Proof. - simpl. intros. unfold propagate_succ. - case (L.beq (st_in st) !! n (L.lub (st_in st) !! n out)). - right; auto. - case (peq s n); intro. - subst s. left. simpl. rewrite NS.add_spec. auto. - right. simpl. apply PMap.gso. auto. + intros. unfold start_state; constructor; simpl; intros. +- subst n. rewrite NS.add_spec; auto. +- rewrite PTree.gsspec in H. rewrite PTree.gempty in H. + destruct (peq n enode). auto. discriminate. Qed. -Lemma propagate_succ_list_records_changes: - forall out scs st s, - let st' := propagate_succ_list st out scs in - NS.In s st'.(st_wrk) \/ st'.(st_in)!!s = st.(st_in)!!s. +Lemma start_state_nodeset_good: + forall enodes, good_state (start_state_nodeset enodes). Proof. - induction scs; simpl; intros. - right; auto. - elim (propagate_succ_records_changes st out a s); intro. - left. apply propagate_succ_list_incr_worklist. auto. - rewrite <- H. auto. + intros. unfold start_state_nodeset; constructor; simpl; intros. +- left. auto. +- rewrite PTree.gempty in H. congruence. Qed. -Lemma step_state_good: - forall st n instr rem, - NS.pick st.(st_wrk) = Some(n, rem) -> - code!n = Some instr -> - good_state st -> - good_state (propagate_succ_list (mkstate st.(st_in) rem) - (transf n st.(st_in)!!n) - (successors instr)). +Lemma start_state_allnodes_good: + good_state start_state_allnodes. Proof. - unfold good_state. intros st n instr rem WKL INSTR GOOD x. - generalize (NS.pick_some _ _ _ WKL); intro PICK. - set (out := transf n st.(st_in)!!n). - elim (propagate_succ_list_records_changes - out (successors instr) (mkstate st.(st_in) rem) x). - intro; left; auto. - simpl; intros EQ. rewrite EQ. - (* Case 1: x = n *) - destruct (peq x n). subst x. - right; intros. - assert (instr0 = instr) by congruence. subst instr0. - elim (propagate_succ_list_charact out (successors instr) - (mkstate st.(st_in) rem) s); intros. - auto. - (* Case 2: x <> n *) - elim (GOOD x); intro. - (* Case 2.1: x was already in worklist, still is *) - left. apply propagate_succ_list_incr_worklist. - simpl. rewrite PICK in H. elim H; intro. congruence. auto. - (* Case 2.2: x was not in worklist *) - right; intros. - case (In_dec peq s (successors instr)); intro. - (* Case 2.2.1: s is a successor of n, it may have increased *) - apply L.ge_trans with st.(st_in)!!s. - change st.(st_in)!!s with (mkstate st.(st_in) rem).(st_in)!!s. - apply propagate_succ_list_incr. - eauto. - (* Case 2.2.2: s is not a successor of n, it did not change *) - elim (propagate_succ_list_charact out (successors instr) - (mkstate st.(st_in) rem) s); intros. - rewrite H3. simpl. eauto. auto. + unfold start_state_allnodes; constructor; simpl; intros. +- destruct H as [instr CODE]. left. eapply NS.all_nodes_spec; eauto. +- rewrite PTree.gempty in H. congruence. Qed. -Lemma step_state_good_2: - forall st n rem, - good_state st -> - NS.pick (st_wrk st) = Some (n, rem) -> - code!n = None -> - good_state (mkstate st.(st_in) rem). +(** Reachability in final states. *) + +Lemma reachable_visited: + forall st, good_state st -> NS.pick st.(worklist) = None -> + forall p q, reachable code successors p q -> st.(visited) p -> st.(visited) q. Proof. - intros; red; intros; simpl. - destruct (H n0). - erewrite NS.pick_some in H2 by eauto. destruct H2; auto. - subst n0. right; intros; congruence. - right; auto. + intros st [GOOD1 GOOD2] PICK. induction 1; intros. +- auto. +- eapply IHreachable; eauto. + exploit GOOD1; eauto. intros [P | P]. + eelim NS.pick_none; eauto. + exploit P; eauto. intros OGE; inv OGE. eapply GOOD2; eauto. Qed. -(** ** Correctness of the solution returned by [iterate]. *) +(** ** Correctness of the solution returned by [fixpoint]. *) (** As a consequence of the [good_state] invariant, the result of - [fixpoint], if defined, is a solution of the dataflow inequations, - since [st_wrk] is empty when the iteration terminates. *) + [fixpoint], if defined, is a solution of the dataflow inequations. + This assumes that the transfer function maps [L.bot] to [L.bot]. *) Theorem fixpoint_solution: - forall res n instr s, - fixpoint = Some res -> + forall ep ev res n instr s, + fixpoint ep ev = Some res -> code!n = Some instr -> In s (successors instr) -> + (forall n, L.eq (transf n L.bot) L.bot) -> L.ge res!!s (transf n res!!n). Proof. - intros until s. unfold fixpoint. intros PI. revert n instr s. - pattern res. - eapply (PrimIter.iterate_prop _ _ step good_state). - - intros st GOOD. unfold step. - destruct (NS.pick st.(st_wrk)) as [[n rem] | ] eqn:PICK. - destruct (code!n) as [instr | ] eqn:INSTR. - apply step_state_good; auto. - eapply step_state_good_2; eauto. - intros. destruct (GOOD n). elim (NS.pick_none _ n PICK). auto. eauto. + unfold fixpoint; intros. + exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES). + exploit steps_state_good; eauto. apply start_state_good. intros [GOOD1 GOOD2]. + rewrite RES; unfold PMap.get; simpl. + destruct st.(aval)!n as [v|] eqn:STN. +- destruct (GOOD1 n) as [P|P]; eauto. + eelim NS.pick_none; eauto. + exploit P; eauto. unfold abstr_value; rewrite STN. intros OGE; inv OGE. auto. +- apply L.ge_trans with L.bot. apply L.ge_bot. apply L.ge_refl. apply L.eq_sym. eauto. +Qed. + +(** Moreover, the result of [fixpoint], if defined, satisfies the additional + constraint given on the entry point. *) - eauto. apply start_state_good. +Theorem fixpoint_entry: + forall ep ev res, + fixpoint ep ev = Some res -> + L.ge res!!ep ev. +Proof. + unfold fixpoint; intros. + exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES). + exploit (steps_incr ep); eauto. simpl. rewrite PTree.gss. intros [P Q]. + rewrite RES; unfold PMap.get; simpl. inv P; auto. Qed. -(** As a consequence of the monotonicity property, the result of - [fixpoint], if defined, is pointwise greater than or equal the - initial mapping. Therefore, it satisfies the additional constraints - stated in [entrypoints]. *) +(** For [fixpoint_allnodes], we show that the result is a solution + without assuming [transf n L.bot = L.bot]. *) -Lemma start_state_in_entry: - forall ep n v, - In (n, v) ep -> - L.ge (start_state_in ep)!!n v. +Theorem fixpoint_allnodes_solution: + forall res n instr s, + fixpoint_allnodes = Some res -> + code!n = Some instr -> + In s (successors instr) -> + L.ge res!!s (transf n res!!n). Proof. - induction ep; simpl; intros. - elim H. - elim H; intros. - subst a. rewrite PMap.gss. - apply L.ge_lub_right. - destruct a. rewrite PMap.gsspec. case (peq n p); intro. - subst p. apply L.ge_trans with (start_state_in ep)!!n. - apply L.ge_lub_left. auto. - auto. + unfold fixpoint_allnodes; intros. + exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES). + exploit steps_state_good; eauto. apply start_state_allnodes_good. intros [GOOD1 GOOD2]. + exploit (steps_incr n); eauto. simpl. intros [U V]. + exploit (GOOD1 n). apply V. exists instr; auto. intros [P|P]. + eelim NS.pick_none; eauto. + exploit P; eauto. intros OGE. rewrite RES; unfold PMap.get; simpl. + inv OGE. assumption. Qed. -Theorem fixpoint_entry: - forall res n v, - fixpoint = Some res -> - In (n, v) entrypoints -> - L.ge res!!n v. +(** For [fixpoint_nodeset], we show that the result is a solution + at all program points that are reachable from the given entry points. *) + +Theorem fixpoint_nodeset_solution: + forall enodes res e n instr s, + fixpoint_nodeset enodes = Some res -> + NS.In e enodes -> + reachable code successors e n -> + code!n = Some instr -> + In s (successors instr) -> + L.ge res!!s (transf n res!!n). Proof. - intros. - apply L.ge_trans with (start_state_in entrypoints)!!n. - apply fixpoint_incr. auto. - apply start_state_in_entry. auto. + unfold fixpoint_nodeset; intros. + exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES). + exploit steps_state_good; eauto. apply start_state_nodeset_good. intros GOOD. + exploit (steps_incr e); eauto. simpl. intros [U V]. + assert (st.(visited) n). + { eapply reachable_visited; eauto. } + destruct GOOD as [GOOD1 GOOD2]. + exploit (GOOD1 n); eauto. intros [P|P]. + eelim NS.pick_none; eauto. + exploit P; eauto. intros OGE. rewrite RES; unfold PMap.get; simpl. + inv OGE. assumption. Qed. (** ** Preservation of a property over solutions *) -Variable P: L.t -> Prop. -Hypothesis P_bot: P L.bot. -Hypothesis P_lub: forall x y, P x -> P y -> P (L.lub x y). -Hypothesis P_transf: forall pc instr x, code!pc = Some instr -> P x -> P (transf pc x). -Hypothesis P_entrypoints: forall n v, In (n, v) entrypoints -> P v. - Theorem fixpoint_invariant: - forall res pc, - fixpoint = Some res -> + forall ep ev + (P: L.t -> Prop) + (P_bot: P L.bot) + (P_lub: forall x y, P x -> P y -> P (L.lub x y)) + (P_transf: forall pc instr x, code!pc = Some instr -> P x -> P (transf pc x)) + (P_entrypoint: P ev) + res pc, + fixpoint ep ev = Some res -> P res!!pc. Proof. - assert (forall ep, - (forall n v, In (n, v) ep -> P v) -> - forall pc, P (start_state_in ep)!!pc). - induction ep; intros; simpl. - rewrite PMap.gi. auto. - simpl in H. - assert (P (start_state_in ep)!!pc). apply IHep. eauto. - destruct a as [n v]. rewrite PMap.gsspec. destruct (peq pc n). - apply P_lub. subst. auto. eapply H. left; reflexivity. auto. - set (inv := fun st => forall pc, P (st.(st_in)!!pc)). + intros. + set (inv := fun st => forall x, P (abstr_value x st)). + assert (inv (start_state ep ev)). + { + red; simpl; intros. unfold abstr_value, start_state; simpl. + rewrite PTree.gsspec. rewrite PTree.gempty. + destruct (peq x ep). auto. auto. + } assert (forall st v n, inv st -> P v -> inv (propagate_succ st v n)). - unfold inv, propagate_succ. intros. - destruct (LAT.beq (st_in st)!!n (LAT.lub (st_in st)!!n v)). - auto. simpl. rewrite PMap.gsspec. destruct (peq pc n). - apply P_lub. subst n; auto. auto. + { + unfold inv, propagate_succ. intros. + destruct (aval st)!n as [oldl|] eqn:E. + destruct (L.beq oldl (L.lub oldl v)). + auto. + unfold abstr_value. simpl. rewrite PTree.gsspec. destruct (peq x n). + apply P_lub; auto. replace oldl with (abstr_value n st). auto. + unfold abstr_value; rewrite E; auto. + apply H1. + unfold abstr_value. simpl. rewrite PTree.gsspec. destruct (peq x n). auto. + apply H1. + } assert (forall l st v, inv st -> P v -> inv (propagate_succ_list st v l)). + { induction l; intros; simpl. auto. apply IHl; auto. - assert (forall res, fixpoint = Some res -> forall pc, P res!!pc). - unfold fixpoint. intros res0 RES0. pattern res0. - eapply (PrimIter.iterate_prop _ _ step inv). - intros. unfold step. - destruct (NS.pick (st_wrk a)) as [[n rem] | ]. - destruct (code!n) as [instr | ] eqn:INSTR. - apply H1. auto. eapply P_transf; eauto. - assumption. - eauto. - eauto. - unfold inv, start_state; simpl. auto. - intros. auto. + } + assert (forall st1 st2, steps st1 st2 -> inv st1 -> inv st2). + { + induction 1; intros. + auto. + unfold step in e. destruct (NS.pick (worklist s2)) as [[n rem]|]; try discriminate. + destruct (code!n) as [instr|] eqn:INSTR; inv e. + apply H2. apply IHsteps; auto. eapply P_transf; eauto. apply IHsteps; auto. + apply IHsteps; auto. + } + unfold fixpoint in H. exploit fixpoint_from_charact; eauto. + intros (st & STEPS & PICK & RES). + replace (res!!pc) with (abstr_value pc st). eapply H3; eauto. + rewrite RES; auto. Qed. End Kildall. @@ -606,7 +802,12 @@ End Dataflow_Solver. successors. We exploit this observation to cheaply derive a backward solver from the forward solver. *) -(** ** Construction of the predecessor relation *) +(** ** Construction of the reversed flow graph (the predecessor relation) *) + +Definition successors_list (successors: PTree.t (list positive)) (pc: positive) : list positive := + match successors!pc with None => nil | Some l => l end. + +Notation "a !!! b" := (successors_list a b) (at level 1). Section Predecessor. @@ -672,6 +873,17 @@ Proof. contradiction. Qed. +Lemma reachable_predecessors: + forall p q, + reachable code successors p q -> + reachable make_predecessors (fun l => l) q p. +Proof. + induction 1. +- constructor. +- exploit make_predecessors_correct_2; eauto. intros [l [P Q]]. + eapply reachable_right; eauto. +Qed. + End Predecessor. (** ** Solving backward dataflow problems *) @@ -682,33 +894,40 @@ Module Type BACKWARD_DATAFLOW_SOLVER. Declare Module L: SEMILATTICE. + (** [fixpoint successors transf] is the solver. + It returns either an error or a mapping from program points to + values of type [L.t] representing the solution. [successors] + is a finite map returning the list of successors of the given program + point. [transf] is the transfer function. *) + Variable fixpoint: forall {A: Type} (code: PTree.t A) (successors: A -> list positive) - (transf: positive -> L.t -> L.t) - (entrypoints: list (positive * L.t)), + (transf: positive -> L.t -> L.t), option (PMap.t L.t). + (** The [fixpoint_solution] theorem shows that the returned solution, + if any, satisfies the backward dataflow inequations. *) + Hypothesis fixpoint_solution: - forall A (code: PTree.t A) successors transf entrypoints res n instr s, - fixpoint code successors transf entrypoints = Some res -> + forall A (code: PTree.t A) successors transf res n instr s, + fixpoint code successors transf = Some res -> code!n = Some instr -> In s (successors instr) -> + (forall n a, code!n = None -> L.eq (transf n a) L.bot) -> L.ge res!!n (transf s res!!s). - Hypothesis fixpoint_entry: - forall A (code: PTree.t A) successors transf entrypoints res n v, - fixpoint code successors transf entrypoints = Some res -> - In (n, v) entrypoints -> - L.ge res!!n v. + (** [fixpoint_allnodes] is a variant of [fixpoint], less algorithmically + efficient, but correct without any hypothesis on the transfer function. *) - Hypothesis fixpoint_invariant: - forall A (code: PTree.t A) successors transf entrypoints (P: L.t -> Prop), - P L.bot -> - (forall x y, P x -> P y -> P (L.lub x y)) -> - (forall pc x, P x -> P (transf pc x)) -> - (forall n v, In (n, v) entrypoints -> P v) -> - forall res pc, - fixpoint code successors transf entrypoints = Some res -> - P res!!pc. + Variable fixpoint_allnodes: + forall {A: Type} (code: PTree.t A) (successors: A -> list positive) + (transf: positive -> L.t -> L.t), + option (PMap.t L.t). + + Hypothesis fixpoint_allnodes_solution: + forall A (code: PTree.t A) successors transf res n instr s, + fixpoint_allnodes code successors transf = Some res -> + code!n = Some instr -> In s (successors instr) -> + L.ge res!!n (transf s res!!s). End BACKWARD_DATAFLOW_SOLVER. @@ -729,45 +948,127 @@ Context {A: Type}. Variable code: PTree.t A. Variable successors: A -> list positive. Variable transf: positive -> L.t -> L.t. -Variable entrypoints: list (positive * L.t). + +(** Finding entry points for the reverse control-flow graph. *) + +Section Exit_points. + +(** Assuming that the nodes of the CFG [code] are numbered in reverse + postorder (cf. pass [Renumber]), an edge from [n] to [s] is a + normal edge if [s < n] and a back-edge otherwise. + [sequential_node] returns [true] if the given node has at least one + normal outgoing edge. It returns [false] if the given node is an exit + node (no outgoing edges) or the final node of a loop body + (all outgoing edges are back-edges). As we prove later, the set + of all non-sequential node is an appropriate set of entry points + for exploring the reverse CFG. *) + +Definition sequential_node (pc: positive) (instr: A): bool := + existsb (fun s => match code!s with None => false | Some _ => plt s pc end) + (successors instr). + +Definition exit_points : NS.t := + PTree.fold + (fun ep pc instr => + if sequential_node pc instr + then ep + else NS.add pc ep) + code NS.empty. + +Lemma exit_points_charact: + forall n, + NS.In n exit_points <-> exists i, code!n = Some i /\ sequential_node n i = false. +Proof. + intros n. unfold exit_points. eapply PTree_Properties.fold_rec. +- (* extensionality *) + intros. rewrite <- H. auto. +- (* base case *) + simpl. split; intros. + eelim NS.empty_spec; eauto. + destruct H as [i [P Q]]. rewrite PTree.gempty in P. congruence. +- (* inductive case *) + intros. destruct (sequential_node k v) eqn:SN. + + rewrite H1. rewrite PTree.gsspec. destruct (peq n k). + subst. split; intros [i [P Q]]. congruence. inv P. congruence. + tauto. + + rewrite NS.add_spec. rewrite H1. rewrite PTree.gsspec. destruct (peq n k). + subst. split. intros. exists v; auto. auto. + split. intros [P | [i [P Q]]]. congruence. exists i; auto. + intros [i [P Q]]. right; exists i; auto. +Qed. + +Lemma reachable_exit_points: + forall pc i, + code!pc = Some i -> exists x, NS.In x exit_points /\ reachable code successors pc x. +Proof. + intros pc0. pattern pc0. apply (well_founded_ind Plt_wf). + intros pc HR i CODE. + destruct (sequential_node pc i) eqn:SN. +- (* at least one successor that decreases the pc *) + unfold sequential_node in SN. rewrite existsb_exists in SN. + destruct SN as [s [P Q]]. destruct (code!s) as [i'|] eqn:CS; try discriminate. InvBooleans. + exploit (HR s); eauto. intros [x [U V]]. + exists x; split; auto. eapply reachable_left; eauto. +- (* otherwise we are an exit point *) + exists pc; split. + rewrite exit_points_charact. exists i; auto. constructor. +Qed. + +(** The crucial property of exit points is that any nonempty node in the + CFG is reverse-reachable from an exit point. *) + +Lemma reachable_exit_points_predecessor: + forall pc i, + code!pc = Some i -> + exists x, NS.In x exit_points /\ reachable (make_predecessors code successors) (fun l => l) x pc. +Proof. + intros. exploit reachable_exit_points; eauto. intros [x [P Q]]. + exists x; split; auto. apply reachable_predecessors. auto. +Qed. + +End Exit_points. + +(** The efficient backward solver. *) Definition fixpoint := - DS.fixpoint + DS.fixpoint_nodeset (make_predecessors code successors) (fun l => l) - transf entrypoints. + transf exit_points. Theorem fixpoint_solution: forall res n instr s, fixpoint = Some res -> code!n = Some instr -> In s (successors instr) -> + (forall n a, code!n = None -> L.eq (transf n a) L.bot) -> L.ge res!!n (transf s res!!s). Proof. intros. exploit (make_predecessors_correct_2 code); eauto. intros [l [P Q]]. - eapply DS.fixpoint_solution; eauto. + destruct code!s as [instr'|] eqn:CS. +- exploit reachable_exit_points_predecessor. eexact CS. intros (ep & U & V). + unfold fixpoint in H. eapply DS.fixpoint_nodeset_solution; eauto. +- apply L.ge_trans with L.bot. apply L.ge_bot. + apply L.ge_refl. apply L.eq_sym. auto. Qed. -Theorem fixpoint_entry: - forall res n v, - fixpoint = Some res -> - In (n, v) entrypoints -> - L.ge res!!n v. -Proof. - intros. eapply DS.fixpoint_entry. eexact H. auto. -Qed. +(** The alternate solver that starts with all nodes of the CFG instead + of just the exit points. *) -Theorem fixpoint_invariant: - forall (P: L.t -> Prop), - P L.bot -> - (forall x y, P x -> P y -> P (L.lub x y)) -> - (forall pc x, P x -> P (transf pc x)) -> - (forall n v, In (n, v) entrypoints -> P v) -> - forall res pc, - fixpoint = Some res -> - P res!!pc. +Definition fixpoint_allnodes := + DS.fixpoint_allnodes + (make_predecessors code successors) (fun l => l) + transf. + +Theorem fixpoint_allnodes_solution: + forall res n instr s, + fixpoint_allnodes = Some res -> + code!n = Some instr -> In s (successors instr) -> + L.ge res!!n (transf s res!!s). Proof. - intros. - eapply DS.fixpoint_invariant with (code := make_predecessors code successors) (transf := transf); eauto. + intros. + exploit (make_predecessors_correct_2 code); eauto. intros [l [P Q]]. + unfold fixpoint_allnodes in H. + eapply DS.fixpoint_allnodes_solution; eauto. Qed. End Kildall. @@ -861,24 +1162,24 @@ Definition result := PMap.t L.t. *) Record state : Type := mkstate - { st_in: result; st_wrk: list positive }. + { aval: result; worklist: list positive }. (** The ``extended basic block'' algorithm, in pseudo-code, is as follows: << - st_wrk := the set of all points n having multiple predecessors - st_in := the mapping n -> L.top + worklist := the set of all points n having multiple predecessors + aval := the mapping n -> L.top - while st_wrk is not empty, do - extract a node n from st_wrk - compute out = transf n st_in[n] + while worklist is not empty, do + extract a node n from worklist + compute out = transf n aval[n] for each successor s of n: if s has only one predecessor (namely, n): - st_in[s] := out - st_wrk := st_wrk union {s} + aval[s] := out + worklist := worklist union {s} end if end for end while - return st_in + return aval >> **) @@ -892,22 +1193,22 @@ Fixpoint propagate_successors propagate_successors bb sl l st else propagate_successors bb sl l - (mkstate (PMap.set s1 l st.(st_in)) - (s1 :: st.(st_wrk))) + (mkstate (PMap.set s1 l st.(aval)) + (s1 :: st.(worklist))) end. Definition step (bb: bbmap) (st: state) : result + state := - match st.(st_wrk) with - | nil => inl _ st.(st_in) + match st.(worklist) with + | nil => inl _ st.(aval) | pc :: rem => match code!pc with | None => - inr _ (mkstate st.(st_in) rem) + inr _ (mkstate st.(aval) rem) | Some instr => inr _ (propagate_successors bb (successors instr) - (transf pc st.(st_in)!!pc) - (mkstate st.(st_in) rem)) + (transf pc st.(aval)!!pc) + (mkstate st.(aval) rem)) end end. @@ -989,34 +1290,34 @@ Qed. *) Definition state_invariant (st: state) : Prop := - (forall n, basic_block_map n = true -> st.(st_in)!!n = L.top) + (forall n, basic_block_map n = true -> st.(aval)!!n = L.top) /\ (forall n, - In n st.(st_wrk) \/ + In n st.(worklist) \/ (forall instr s, code!n = Some instr -> In s (successors instr) -> - L.ge st.(st_in)!!s (transf n st.(st_in)!!n))). + L.ge st.(aval)!!s (transf n st.(aval)!!n))). Lemma propagate_successors_charact1: forall bb succs l st, - incl st.(st_wrk) - (propagate_successors bb succs l st).(st_wrk). + incl st.(worklist) + (propagate_successors bb succs l st).(worklist). Proof. induction succs; simpl; intros. apply incl_refl. case (bb a). auto. - apply incl_tran with (a :: st_wrk st). + apply incl_tran with (a :: worklist st). apply incl_tl. apply incl_refl. - set (st1 := (mkstate (PMap.set a l (st_in st)) (a :: st_wrk st))). - change (a :: st_wrk st) with (st_wrk st1). + set (st1 := (mkstate (PMap.set a l (aval st)) (a :: worklist st))). + change (a :: worklist st) with (worklist st1). auto. Qed. Lemma propagate_successors_charact2: forall bb succs l st n, let st' := propagate_successors bb succs l st in - (In n succs -> bb n = false -> In n st'.(st_wrk) /\ st'.(st_in)!!n = l) -/\ (~In n succs \/ bb n = true -> st'.(st_in)!!n = st.(st_in)!!n). + (In n succs -> bb n = false -> In n st'.(worklist) /\ st'.(aval)!!n = l) +/\ (~In n succs \/ bb n = true -> st'.(aval)!!n = st.(aval)!!n). Proof. induction succs; simpl; intros. (* Base case *) @@ -1027,7 +1328,7 @@ Proof. split; intros. apply U; auto. elim H0; intro. subst a. congruence. auto. apply V. tauto. - set (st1 := mkstate (PMap.set a l (st_in st)) (a :: st_wrk st)). + set (st1 := mkstate (PMap.set a l (aval st)) (a :: worklist st)). elim (IHsuccs l st1 n); intros U V. split; intros. elim H0; intros. @@ -1069,7 +1370,7 @@ Proof. right; intros. assert (instr0 = instr) by congruence. subst instr0. elim (U s); intros C D. - replace (st1.(st_in)!!pc) with res!!pc. fold l. + replace (st1.(aval)!!pc) with res!!pc. fold l. destruct (basic_block_map s) eqn:BB. rewrite D. simpl. rewrite INV1. apply L.top_ge. auto. tauto. elim (C H0 (refl_equal _)). intros X Y. rewrite Y. apply L.refl_ge. @@ -1082,7 +1383,7 @@ Proof. (* Case 2.1: n was already in worklist, still is *) left. apply V. simpl. tauto. (* Case 2.2: n was not in worklist *) - assert (INV3: forall s instr', code!n = Some instr' -> In s (successors instr') -> st1.(st_in)!!s = res!!s). + assert (INV3: forall s instr', code!n = Some instr' -> In s (successors instr') -> st1.(aval)!!s = res!!s). (* Amazingly, successors of n do not change. The only way they could change is if they were successors of pc as well, but that gives them two different predecessors, so @@ -1181,7 +1482,7 @@ Qed. (** ** Preservation of a property over solutions *) Definition Pstate (st: state) : Prop := - forall pc, P st.(st_in)!!pc. + forall pc, P st.(aval)!!pc. Lemma propagate_successors_P: forall bb l, @@ -1204,9 +1505,9 @@ Proof. unfold fixpoint; intros. pattern res. eapply (PrimIter.iterate_prop _ _ (step basic_block_map) Pstate). - intros st PS. unfold step. destruct (st.(st_wrk)). + intros st PS. unfold step. destruct (st.(worklist)). apply PS. - assert (PS2: Pstate (mkstate st.(st_in) l)). + assert (PS2: Pstate (mkstate st.(aval) l)). red; intro; simpl. apply PS. destruct (code!p) as [instr|] eqn:CODE. apply propagate_successors_P. eauto. auto. @@ -1239,16 +1540,23 @@ Require Import Heaps. Module NodeSetForward <: NODE_SET. Definition t := PHeap.t. + Definition empty := PHeap.empty. Definition add (n: positive) (s: t) : t := PHeap.insert n s. Definition pick (s: t) := match PHeap.findMax s with | Some n => Some(n, PHeap.deleteMax s) | None => None end. - Definition initial {A: Type} (code: PTree.t A) := + Definition all_nodes {A: Type} (code: PTree.t A) := PTree.fold (fun s pc instr => PHeap.insert pc s) code PHeap.empty. Definition In := PHeap.In. + Lemma empty_spec: + forall n, ~In n empty. + Proof. + intros. apply PHeap.In_empty. + Qed. + Lemma add_spec: forall n n' s, In n' (add n s) <-> n = n' \/ In n' s. Proof. @@ -1273,9 +1581,9 @@ Module NodeSetForward <: NODE_SET. congruence. Qed. - Lemma initial_spec: + Lemma all_nodes_spec: forall A (code: PTree.t A) n instr, - code!n = Some instr -> In n (initial code). + code!n = Some instr -> In n (all_nodes code). Proof. intros A code n instr. apply PTree_Properties.fold_rec with @@ -1292,16 +1600,21 @@ End NodeSetForward. Module NodeSetBackward <: NODE_SET. Definition t := PHeap.t. + Definition empty := PHeap.empty. Definition add (n: positive) (s: t) : t := PHeap.insert n s. Definition pick (s: t) := match PHeap.findMin s with | Some n => Some(n, PHeap.deleteMin s) | None => None end. - Definition initial {A: Type} (code: PTree.t A) := + Definition all_nodes {A: Type} (code: PTree.t A) := PTree.fold (fun s pc instr => PHeap.insert pc s) code PHeap.empty. Definition In := PHeap.In. + Lemma empty_spec: + forall n, ~In n empty. + Proof NodeSetForward.empty_spec. + Lemma add_spec: forall n n' s, In n' (add n s) <-> n = n' \/ In n' s. Proof NodeSetForward.add_spec. @@ -1324,9 +1637,9 @@ Module NodeSetBackward <: NODE_SET. congruence. Qed. - Lemma initial_spec: + Lemma all_nodes_spec: forall A (code: PTree.t A) n instr, - code!n = Some instr -> In n (initial code). - Proof NodeSetForward.initial_spec. + code!n = Some instr -> In n (all_nodes code). + Proof NodeSetForward.all_nodes_spec. End NodeSetBackward. diff --git a/backend/Linearize.v b/backend/Linearize.v index a4d1c0d..b1102e2 100644 --- a/backend/Linearize.v +++ b/backend/Linearize.v @@ -93,7 +93,7 @@ Definition reachable_aux (f: LTL.function) : option (PMap.t bool) := DS.fixpoint (LTL.fn_code f) successors_block (fun pc r => r) - ((f.(fn_entrypoint), true) :: nil). + f.(fn_entrypoint) true. Definition reachable (f: LTL.function) : PMap.t bool := match reachable_aux f with diff --git a/backend/Linearizeproof.v b/backend/Linearizeproof.v index 93d38dd..3b22fc6 100644 --- a/backend/Linearizeproof.v +++ b/backend/Linearizeproof.v @@ -108,7 +108,7 @@ Proof. caseEq (reachable_aux f). unfold reachable_aux; intros reach A. assert (LBoolean.ge reach!!(f.(fn_entrypoint)) true). - eapply DS.fixpoint_entry. eexact A. auto with coqlib. + eapply DS.fixpoint_entry. eexact A. auto. unfold LBoolean.ge in H. tauto. intros. apply PMap.gi. Qed. @@ -126,7 +126,7 @@ Proof. unfold reachable_aux. intro reach; intros. assert (LBoolean.ge reach!!pc' reach!!pc). change (reach!!pc) with ((fun pc r => r) pc (reach!!pc)). - eapply DS.fixpoint_solution; eauto. + eapply DS.fixpoint_solution; eauto. intros; apply DS.L.eq_refl. elim H3; intro. congruence. auto. intros. apply PMap.gi. Qed. diff --git a/backend/Liveness.v b/backend/Liveness.v index 23faf41..3a5bde9 100644 --- a/backend/Liveness.v +++ b/backend/Liveness.v @@ -110,7 +110,7 @@ Module RegsetLat := LFSet(Regset). Module DS := Backward_Dataflow_Solver(RegsetLat)(NodeSetBackward). Definition analyze (f: function): option (PMap.t Regset.t) := - DS.fixpoint f.(fn_code) successors_instr (transfer f) nil. + DS.fixpoint f.(fn_code) successors_instr (transfer f). (** Basic property of the liveness information computed by [analyze]. *) @@ -122,6 +122,7 @@ Lemma analyze_solution: Regset.Subset (transfer f s live!!s) live!!n. Proof. unfold analyze; intros. eapply DS.fixpoint_solution; eauto. + intros. unfold transfer; rewrite H2. apply DS.L.eq_refl. Qed. (** Given an RTL function, compute (for every PC) the list of diff --git a/backend/NeedDomain.v b/backend/NeedDomain.v new file mode 100644 index 0000000..568c80e --- /dev/null +++ b/backend/NeedDomain.v @@ -0,0 +1,1515 @@ +(* *********************************************************************) +(* *) +(* 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. *) +(* *) +(* *********************************************************************) + +(** Abstract domain for neededness analysis *) + +Require Import Coqlib. +Require Import Maps. +Require Import IntvSets. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Globalenvs. +Require Import Events. +Require Import Lattice. +Require Import Registers. +Require Import ValueDomain. +Require Import Op. +Require Import RTL. + +(** * Neededness for values *) + +Inductive nval : Type := + | Nothing (**r value is entirely unused *) + | I (m: int) (**r only need the bits that are 1 in [m] *) + | Fsingle (**r only need the value after conversion to single float *) + | All. (**r every bit of the value is used *) + +Definition eq_nval (x y: nval) : {x=y} + {x<>y}. +Proof. + decide equality. apply Int.eq_dec. +Defined. + +(** ** Agreement between two values relative to a need. *) + +Definition iagree (p q mask: int) : Prop := + forall i, 0 <= i < Int.zwordsize -> Int.testbit mask i = true -> + Int.testbit p i = Int.testbit q i. + +Fixpoint vagree (v w: val) (x: nval) {struct x}: Prop := + match x with + | Nothing => True + | I m => + match v, w with + | Vint p, Vint q => iagree p q m + | Vint p, _ => False + | _, _ => True + end + | Fsingle => + match v, w with + | Vfloat f, Vfloat g => Float.singleoffloat f = Float.singleoffloat g + | Vfloat _, _ => False + | _, _ => True + end + | All => Val.lessdef v w + end. + +Lemma vagree_same: forall v x, vagree v v x. +Proof. + intros. destruct x; simpl; auto; destruct v; auto. red; auto. +Qed. + +Lemma vagree_lessdef: forall v w x, Val.lessdef v w -> vagree v w x. +Proof. + intros. inv H. apply vagree_same. destruct x; simpl; auto. +Qed. + +Lemma lessdef_vagree: forall v w, vagree v w All -> Val.lessdef v w. +Proof. + intros. simpl in H. auto. +Qed. + +Hint Resolve vagree_same vagree_lessdef lessdef_vagree: na. + +Definition vagree_list (vl1 vl2: list val) (nv: nval) : Prop := + list_forall2 (fun v1 v2 => vagree v1 v2 nv) vl1 vl2. + +Lemma lessdef_vagree_list: + forall vl1 vl2, vagree_list vl1 vl2 All -> Val.lessdef_list vl1 vl2. +Proof. + induction 1; constructor; auto with na. +Qed. + +Lemma vagree_lessdef_list: + forall x vl1 vl2, Val.lessdef_list vl1 vl2 -> vagree_list vl1 vl2 x. +Proof. + induction 1; constructor; auto with na. +Qed. + +Hint Resolve lessdef_vagree_list vagree_lessdef_list: na. + +(** ** Ordering and least upper bound between value needs *) + +Inductive nge: nval -> nval -> Prop := + | nge_nothing: forall x, nge All x + | nge_all: forall x, nge x Nothing + | nge_int: forall m1 m2, + (forall i, 0 <= i < Int.zwordsize -> Int.testbit m2 i = true -> Int.testbit m1 i = true) -> + nge (I m1) (I m2) + | nge_single: + nge Fsingle Fsingle. + +Lemma nge_refl: forall x, nge x x. +Proof. + destruct x; constructor; auto. +Qed. + +Hint Constructors nge: na. +Hint Resolve nge_refl: na. + +Lemma nge_trans: forall x y, nge x y -> forall z, nge y z -> nge x z. +Proof. + induction 1; intros w VG; inv VG; eauto with na. +Qed. + +Lemma nge_agree: + forall v w x y, nge x y -> vagree v w x -> vagree v w y. +Proof. + induction 1; simpl; auto. +- destruct v; auto with na. +- destruct v, w; intuition. red; auto. +Qed. + +Definition nlub (x y: nval) : nval := + match x, y with + | Nothing, _ => y + | _, Nothing => x + | I m1, I m2 => I (Int.or m1 m2) + | Fsingle, Fsingle => Fsingle + | _, _ => All + end. + +Lemma nge_lub_l: + forall x y, nge (nlub x y) x. +Proof. + unfold nlub; destruct x, y; auto with na. + constructor. intros. autorewrite with ints; auto. rewrite H0; auto. +Qed. + +Lemma nge_lub_r: + forall x y, nge (nlub x y) y. +Proof. + unfold nlub; destruct x, y; auto with na. + constructor. intros. autorewrite with ints; auto. rewrite H0. apply orb_true_r; auto. +Qed. + +(** ** Properties of agreement between integers *) + +Lemma iagree_refl: + forall p m, iagree p p m. +Proof. + intros; red; auto. +Qed. + +Remark eq_same_bits: + forall i x y, x = y -> Int.testbit x i = Int.testbit y i. +Proof. + intros; congruence. +Qed. + +Lemma iagree_and_eq: + forall x y mask, + iagree x y mask <-> Int.and x mask = Int.and y mask. +Proof. + intros; split; intros. +- Int.bit_solve. specialize (H i H0). + destruct (Int.testbit mask i). + rewrite ! andb_true_r; auto. + rewrite ! andb_false_r; auto. +- red; intros. exploit (eq_same_bits i); eauto; autorewrite with ints; auto. + rewrite H1. rewrite ! andb_true_r; auto. +Qed. + +Lemma iagree_mone: + forall p q, iagree p q Int.mone -> p = q. +Proof. + intros. rewrite iagree_and_eq in H. rewrite ! Int.and_mone in H. auto. +Qed. + +Lemma iagree_zero: + forall p q, iagree p q Int.zero. +Proof. + intros. rewrite iagree_and_eq. rewrite ! Int.and_zero; auto. +Qed. + +Lemma iagree_and: + forall x y n m, + iagree x y (Int.and m n) -> iagree (Int.and x n) (Int.and y n) m. +Proof. + intros. rewrite iagree_and_eq in *. rewrite ! Int.and_assoc. + rewrite (Int.and_commut n). auto. +Qed. + +Lemma iagree_not: + forall x y m, iagree x y m -> iagree (Int.not x) (Int.not y) m. +Proof. + intros; red; intros; autorewrite with ints; auto. f_equal; auto. +Qed. + +Lemma iagree_not': + forall x y m, iagree (Int.not x) (Int.not y) m -> iagree x y m. +Proof. + intros. rewrite <- (Int.not_involutive x). rewrite <- (Int.not_involutive y). + apply iagree_not; auto. +Qed. + +Lemma iagree_or: + forall x y n m, + iagree x y (Int.and m (Int.not n)) -> iagree (Int.or x n) (Int.or y n) m. +Proof. + intros. apply iagree_not'. rewrite ! Int.not_or_and_not. apply iagree_and. + apply iagree_not; auto. +Qed. + +Lemma iagree_bitwise_binop: + forall sem f, + (forall x y i, 0 <= i < Int.zwordsize -> + Int.testbit (f x y) i = sem (Int.testbit x i) (Int.testbit y i)) -> + forall x1 x2 y1 y2 m, + iagree x1 y1 m -> iagree x2 y2 m -> iagree (f x1 x2) (f y1 y2) m. +Proof. + intros; red; intros. rewrite ! H by auto. f_equal; auto. +Qed. + +Lemma iagree_shl: + forall x y m n, + iagree x y (Int.shru m n) -> iagree (Int.shl x n) (Int.shl y n) m. +Proof. + intros; red; intros. autorewrite with ints; auto. + destruct (zlt i (Int.unsigned n)). +- auto. +- generalize (Int.unsigned_range n); intros. + apply H. omega. rewrite Int.bits_shru by omega. + replace (i - Int.unsigned n + Int.unsigned n) with i by omega. + rewrite zlt_true by omega. auto. +Qed. + +Lemma iagree_shru: + forall x y m n, + iagree x y (Int.shl m n) -> iagree (Int.shru x n) (Int.shru y n) m. +Proof. + intros; red; intros. autorewrite with ints; auto. + destruct (zlt (i + Int.unsigned n) Int.zwordsize). +- generalize (Int.unsigned_range n); intros. + apply H. omega. rewrite Int.bits_shl by omega. + replace (i + Int.unsigned n - Int.unsigned n) with i by omega. + rewrite zlt_false by omega. auto. +- auto. +Qed. + +Lemma iagree_shr_1: + forall x y m n, + Int.shru (Int.shl m n) n = m -> + iagree x y (Int.shl m n) -> iagree (Int.shr x n) (Int.shr y n) m. +Proof. + intros; red; intros. rewrite <- H in H2. rewrite Int.bits_shru in H2 by auto. + rewrite ! Int.bits_shr by auto. + destruct (zlt (i + Int.unsigned n) Int.zwordsize). +- apply H0; auto. generalize (Int.unsigned_range n); omega. +- discriminate. +Qed. + +Lemma iagree_shr: + forall x y m n, + iagree x y (Int.or (Int.shl m n) (Int.repr Int.min_signed)) -> + iagree (Int.shr x n) (Int.shr y n) m. +Proof. + intros; red; intros. rewrite ! Int.bits_shr by auto. + generalize (Int.unsigned_range n); intros. + set (j := if zlt (i + Int.unsigned n) Int.zwordsize + then i + Int.unsigned n + else Int.zwordsize - 1). + assert (0 <= j < Int.zwordsize). + { unfold j; destruct (zlt (i + Int.unsigned n) Int.zwordsize); omega. } + apply H; auto. autorewrite with ints; auto. apply orb_true_intro. + unfold j; destruct (zlt (i + Int.unsigned n) Int.zwordsize). +- left. rewrite zlt_false by omega. + replace (i + Int.unsigned n - Int.unsigned n) with i by omega. + auto. +- right. reflexivity. +Qed. + +Lemma iagree_rol: + forall p q m amount, + iagree p q (Int.ror m amount) -> + iagree (Int.rol p amount) (Int.rol q amount) m. +Proof. + intros. assert (Int.zwordsize > 0) by (compute; auto). + red; intros. rewrite ! Int.bits_rol by auto. apply H. + apply Z_mod_lt; auto. + rewrite Int.bits_ror. + replace (((i - Int.unsigned amount) mod Int.zwordsize + Int.unsigned amount) + mod Int.zwordsize) with i. auto. + apply Int.eqmod_small_eq with Int.zwordsize; auto. + apply Int.eqmod_trans with ((i - Int.unsigned amount) + Int.unsigned amount). + apply Int.eqmod_refl2; omega. + eapply Int.eqmod_trans. 2: apply Int.eqmod_mod; auto. + apply Int.eqmod_add. + apply Int.eqmod_mod; auto. + apply Int.eqmod_refl. + apply Z_mod_lt; auto. + apply Z_mod_lt; auto. +Qed. + +Lemma iagree_ror: + forall p q m amount, + iagree p q (Int.rol m amount) -> + iagree (Int.ror p amount) (Int.ror q amount) m. +Proof. + intros. rewrite ! Int.ror_rol_neg by apply int_wordsize_divides_modulus. + apply iagree_rol. + rewrite Int.ror_rol_neg by apply int_wordsize_divides_modulus. + rewrite Int.neg_involutive; auto. +Qed. + +Lemma eqmod_iagree: + forall m x y, + Int.eqmod (two_p (Int.size m)) x y -> + iagree (Int.repr x) (Int.repr y) m. +Proof. + intros. set (p := nat_of_Z (Int.size m)). + generalize (Int.size_range m); intros RANGE. + assert (EQ: Int.size m = Z_of_nat p). { symmetry; apply nat_of_Z_eq. omega. } + rewrite EQ in H; rewrite <- two_power_nat_two_p in H. + red; intros. rewrite ! Int.testbit_repr by auto. + destruct (zlt i (Int.size m)). + eapply Int.same_bits_eqmod; eauto. omega. + assert (Int.testbit m i = false) by (eapply Int.bits_size_2; omega). + congruence. +Qed. + +Definition complete_mask (m: int) := Int.zero_ext (Int.size m) Int.mone. + +Lemma iagree_eqmod: + forall x y m, + iagree x y (complete_mask m) -> + Int.eqmod (two_p (Int.size m)) (Int.unsigned x) (Int.unsigned y). +Proof. + intros. set (p := nat_of_Z (Int.size m)). + generalize (Int.size_range m); intros RANGE. + assert (EQ: Int.size m = Z_of_nat p). { symmetry; apply nat_of_Z_eq. omega. } + rewrite EQ; rewrite <- two_power_nat_two_p. + apply Int.eqmod_same_bits. intros. apply H. omega. + unfold complete_mask. rewrite Int.bits_zero_ext by omega. + rewrite zlt_true by omega. rewrite Int.bits_mone by omega. auto. +Qed. + +Lemma complete_mask_idem: + forall m, complete_mask (complete_mask m) = complete_mask m. +Proof. + unfold complete_mask; intros. destruct (Int.eq_dec m Int.zero). ++ subst m; reflexivity. ++ assert (Int.unsigned m <> 0). + { red; intros; elim n. rewrite <- (Int.repr_unsigned m). rewrite H; auto. } + assert (0 < Int.size m). + { apply Int.Zsize_pos'. generalize (Int.unsigned_range m); omega. } + generalize (Int.size_range m); intros. + f_equal. apply Int.bits_size_4. tauto. + rewrite Int.bits_zero_ext by omega. rewrite zlt_true by omega. + apply Int.bits_mone; omega. + intros. rewrite Int.bits_zero_ext by omega. apply zlt_false; omega. +Qed. + +(** ** Abstract operations over value needs. *) + +Ltac InvAgree := + simpl vagree in *; + repeat ( + auto || exact Logic.I || + match goal with + | [ H: False |- _ ] => contradiction + | [ H: match ?v with Vundef => _ | Vint _ => _ | Vlong _ => _ | Vfloat _ => _ | Vptr _ _ => _ end |- _ ] => destruct v + end). + +(** And immediate, or immediate *) + +Definition andimm (x: nval) (n: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.and m n) + | All => I n + end. + +Lemma andimm_sound: + forall v w x n, + vagree v w (andimm x n) -> + vagree (Val.and v (Vint n)) (Val.and w (Vint n)) x. +Proof. + unfold andimm; intros; destruct x; simpl in *; unfold Val.and. +- auto. +- InvAgree. apply iagree_and; auto. +- destruct v; destruct w; tauto. +- InvAgree. rewrite iagree_and_eq in H. rewrite H; auto. +Qed. + +Definition orimm (x: nval) (n: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.and m (Int.not n)) + | _ => I (Int.not n) + end. + +Lemma orimm_sound: + forall v w x n, + vagree v w (orimm x n) -> + vagree (Val.or v (Vint n)) (Val.or w (Vint n)) x. +Proof. + unfold orimm; intros; destruct x; simpl in *. +- auto. +- unfold Val.or; InvAgree. apply iagree_or; auto. +- destruct v; destruct w; tauto. +- InvAgree. simpl. apply Val.lessdef_same. f_equal. apply iagree_mone. + apply iagree_or. rewrite Int.and_commut. rewrite Int.and_mone. auto. +Qed. + +(** Bitwise operations: and, or, xor, not *) + +Definition bitwise (x: nval) := + match x with + | Fsingle => Nothing + | _ => x + end. + +Remark bitwise_idem: forall nv, bitwise (bitwise nv) = bitwise nv. +Proof. destruct nv; auto. Qed. + +Lemma vagree_bitwise_binop: + forall f, + (forall p1 p2 q1 q2 m, + iagree p1 q1 m -> iagree p2 q2 m -> iagree (f p1 p2) (f q1 q2) m) -> + forall v1 w1 v2 w2 x, + vagree v1 w1 (bitwise x) -> vagree v2 w2 (bitwise x) -> + vagree (match v1, v2 with Vint i1, Vint i2 => Vint(f i1 i2) | _, _ => Vundef end) + (match w1, w2 with Vint i1, Vint i2 => Vint(f i1 i2) | _, _ => Vundef end) + x. +Proof. + unfold bitwise; intros. destruct x; simpl in *. +- auto. +- InvAgree. +- destruct v1; auto. destruct v2; auto. +- inv H0; auto. inv H1; auto. destruct w1; auto. +Qed. + +Lemma and_sound: + forall v1 w1 v2 w2 x, + vagree v1 w1 (bitwise x) -> vagree v2 w2 (bitwise x) -> + vagree (Val.and v1 v2) (Val.and w1 w2) x. +Proof (vagree_bitwise_binop Int.and (iagree_bitwise_binop andb Int.and Int.bits_and)). + +Lemma or_sound: + forall v1 w1 v2 w2 x, + vagree v1 w1 (bitwise x) -> vagree v2 w2 (bitwise x) -> + vagree (Val.or v1 v2) (Val.or w1 w2) x. +Proof (vagree_bitwise_binop Int.or (iagree_bitwise_binop orb Int.or Int.bits_or)). + +Lemma xor_sound: + forall v1 w1 v2 w2 x, + vagree v1 w1 (bitwise x) -> vagree v2 w2 (bitwise x) -> + vagree (Val.xor v1 v2) (Val.xor w1 w2) x. +Proof (vagree_bitwise_binop Int.xor (iagree_bitwise_binop xorb Int.xor Int.bits_xor)). + +Lemma notint_sound: + forall v w x, + vagree v w (bitwise x) -> vagree (Val.notint v) (Val.notint w) x. +Proof. + intros. rewrite ! Val.not_xor. apply xor_sound; auto with na. +Qed. + +(** Shifts and rotates *) + +Definition shlimm (x: nval) (n: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.shru m n) + | All => I (Int.shru Int.mone n) + end. + +Lemma shlimm_sound: + forall v w x n, + vagree v w (shlimm x n) -> + vagree (Val.shl v (Vint n)) (Val.shl w (Vint n)) x. +Proof. + unfold shlimm; intros. unfold Val.shl. + destruct (Int.ltu n Int.iwordsize). + destruct x; simpl in *. +- auto. +- InvAgree. apply iagree_shl; auto. +- destruct v; destruct w; auto. +- InvAgree. apply Val.lessdef_same. f_equal. apply iagree_mone. apply iagree_shl; auto. +- destruct v; auto with na. +Qed. + +Definition shruimm (x: nval) (n: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.shl m n) + | All => I (Int.shl Int.mone n) + end. + +Lemma shruimm_sound: + forall v w x n, + vagree v w (shruimm x n) -> + vagree (Val.shru v (Vint n)) (Val.shru w (Vint n)) x. +Proof. + unfold shruimm; intros. unfold Val.shru. + destruct (Int.ltu n Int.iwordsize). + destruct x; simpl in *. +- auto. +- InvAgree. apply iagree_shru; auto. +- destruct v; destruct w; auto. +- InvAgree. apply Val.lessdef_same. f_equal. apply iagree_mone. apply iagree_shru; auto. +- destruct v; auto with na. +Qed. + +Definition shrimm (x: nval) (n: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (let m' := Int.shl m n in + if Int.eq_dec (Int.shru m' n) m + then m' + else Int.or m' (Int.repr Int.min_signed)) + | All => I (Int.or (Int.shl Int.mone n) (Int.repr Int.min_signed)) + end. + +Lemma shrimm_sound: + forall v w x n, + vagree v w (shrimm x n) -> + vagree (Val.shr v (Vint n)) (Val.shr w (Vint n)) x. +Proof. + unfold shrimm; intros. unfold Val.shr. + destruct (Int.ltu n Int.iwordsize). + destruct x; simpl in *. +- auto. +- InvAgree. + destruct (Int.eq_dec (Int.shru (Int.shl m n) n) m). + apply iagree_shr_1; auto. + apply iagree_shr; auto. +- destruct v; destruct w; auto. +- InvAgree. apply Val.lessdef_same. f_equal. apply iagree_mone. apply iagree_shr. auto. +- destruct v; auto with na. +Qed. + +Definition rolm (x: nval) (amount mask: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.ror (Int.and m mask) amount) + | _ => I (Int.ror mask amount) + end. + +Lemma rolm_sound: + forall v w x amount mask, + vagree v w (rolm x amount mask) -> + vagree (Val.rolm v amount mask) (Val.rolm w amount mask) x. +Proof. + unfold rolm; intros; destruct x; simpl in *. +- auto. +- unfold Val.rolm; InvAgree. unfold Int.rolm. + apply iagree_and. apply iagree_rol. auto. +- unfold Val.rolm; destruct v, w; auto. +- unfold Val.rolm; InvAgree. apply Val.lessdef_same. f_equal. apply iagree_mone. + unfold Int.rolm. apply iagree_and. apply iagree_rol. rewrite Int.and_commut. + rewrite Int.and_mone. auto. +Qed. + +Definition ror (x: nval) (amount: int) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.rol m amount) + | All => All + end. + +Lemma ror_sound: + forall v w x n, + vagree v w (ror x n) -> + vagree (Val.ror v (Vint n)) (Val.ror w (Vint n)) x. +Proof. + unfold ror; intros. unfold Val.ror. + destruct (Int.ltu n Int.iwordsize). + destruct x; simpl in *. +- auto. +- InvAgree. apply iagree_ror; auto. +- destruct v, w; auto. +- inv H; auto. +- destruct v; auto with na. +Qed. + +(** Modular arithmetic operations: add, mul. + (But not subtraction because of the pointer - pointer case. *) + +Definition modarith (x: nval) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (complete_mask m) + | All => All + end. + +Lemma add_sound: + forall v1 w1 v2 w2 x, + vagree v1 w1 (modarith x) -> vagree v2 w2 (modarith x) -> + vagree (Val.add v1 v2) (Val.add w1 w2) x. +Proof. + unfold modarith; intros. destruct x; simpl in *. +- auto. +- unfold Val.add; InvAgree. apply eqmod_iagree. apply Int.eqmod_add; apply iagree_eqmod; auto. +- unfold Val.add; destruct v1, w1; auto; destruct v2, w2; auto. +- inv H; auto. inv H0; auto. destruct w1; auto. +Qed. + +Remark modarith_idem: forall nv, modarith (modarith nv) = modarith nv. +Proof. + destruct nv; simpl; auto. f_equal; apply complete_mask_idem. +Qed. + +Lemma add_sound_2: + forall v1 w1 v2 w2 x, + vagree v1 w1 (modarith x) -> vagree v2 w2 (modarith x) -> + vagree (Val.add v1 v2) (Val.add w1 w2) (modarith x). +Proof. + intros. apply add_sound; rewrite modarith_idem; auto. +Qed. + +Lemma mul_sound: + forall v1 w1 v2 w2 x, + vagree v1 w1 (modarith x) -> vagree v2 w2 (modarith x) -> + vagree (Val.mul v1 v2) (Val.mul w1 w2) x. +Proof. + unfold mul, add; intros. destruct x; simpl in *. +- auto. +- unfold Val.mul; InvAgree. apply eqmod_iagree. apply Int.eqmod_mult; apply iagree_eqmod; auto. +- unfold Val.mul; destruct v1, w1; auto; destruct v2, w2; auto. +- inv H; auto. inv H0; auto. destruct w1; auto. +Qed. + +Lemma mul_sound_2: + forall v1 w1 v2 w2 x, + vagree v1 w1 (modarith x) -> vagree v2 w2 (modarith x) -> + vagree (Val.mul v1 v2) (Val.mul w1 w2) (modarith x). +Proof. + intros. apply mul_sound; rewrite modarith_idem; auto. +Qed. + +(** Conversions: zero extension, sign extension, single-of-float *) + +Definition zero_ext (n: Z) (x: nval) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.zero_ext n m) + | All => I (Int.zero_ext n Int.mone) + end. + +Lemma zero_ext_sound: + forall v w x n, + vagree v w (zero_ext n x) -> + 0 <= n -> + vagree (Val.zero_ext n v) (Val.zero_ext n w) x. +Proof. + unfold zero_ext; intros. + destruct x; simpl in *. +- auto. +- unfold Val.zero_ext; InvAgree. + red; intros. autorewrite with ints; try omega. + destruct (zlt i1 n); auto. apply H; auto. + autorewrite with ints; try omega. rewrite zlt_true; auto. +- unfold Val.zero_ext; destruct v; destruct w; auto. +- unfold Val.zero_ext; InvAgree; auto. apply Val.lessdef_same. f_equal. + Int.bit_solve; try omega. destruct (zlt i1 n); auto. apply H; auto. + autorewrite with ints; try omega. apply zlt_true; auto. +Qed. + +Definition sign_ext (n: Z) (x: nval) := + match x with + | Nothing | Fsingle => Nothing + | I m => I (Int.or (Int.zero_ext n m) (Int.shl Int.one (Int.repr (n - 1)))) + | All => I (Int.zero_ext n Int.mone) + end. + +Lemma sign_ext_sound: + forall v w x n, + vagree v w (sign_ext n x) -> + 0 < n < Int.zwordsize -> + vagree (Val.sign_ext n v) (Val.sign_ext n w) x. +Proof. + unfold sign_ext; intros. destruct x; simpl in *. +- auto. +- unfold Val.sign_ext; InvAgree. + red; intros. autorewrite with ints; try omega. + set (j := if zlt i1 n then i1 else n - 1). + assert (0 <= j < Int.zwordsize). + { unfold j; destruct (zlt i1 n); omega. } + apply H; auto. + autorewrite with ints; try omega. apply orb_true_intro. + unfold j; destruct (zlt i1 n). + left. rewrite zlt_true; auto. + right. rewrite Int.unsigned_repr. rewrite zlt_false by omega. + replace (n - 1 - (n - 1)) with 0 by omega. reflexivity. + generalize Int.wordsize_max_unsigned; omega. +- unfold Val.sign_ext; destruct v; destruct w; auto. +- unfold Val.sign_ext; InvAgree; auto. apply Val.lessdef_same. f_equal. + Int.bit_solve; try omega. + set (j := if zlt i1 n then i1 else n - 1). + assert (0 <= j < Int.zwordsize). + { unfold j; destruct (zlt i1 n); omega. } + apply H; auto. rewrite Int.bits_zero_ext; try omega. + rewrite zlt_true. apply Int.bits_mone; auto. + unfold j. destruct (zlt i1 n); omega. +Qed. + +Definition singleoffloat (x: nval) := + match x with + | Nothing | I _ => Nothing + | Fsingle | All => Fsingle + end. + +Lemma singleoffloat_sound: + forall v w x, + vagree v w (singleoffloat x) -> + vagree (Val.singleoffloat v) (Val.singleoffloat w) x. +Proof. + unfold singleoffloat; intros. destruct x; simpl in *. +- auto. +- unfold Val.singleoffloat; destruct v, w; auto. +- unfold Val.singleoffloat; InvAgree. congruence. +- unfold Val.singleoffloat; InvAgree; auto. rewrite H; auto. +Qed. + +(** The needs of a memory store concerning the value being stored. *) + +Definition store_argument (chunk: memory_chunk) := + match chunk with + | Mint8signed | Mint8unsigned => I (Int.repr 255) + | Mint16signed | Mint16unsigned => I (Int.repr 65535) + | Mfloat32 => Fsingle + | _ => All + end. + +Lemma store_argument_sound: + forall chunk v w, + vagree v w (store_argument chunk) -> + list_forall2 memval_lessdef (encode_val chunk v) (encode_val chunk w). +Proof. + intros. + assert (UNDEF: list_forall2 memval_lessdef + (list_repeat (size_chunk_nat chunk) Undef) + (encode_val chunk w)). + { + rewrite <- (encode_val_length chunk w). + apply repeat_Undef_inject_any. + } + assert (SAME: forall vl1 vl2, + vl1 = vl2 -> + list_forall2 memval_lessdef vl1 vl2). + { + intros. subst vl2. revert vl1. induction vl1; constructor; auto. + apply memval_lessdef_refl. + } + + intros. unfold store_argument in H; destruct chunk. +- InvAgree. apply SAME. simpl; f_equal. apply encode_int_8_mod. + change 8 with (Int.size (Int.repr 255)). apply iagree_eqmod; auto. +- InvAgree. apply SAME. simpl; f_equal. apply encode_int_8_mod. + change 8 with (Int.size (Int.repr 255)). apply iagree_eqmod; auto. +- InvAgree. apply SAME. simpl; f_equal. apply encode_int_16_mod. + change 16 with (Int.size (Int.repr 65535)). apply iagree_eqmod; auto. +- InvAgree. apply SAME. simpl; f_equal. apply encode_int_16_mod. + change 16 with (Int.size (Int.repr 65535)). apply iagree_eqmod; auto. +- apply encode_val_inject. rewrite val_inject_id; auto. +- apply encode_val_inject. rewrite val_inject_id; auto. +- InvAgree. apply SAME. simpl. + rewrite <- (Float.bits_of_singleoffloat f). + rewrite <- (Float.bits_of_singleoffloat f0). + congruence. +- apply encode_val_inject. rewrite val_inject_id; auto. +- apply encode_val_inject. rewrite val_inject_id; auto. +Qed. + +Lemma store_argument_load_result: + forall chunk v w, + vagree v w (store_argument chunk) -> + Val.lessdef (Val.load_result chunk v) (Val.load_result chunk w). +Proof. + unfold store_argument; intros; destruct chunk; + auto using Val.load_result_lessdef; InvAgree; simpl. +- apply sign_ext_sound with (v := Vint i) (w := Vint i0) (x := All) (n := 8). + auto. compute; auto. +- apply zero_ext_sound with (v := Vint i) (w := Vint i0) (x := All) (n := 8). + auto. omega. +- apply sign_ext_sound with (v := Vint i) (w := Vint i0) (x := All) (n := 16). + auto. compute; auto. +- apply zero_ext_sound with (v := Vint i) (w := Vint i0) (x := All) (n := 16). + auto. omega. +- apply singleoffloat_sound with (v := Vfloat f) (w := Vfloat f0) (x := All). + auto. +Qed. + +(** The needs of a comparison *) + +Definition maskzero (n: int) := I n. + +Lemma maskzero_sound: + forall v w n b, + vagree v w (maskzero n) -> + Val.maskzero_bool v n = Some b -> + Val.maskzero_bool w n = Some b. +Proof. + unfold maskzero; intros. + unfold Val.maskzero_bool; InvAgree; try discriminate. + inv H0. rewrite iagree_and_eq in H. rewrite H. auto. +Qed. + +(** The default abstraction: if the result is unused, the arguments are + unused; otherwise, the arguments are needed in full. *) + +Definition default (x: nval) := + match x with + | Nothing => Nothing + | _ => All + end. + +Section DEFAULT. + +Variable ge: genv. +Variable sp: block. +Variables m1 m2: mem. +Hypothesis PERM: forall b ofs k p, Mem.perm m1 b ofs k p -> Mem.perm m2 b ofs k p. + +Let valid_pointer_inj: + forall b1 ofs b2 delta, + inject_id b1 = Some(b2, delta) -> + Mem.valid_pointer m1 b1 (Int.unsigned ofs) = true -> + Mem.valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true. +Proof. + unfold inject_id; intros. inv H. rewrite Int.add_zero. + rewrite Mem.valid_pointer_nonempty_perm in *. eauto. +Qed. + +Let weak_valid_pointer_inj: + forall b1 ofs b2 delta, + inject_id b1 = Some(b2, delta) -> + Mem.weak_valid_pointer m1 b1 (Int.unsigned ofs) = true -> + Mem.weak_valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true. +Proof. + unfold inject_id; intros. inv H. rewrite Int.add_zero. + rewrite Mem.weak_valid_pointer_spec in *. + rewrite ! Mem.valid_pointer_nonempty_perm in *. + destruct H0; [left|right]; eauto. +Qed. + +Let weak_valid_pointer_no_overflow: + forall b1 ofs b2 delta, + inject_id b1 = Some(b2, delta) -> + Mem.weak_valid_pointer m1 b1 (Int.unsigned ofs) = true -> + 0 <= Int.unsigned ofs + Int.unsigned (Int.repr delta) <= Int.max_unsigned. +Proof. + unfold inject_id; intros. inv H. rewrite Zplus_0_r. apply Int.unsigned_range_2. +Qed. + +Let valid_different_pointers_inj: + forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2, + b1 <> b2 -> + Mem.valid_pointer m1 b1 (Int.unsigned ofs1) = true -> + Mem.valid_pointer m1 b2 (Int.unsigned ofs2) = true -> + inject_id b1 = Some (b1', delta1) -> + inject_id b2 = Some (b2', delta2) -> + b1' <> b2' \/ + Int.unsigned (Int.add ofs1 (Int.repr delta1)) <> Int.unsigned (Int.add ofs2 (Int.repr delta2)). +Proof. + unfold inject_id; intros. left; congruence. +Qed. + +Lemma default_needs_of_condition_sound: + forall cond args1 b args2, + eval_condition cond args1 m1 = Some b -> + vagree_list args1 args2 All -> + eval_condition cond args2 m2 = Some b. +Proof. + intros. apply eval_condition_inj with (f := inject_id) (m1 := m1) (vl1 := args1); auto. + apply val_list_inject_lessdef. apply lessdef_vagree_list. auto. +Qed. + +Lemma default_needs_of_operation_sound: + forall op args1 v1 args2 nv, + eval_operation ge (Vptr sp Int.zero) op args1 m1 = Some v1 -> + vagree_list args1 args2 (default nv) -> + nv <> Nothing -> + exists v2, + eval_operation ge (Vptr sp Int.zero) op args2 m2 = Some v2 + /\ vagree v1 v2 nv. +Proof. + intros. assert (default nv = All) by (destruct nv; simpl; congruence). + rewrite H2 in H0. + exploit (@eval_operation_inj _ _ ge inject_id). + intros. apply val_inject_lessdef. auto. + eassumption. auto. auto. auto. + apply val_inject_lessdef. instantiate (1 := Vptr sp Int.zero). instantiate (1 := Vptr sp Int.zero). auto. + apply val_list_inject_lessdef. apply lessdef_vagree_list. eauto. + eauto. + intros (v2 & A & B). exists v2; split; auto. + apply vagree_lessdef. apply val_inject_lessdef. auto. +Qed. + +End DEFAULT. + +(** ** Detecting operations that are redundant and can be turned into a move *) + +Definition andimm_redundant (x: nval) (n: int) := + match x with + | Nothing => true + | I m => Int.eq_dec (Int.and m (Int.not n)) Int.zero + | _ => false + end. + +Lemma andimm_redundant_sound: + forall v w x n, + andimm_redundant x n = true -> + vagree v w (andimm x n) -> + vagree (Val.and v (Vint n)) w x. +Proof. + unfold andimm_redundant; intros. destruct x; try discriminate. +- simpl; auto. +- InvBooleans. simpl in *; unfold Val.and; InvAgree. + red; intros. exploit (eq_same_bits i1); eauto. + autorewrite with ints; auto. rewrite H2; simpl; intros. + destruct (Int.testbit n i1) eqn:N; try discriminate. + rewrite andb_true_r. apply H0; auto. autorewrite with ints; auto. + rewrite H2, N; auto. +Qed. + +Definition orimm_redundant (x: nval) (n: int) := + match x with + | Nothing => true + | I m => Int.eq_dec (Int.and m n) Int.zero + | _ => false + end. + +Lemma orimm_redundant_sound: + forall v w x n, + orimm_redundant x n = true -> + vagree v w (orimm x n) -> + vagree (Val.or v (Vint n)) w x. +Proof. + unfold orimm_redundant; intros. destruct x; try discriminate. +- auto. +- InvBooleans. simpl in *; unfold Val.or; InvAgree. + apply iagree_not'. rewrite Int.not_or_and_not. + apply (andimm_redundant_sound (Vint (Int.not i)) (Vint (Int.not i0)) (I m) (Int.not n)). + simpl. rewrite Int.not_involutive. apply proj_sumbool_is_true. auto. + simpl. apply iagree_not; auto. +Qed. + +Definition rolm_redundant (x: nval) (amount mask: int) := + Int.eq_dec amount Int.zero && andimm_redundant x mask. + +Lemma rolm_redundant_sound: + forall v w x amount mask, + rolm_redundant x amount mask = true -> + vagree v w (rolm x amount mask) -> + vagree (Val.rolm v amount mask) w x. +Proof. + unfold rolm_redundant; intros; InvBooleans. subst amount. rewrite Val.rolm_zero. + apply andimm_redundant_sound; auto. + assert (forall n, Int.ror n Int.zero = n). + { intros. rewrite Int.ror_rol_neg by apply int_wordsize_divides_modulus. + rewrite Int.neg_zero. apply Int.rol_zero. } + unfold rolm, andimm in *. destruct x; auto. + rewrite H in H0. auto. + rewrite H in H0. auto. +Qed. + +Definition zero_ext_redundant (n: Z) (x: nval) := + match x with + | Nothing => true + | I m => Int.eq_dec (Int.zero_ext n m) m + | _ => false + end. + +Lemma zero_ext_redundant_sound: + forall v w x n, + zero_ext_redundant n x = true -> + vagree v w (zero_ext n x) -> + 0 <= n -> + vagree (Val.zero_ext n v) w x. +Proof. + unfold zero_ext_redundant; intros. destruct x; try discriminate. +- auto. +- simpl in *; InvAgree. simpl. InvBooleans. rewrite <- H. + red; intros; autorewrite with ints; try omega. + destruct (zlt i1 n). apply H0; auto. + rewrite Int.bits_zero_ext in H3 by omega. rewrite zlt_false in H3 by auto. discriminate. +Qed. + +Definition sign_ext_redundant (n: Z) (x: nval) := + match x with + | Nothing => true + | I m => Int.eq_dec (Int.zero_ext n m) m + | _ => false + end. + +Lemma sign_ext_redundant_sound: + forall v w x n, + sign_ext_redundant n x = true -> + vagree v w (sign_ext n x) -> + 0 < n -> + vagree (Val.sign_ext n v) w x. +Proof. + unfold sign_ext_redundant; intros. destruct x; try discriminate. +- auto. +- simpl in *; InvAgree. simpl. InvBooleans. rewrite <- H. + red; intros; autorewrite with ints; try omega. + destruct (zlt i1 n). apply H0; auto. + rewrite Int.bits_or; auto. rewrite H3; auto. + rewrite Int.bits_zero_ext in H3 by omega. rewrite zlt_false in H3 by auto. discriminate. +Qed. + +Definition singleoffloat_redundant (x: nval) := + match x with + | Nothing => true + | Fsingle => true + | _ => false + end. + +Lemma singleoffloat_redundant_sound: + forall v w x, + singleoffloat_redundant x = true -> + vagree v w (singleoffloat x) -> + vagree (Val.singleoffloat v) w x. +Proof. + unfold singleoffloat; intros. destruct x; try discriminate. +- auto. +- simpl in *; InvAgree. simpl. rewrite Float.singleoffloat_idem; auto. +Qed. + +(** * Neededness for register environments *) + +Module NVal <: SEMILATTICE. + + Definition t := nval. + Definition eq (x y: t) := (x = y). + Definition eq_refl: forall x, eq x x := (@refl_equal t). + Definition eq_sym: forall x y, eq x y -> eq y x := (@sym_equal t). + Definition eq_trans: forall x y z, eq x y -> eq y z -> eq x z := (@trans_equal t). + Definition beq (x y: t) : bool := proj_sumbool (eq_nval x y). + Lemma beq_correct: forall x y, beq x y = true -> eq x y. + Proof. unfold beq; intros. InvBooleans. auto. Qed. + Definition ge := nge. + Lemma ge_refl: forall x y, eq x y -> ge x y. + Proof. unfold eq, ge; intros. subst y. apply nge_refl. Qed. + Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z. + Proof. unfold ge; intros. eapply nge_trans; eauto. Qed. + Definition bot : t := Nothing. + Lemma ge_bot: forall x, ge x bot. + Proof. intros. constructor. Qed. + Definition lub := nlub. + Lemma ge_lub_left: forall x y, ge (lub x y) x. + Proof nge_lub_l. + Lemma ge_lub_right: forall x y, ge (lub x y) y. + Proof nge_lub_r. +End NVal. + +Module NE := LPMap1(NVal). + +Definition nenv := NE.t. + +Definition nreg (ne: nenv) (r: reg) := NE.get r ne. + +Definition eagree (e1 e2: regset) (ne: nenv) : Prop := + forall r, vagree e1#r e2#r (NE.get r ne). + +Lemma nreg_agree: + forall rs1 rs2 ne r, eagree rs1 rs2 ne -> vagree rs1#r rs2#r (nreg ne r). +Proof. + intros. apply H. +Qed. + +Hint Resolve nreg_agree: na. + +Lemma eagree_ge: + forall e1 e2 ne ne', + eagree e1 e2 ne -> NE.ge ne ne' -> eagree e1 e2 ne'. +Proof. + intros; red; intros. apply nge_agree with (NE.get r ne); auto. apply H0. +Qed. + +Lemma eagree_bot: + forall e1 e2, eagree e1 e2 NE.bot. +Proof. + intros; red; intros. rewrite NE.get_bot. exact Logic.I. +Qed. + +Lemma eagree_same: + forall e ne, eagree e e ne. +Proof. + intros; red; intros. apply vagree_same. +Qed. + +Lemma eagree_update_1: + forall e1 e2 ne v1 v2 nv r, + eagree e1 e2 ne -> vagree v1 v2 nv -> eagree (e1#r <- v1) (e2#r <- v2) (NE.set r nv ne). +Proof. + intros; red; intros. rewrite NE.gsspec. rewrite ! PMap.gsspec. + destruct (peq r0 r); auto. +Qed. + +Lemma eagree_update: + forall e1 e2 ne v1 v2 r, + vagree v1 v2 (nreg ne r) -> + eagree e1 e2 (NE.set r Nothing ne) -> + eagree (e1#r <- v1) (e2#r <- v2) ne. +Proof. + intros; red; intros. specialize (H0 r0). rewrite NE.gsspec in H0. + rewrite ! PMap.gsspec. destruct (peq r0 r). + subst r0. auto. + auto. +Qed. + +Lemma eagree_update_dead: + forall e1 e2 ne v1 r, + nreg ne r = Nothing -> + eagree e1 e2 ne -> eagree (e1#r <- v1) e2 ne. +Proof. + intros; red; intros. rewrite PMap.gsspec. + destruct (peq r0 r); auto. subst. unfold nreg in H. rewrite H. red; auto. +Qed. + +(** * Neededness for memory locations *) + +Inductive nmem : Type := + | NMemDead + | NMem (stk: ISet.t) (gl: PTree.t ISet.t). + +(** Interpretation of [nmem]: +- [NMemDead]: all memory locations are unused (dead). Acts as bottom. +- [NMem stk gl]: all memory locations are used, except: + - the stack locations whose offset is in the interval [stk] + - the global locations whose offset is in the corresponding entry of [gl]. +*) + +Section LOCATIONS. + +Variable ge: genv. +Variable sp: block. + +Inductive nlive: nmem -> block -> Z -> Prop := + | nlive_intro: forall stk gl b ofs + (STK: b = sp -> ~ISet.In ofs stk) + (GL: forall id iv, + Genv.find_symbol ge id = Some b -> + gl!id = Some iv -> + ~ISet.In ofs iv), + nlive (NMem stk gl) b ofs. + +(** All locations are live *) + +Definition nmem_all := NMem ISet.empty (PTree.empty _). + +Lemma nlive_all: forall b ofs, nlive nmem_all b ofs. +Proof. + intros; constructor; intros. + apply ISet.In_empty. + rewrite PTree.gempty in H0; discriminate. +Qed. + +(** Add a range of live locations to [nm]. The range starts at + the abstract pointer [p] and has length [sz]. *) + +Definition nmem_add (nm: nmem) (p: aptr) (sz: Z) : nmem := + match nm with + | NMemDead => nmem_all (**r very conservative, should never happen *) + | NMem stk gl => + match p with + | Gl id ofs => + match gl!id with + | Some iv => NMem stk (PTree.set id (ISet.remove (Int.unsigned ofs) (Int.unsigned ofs + sz) iv) gl) + | None => nm + end + | Glo id => + NMem stk (PTree.remove id gl) + | Stk ofs => + NMem (ISet.remove (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) gl + | Stack => + NMem ISet.empty gl + | _ => nmem_all + end + end. + +Lemma nlive_add: + forall bc b ofs p nm sz i, + genv_match bc ge -> + bc sp = BCstack -> + pmatch bc b ofs p -> + Int.unsigned ofs <= i < Int.unsigned ofs + sz -> + nlive (nmem_add nm p sz) b i. +Proof. + intros. unfold nmem_add. destruct nm. apply nlive_all. + inv H1; try (apply nlive_all). + - (* Gl id ofs *) + assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). + destruct gl!id as [iv|] eqn:NG. + + constructor; simpl; intros. + congruence. + assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. + rewrite PTree.gss in H5. inv H5. rewrite ISet.In_remove. + intros [A B]. elim A; auto. + + constructor; simpl; intros. + congruence. + assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. + congruence. + - (* Glo id *) + assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). + constructor; simpl; intros. + congruence. + assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. + rewrite PTree.grs in H5. congruence. + - (* Stk ofs *) + constructor; simpl; intros. + rewrite ISet.In_remove. intros [A B]. elim A; auto. + assert (bc b = BCglob id) by (eapply H; eauto). congruence. + - (* Stack *) + constructor; simpl; intros. + apply ISet.In_empty. + assert (bc b = BCglob id) by (eapply H; eauto). congruence. +Qed. + +Lemma incl_nmem_add: + forall nm b i p sz, + nlive nm b i -> nlive (nmem_add nm p sz) b i. +Proof. + intros. inversion H; subst. unfold nmem_add; destruct p; try (apply nlive_all). +- (* Gl id ofs *) + destruct gl!id as [iv|] eqn:NG. + + split; simpl; intros. auto. + rewrite PTree.gsspec in H1. destruct (peq id0 id); eauto. inv H1. + rewrite ISet.In_remove. intros [P Q]. eelim GL; eauto. + + auto. +- (* Glo id *) + split; simpl; intros. auto. + rewrite PTree.grspec in H1. destruct (PTree.elt_eq id0 id). congruence. eauto. +- (* Stk ofs *) + split; simpl; intros. + rewrite ISet.In_remove. intros [P Q]. eelim STK; eauto. + eauto. +- (* Stack *) + split; simpl; intros. + apply ISet.In_empty. + eauto. +Qed. + +(** Remove a range of locations from [nm], marking these locations as dead. + The range starts at the abstract pointer [p] and has length [sz]. *) + +Definition nmem_remove (nm: nmem) (p: aptr) (sz: Z) : nmem := + match nm with + | NMemDead => NMemDead + | NMem stk gl => + match p with + | Gl id ofs => + let iv' := + match gl!id with + | Some iv => ISet.add (Int.unsigned ofs) (Int.unsigned ofs + sz) iv + | None => ISet.interval (Int.unsigned ofs) (Int.unsigned ofs + sz) + end in + NMem stk (PTree.set id iv' gl) + | Stk ofs => + NMem (ISet.add (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) gl + | _ => nm + end + end. + +Lemma nlive_remove: + forall bc b ofs p nm sz b' i, + genv_match bc ge -> + bc sp = BCstack -> + pmatch bc b ofs p -> + nlive nm b' i -> + b' <> b \/ i < Int.unsigned ofs \/ Int.unsigned ofs + sz <= i -> + nlive (nmem_remove nm p sz) b' i. +Proof. + intros. inversion H2; subst. unfold nmem_remove; inv H1; auto. +- (* Gl id ofs *) + set (iv' := match gl!id with + | Some iv => + ISet.add (Int.unsigned ofs) (Int.unsigned ofs + sz) iv + | None => + ISet.interval (Int.unsigned ofs) + (Int.unsigned ofs + sz) + end). + assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). + split; simpl; auto; intros. + rewrite PTree.gsspec in H6. destruct (peq id0 id). ++ inv H6. destruct H3. congruence. destruct gl!id as [iv0|] eqn:NG. + rewrite ISet.In_add. intros [P|P]. omega. eelim GL; eauto. + rewrite ISet.In_interval. omega. ++ eauto. +- (* Stk ofs *) + split; simpl; auto; intros. destruct H3. + elim H3. subst b'. eapply bc_stack; eauto. + rewrite ISet.In_add. intros [P|P]. omega. eapply STK; eauto. +Qed. + +(** Test (conservatively) whether some locations in the range delimited + by [p] and [sz] can be live in [nm]. *) + +Definition nmem_contains (nm: nmem) (p: aptr) (sz: Z) := + match nm with + | NMemDead => false + | NMem stk gl => + match p with + | Gl id ofs => + match gl!id with + | Some iv => negb (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) iv) + | None => true + end + | Stk ofs => + negb (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) + | _ => true (**r conservative answer *) + end + end. + +Lemma nlive_contains: + forall bc b ofs p nm sz i, + genv_match bc ge -> + bc sp = BCstack -> + pmatch bc b ofs p -> + nmem_contains nm p sz = false -> + Int.unsigned ofs <= i < Int.unsigned ofs + sz -> + ~(nlive nm b i). +Proof. + unfold nmem_contains; intros. red; intros L; inv L. + inv H1; try discriminate. +- (* Gl id ofs *) + assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). + destruct gl!id as [iv|] eqn:HG; inv H2. + destruct (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) iv) eqn:IC; try discriminate. + rewrite ISet.contains_spec in IC. eelim GL; eauto. +- (* Stk ofs *) + destruct (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) eqn:IC; try discriminate. + rewrite ISet.contains_spec in IC. eelim STK; eauto. eapply bc_stack; eauto. +Qed. + +(** Kill all stack locations between 0 and [sz], and mark everything else + as live. This reflects the effect of freeing the stack block at + a [Ireturn] or [Itailcall] instruction. *) + +Definition nmem_dead_stack (sz: Z) := + NMem (ISet.interval 0 sz) (PTree.empty _). + +Lemma nlive_dead_stack: + forall sz b' i, b' <> sp \/ ~(0 <= i < sz) -> nlive (nmem_dead_stack sz) b' i. +Proof. + intros; constructor; simpl; intros. +- rewrite ISet.In_interval. intuition. +- rewrite PTree.gempty in H1; discriminate. +Qed. + +(** Least upper bound *) + +Definition nmem_lub (nm1 nm2: nmem) : nmem := + match nm1, nm2 with + | NMemDead, _ => nm2 + | _, NMemDead => nm1 + | NMem stk1 gl1, NMem stk2 gl2 => + NMem (ISet.inter stk1 stk2) + (PTree.combine + (fun o1 o2 => + match o1, o2 with + | Some iv1, Some iv2 => Some(ISet.inter iv1 iv2) + | _, _ => None + end) + gl1 gl2) + end. + +Lemma nlive_lub_l: + forall nm1 nm2 b i, nlive nm1 b i -> nlive (nmem_lub nm1 nm2) b i. +Proof. + intros. inversion H; subst. destruct nm2; simpl. auto. + constructor; simpl; intros. +- rewrite ISet.In_inter. intros [P Q]. eelim STK; eauto. +- rewrite PTree.gcombine in H1 by auto. + destruct gl!id as [iv1|] eqn:NG1; try discriminate; + destruct gl0!id as [iv2|] eqn:NG2; inv H1. + rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto. +Qed. + +Lemma nlive_lub_r: + forall nm1 nm2 b i, nlive nm2 b i -> nlive (nmem_lub nm1 nm2) b i. +Proof. + intros. inversion H; subst. destruct nm1; simpl. auto. + constructor; simpl; intros. +- rewrite ISet.In_inter. intros [P Q]. eelim STK; eauto. +- rewrite PTree.gcombine in H1 by auto. + destruct gl0!id as [iv1|] eqn:NG1; try discriminate; + destruct gl!id as [iv2|] eqn:NG2; inv H1. + rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto. +Qed. + +(** Boolean-valued equality test *) + +Definition nmem_beq (nm1 nm2: nmem) : bool := + match nm1, nm2 with + | NMemDead, NMemDead => true + | NMem stk1 gl1, NMem stk2 gl2 => ISet.beq stk1 stk2 && PTree.beq ISet.beq gl1 gl2 + | _, _ => false + end. + +Lemma nmem_beq_sound: + forall nm1 nm2 b ofs, + nmem_beq nm1 nm2 = true -> + (nlive nm1 b ofs <-> nlive nm2 b ofs). +Proof. + unfold nmem_beq; intros. + destruct nm1 as [ | stk1 gl1]; destruct nm2 as [ | stk2 gl2]; try discriminate. +- split; intros L; inv L. +- InvBooleans. rewrite ISet.beq_spec in H0. rewrite PTree.beq_correct in H1. + split; intros L; inv L; constructor; intros. ++ rewrite <- H0. eauto. ++ specialize (H1 id). rewrite H2 in H1. destruct gl1!id as [iv1|] eqn: NG; try contradiction. + rewrite ISet.beq_spec in H1. rewrite <- H1. eauto. ++ rewrite H0. eauto. ++ specialize (H1 id). rewrite H2 in H1. destruct gl2!id as [iv2|] eqn: NG; try contradiction. + rewrite ISet.beq_spec in H1. rewrite H1. eauto. +Qed. + +End LOCATIONS. + + +(** * The lattice for dataflow analysis *) + +Module NA <: SEMILATTICE. + + Definition t := (nenv * nmem)%type. + + Definition eq (x y: t) := + NE.eq (fst x) (fst y) /\ + (forall ge sp b ofs, nlive ge sp (snd x) b ofs <-> nlive ge sp (snd y) b ofs). + + Lemma eq_refl: forall x, eq x x. + Proof. + unfold eq; destruct x; simpl; split. apply NE.eq_refl. tauto. + Qed. + Lemma eq_sym: forall x y, eq x y -> eq y x. + Proof. + unfold eq; destruct x, y; simpl. intros [A B]. + split. apply NE.eq_sym; auto. + intros. rewrite B. tauto. + Qed. + Lemma eq_trans: forall x y z, eq x y -> eq y z -> eq x z. + Proof. + unfold eq; destruct x, y, z; simpl. intros [A B] [C D]; split. + eapply NE.eq_trans; eauto. + intros. rewrite B; auto. + Qed. + + Definition beq (x y: t) : bool := + NE.beq (fst x) (fst y) && nmem_beq (snd x) (snd y). + + Lemma beq_correct: forall x y, beq x y = true -> eq x y. + Proof. + unfold beq, eq; destruct x, y; simpl; intros. InvBooleans. split. + apply NE.beq_correct; auto. + intros. apply nmem_beq_sound; auto. + Qed. + + Definition ge (x y: t) : Prop := + NE.ge (fst x) (fst y) /\ + (forall ge sp b ofs, nlive ge sp (snd y) b ofs -> nlive ge sp (snd x) b ofs). + + Lemma ge_refl: forall x y, eq x y -> ge x y. + Proof. + unfold eq, ge; destruct x, y; simpl. intros [A B]; split. + apply NE.ge_refl; auto. + intros. apply B; auto. + Qed. + Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z. + Proof. + unfold ge; destruct x, y, z; simpl. intros [A B] [C D]; split. + eapply NE.ge_trans; eauto. + auto. + Qed. + + Definition bot : t := (NE.bot, NMemDead). + + Lemma ge_bot: forall x, ge x bot. + Proof. + unfold ge, bot; destruct x; simpl. split. + apply NE.ge_bot. + intros. inv H. + Qed. + + Definition lub (x y: t) : t := + (NE.lub (fst x) (fst y), nmem_lub (snd x) (snd y)). + + Lemma ge_lub_left: forall x y, ge (lub x y) x. + Proof. + unfold ge; destruct x, y; simpl; split. + apply NE.ge_lub_left. + intros; apply nlive_lub_l; auto. + Qed. + Lemma ge_lub_right: forall x y, ge (lub x y) y. + Proof. + unfold ge; destruct x, y; simpl; split. + apply NE.ge_lub_right. + intros; apply nlive_lub_r; auto. + Qed. + +End NA. + diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml index 429199e..137f65b 100644 --- a/backend/PrintRTL.ml +++ b/backend/PrintRTL.ml @@ -133,3 +133,6 @@ let print_constprop = print_if destination_constprop let destination_cse : string option ref = ref None let print_cse = print_if destination_cse +let destination_deadcode : string option ref = ref None +let print_deadcode = print_if destination_deadcode + diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml index 8a3c05e..5c68602 100644 --- a/backend/Regalloc.ml +++ b/backend/Regalloc.ml @@ -278,7 +278,7 @@ end module Liveness_Solver = Backward_Dataflow_Solver(VSetLat)(NodeSetBackward) let liveness_analysis f = - match Liveness_Solver.fixpoint f.fn_code successors_block (transfer_live f) [] with + match Liveness_Solver.fixpoint f.fn_code successors_block (transfer_live f) with | None -> assert false | Some lv -> lv diff --git a/backend/Splitting.ml b/backend/Splitting.ml index b238cef..3530ba9 100644 --- a/backend/Splitting.ml +++ b/backend/Splitting.ml @@ -128,7 +128,7 @@ let transfer f pc after = (* The live range analysis *) -let analysis f = Solver.fixpoint f.fn_code successors_instr (transfer f) [] +let analysis f = Solver.fixpoint f.fn_code successors_instr (transfer f) (* Produce renamed registers for each instruction. *) diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v new file mode 100644 index 0000000..396d8d4 --- /dev/null +++ b/backend/ValueAnalysis.v @@ -0,0 +1,1812 @@ +Require Import Coqlib. +Require Import Maps. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Globalenvs. +Require Import Events. +Require Import Lattice. +Require Import Kildall. +Require Import Registers. +Require Import Op. +Require Import RTL. +Require Import ValueDomain. +Require Import ValueAOp. +Require Import Liveness. + +(** * The dataflow analysis *) + +Definition areg (ae: aenv) (r: reg) : aval := AE.get r ae. + +Definition aregs (ae: aenv) (rl: list reg) : list aval := List.map (areg ae) rl. + +Definition mafter_public_call : amem := mtop. + +Definition mafter_private_call (am_before: amem) : amem := + {| am_stack := am_before.(am_stack); + am_glob := PTree.empty _; + am_nonstack := Nonstack; + am_top := plub (ab_summary (am_stack am_before)) Nonstack |}. + +Definition transfer_call (ae: aenv) (am: amem) (args: list reg) (res: reg) := + if pincl am.(am_nonstack) Nonstack + && forallb (fun r => vpincl (areg ae r) Nonstack) args + then + VA.State (AE.set res (Ifptr Nonstack) ae) (mafter_private_call am) + else + VA.State (AE.set res Vtop ae) mafter_public_call. + +Inductive builtin_kind : Type := + | Builtin_vload (chunk: memory_chunk) (aaddr: aval) + | Builtin_vstore (chunk: memory_chunk) (aaddr av: aval) + | Builtin_memcpy (sz al: Z) (adst asrc: aval) + | Builtin_annot + | Builtin_annot_val (av: aval) + | Builtin_default. + +Definition classify_builtin (ef: external_function) (args: list reg) (ae: aenv) := + match ef, args with + | EF_vload chunk, a1::nil => Builtin_vload chunk (areg ae a1) + | EF_vload_global chunk id ofs, nil => Builtin_vload chunk (Ptr (Gl id ofs)) + | EF_vstore chunk, a1::a2::nil => Builtin_vstore chunk (areg ae a1) (areg ae a2) + | EF_vstore_global chunk id ofs, a1::nil => Builtin_vstore chunk (Ptr (Gl id ofs)) (areg ae a1) + | EF_memcpy sz al, a1::a2::nil => Builtin_memcpy sz al (areg ae a1) (areg ae a2) + | EF_annot _ _, _ => Builtin_annot + | EF_annot_val _ _, a1::nil => Builtin_annot_val (areg ae a1) + | _, _ => Builtin_default + end. + +Definition transfer_builtin (ae: aenv) (am: amem) (rm: romem) (ef: external_function) (args: list reg) (res: reg) := + match classify_builtin ef args ae with + | Builtin_vload chunk aaddr => + let a := + if strict + then vlub (loadv chunk rm am aaddr) (vnormalize chunk (Ifptr Glob)) + else vnormalize chunk Vtop in + VA.State (AE.set res a ae) am + | Builtin_vstore chunk aaddr av => + let am' := storev chunk am aaddr av in + VA.State (AE.set res itop ae) (mlub am am') + | Builtin_memcpy sz al adst asrc => + let p := loadbytes am rm (aptr_of_aval asrc) in + let am' := storebytes am (aptr_of_aval adst) sz p in + VA.State (AE.set res itop ae) am' + | Builtin_annot => + VA.State (AE.set res itop ae) am + | Builtin_annot_val av => + VA.State (AE.set res av ae) am + | Builtin_default => + transfer_call ae am args res + end. + +Definition transfer (f: function) (rm: romem) (pc: node) (ae: aenv) (am: amem) : VA.t := + match f.(fn_code)!pc with + | None => + VA.Bot + | Some(Inop s) => + VA.State ae am + | Some(Iop op args res s) => + let a := eval_static_operation op (aregs ae args) in + VA.State (AE.set res a ae) am + | Some(Iload chunk addr args dst s) => + let a := loadv chunk rm am (eval_static_addressing addr (aregs ae args)) in + VA.State (AE.set dst a ae) am + | Some(Istore chunk addr args src s) => + let am' := storev chunk am (eval_static_addressing addr (aregs ae args)) (areg ae src) in + VA.State ae am' + | Some(Icall sig ros args res s) => + transfer_call ae am args res + | Some(Itailcall sig ros args) => + VA.Bot + | Some(Ibuiltin ef args res s) => + transfer_builtin ae am rm ef args res + | Some(Icond cond args s1 s2) => + VA.State ae am + | Some(Ijumptable arg tbl) => + VA.State ae am + | Some(Ireturn arg) => + VA.Bot + end. + +Definition transfer' (f: function) (lastuses: PTree.t (list reg)) (rm: romem) + (pc: node) (before: VA.t) : VA.t := + match before with + | VA.Bot => VA.Bot + | VA.State ae am => + match transfer f rm pc ae am with + | VA.Bot => VA.Bot + | VA.State ae' am' => + let ae'' := + match lastuses!pc with + | None => ae' + | Some regs => eforget regs ae' + end in + VA.State ae'' am' + end + end. + +Module DS := Dataflow_Solver(VA)(NodeSetForward). + +Definition mfunction_entry := + {| am_stack := ablock_init Pbot; + am_glob := PTree.empty _; + am_nonstack := Nonstack; + am_top := Nonstack |}. + +Definition analyze (rm: romem) (f: function): PMap.t VA.t := + let lu := Liveness.last_uses f in + let entry := VA.State (einit_regs f.(fn_params)) mfunction_entry in + match DS.fixpoint f.(fn_code) successors_instr (transfer' f lu rm) + f.(fn_entrypoint) entry with + | None => PMap.init (VA.State AE.top mtop) + | Some res => res + end. + +(** Constructing the approximation of read-only globals *) + +Definition store_init_data (ab: ablock) (p: Z) (id: init_data) : ablock := + match id with + | Init_int8 n => ablock_store Mint8unsigned ab p (I n) + | Init_int16 n => ablock_store Mint16unsigned ab p (I n) + | Init_int32 n => ablock_store Mint32 ab p (I n) + | Init_int64 n => ablock_store Mint64 ab p (L n) + | Init_float32 n => ablock_store Mfloat32 ab p + (if propagate_float_constants tt then F n else ftop) + | Init_float64 n => ablock_store Mfloat64 ab p + (if propagate_float_constants tt then F n else ftop) + | Init_addrof symb ofs => ablock_store Mint32 ab p (Ptr (Gl symb ofs)) + | Init_space n => ab + end. + +Fixpoint store_init_data_list (ab: ablock) (p: Z) (idl: list init_data) + {struct idl}: ablock := + match idl with + | nil => ab + | id :: idl' => store_init_data_list (store_init_data ab p id) (p + Genv.init_data_size id) idl' + end. + +Definition alloc_global (rm: romem) (idg: ident * globdef fundef unit): romem := + match idg with + | (id, Gfun f) => + PTree.remove id rm + | (id, Gvar v) => + if v.(gvar_readonly) && negb v.(gvar_volatile) + then PTree.set id (store_init_data_list (ablock_init Pbot) 0 v.(gvar_init)) rm + else PTree.remove id rm + end. + +Definition romem_for_program (p: program) : romem := + List.fold_left alloc_global p.(prog_defs) (PTree.empty _). + +(** * Soundness proof *) + +(** Properties of the dataflow solution. *) + +Lemma analyze_entrypoint: + forall rm f vl m bc, + (forall v, In v vl -> vmatch bc v (Ifptr Nonstack)) -> + mmatch bc m mfunction_entry -> + exists ae am, + (analyze rm f)!!(fn_entrypoint f) = VA.State ae am + /\ ematch bc (init_regs vl (fn_params f)) ae + /\ mmatch bc m am. +Proof. + intros. + unfold analyze. + set (lu := Liveness.last_uses f). + set (entry := VA.State (einit_regs f.(fn_params)) mfunction_entry). + destruct (DS.fixpoint (fn_code f) successors_instr (transfer' f lu rm) + (fn_entrypoint f) entry) as [res|] eqn:FIX. +- assert (A: VA.ge res!!(fn_entrypoint f) entry) by (eapply DS.fixpoint_entry; eauto). + destruct (res!!(fn_entrypoint f)) as [ | ae am ]; simpl in A. contradiction. + destruct A as [A1 A2]. + exists ae, am. + split. auto. + split. eapply ematch_ge; eauto. apply ematch_init; auto. + auto. +- exists AE.top, mtop. + split. apply PMap.gi. + split. apply ematch_ge with (einit_regs (fn_params f)). + apply ematch_init; auto. apply AE.ge_top. + eapply mmatch_top'; eauto. +Qed. + +Lemma analyze_successor: + forall f n ae am instr s rm ae' am', + (analyze rm f)!!n = VA.State ae am -> + f.(fn_code)!n = Some instr -> + In s (successors_instr instr) -> + transfer f rm n ae am = VA.State ae' am' -> + VA.ge (analyze rm f)!!s (transfer f rm n ae am). +Proof. + unfold analyze; intros. + set (lu := Liveness.last_uses f) in *. + set (entry := VA.State (einit_regs f.(fn_params)) mfunction_entry) in *. + destruct (DS.fixpoint (fn_code f) successors_instr (transfer' f lu rm) + (fn_entrypoint f) entry) as [res|] eqn:FIX. +- assert (A: VA.ge res!!s (transfer' f lu rm n res#n)). + { eapply DS.fixpoint_solution; eauto with coqlib. + intros. unfold transfer'. simpl. auto. } + rewrite H in A. unfold transfer' in A. rewrite H2 in A. rewrite H2. + destruct lu!n. + eapply VA.ge_trans. eauto. split; auto. apply eforget_ge. + auto. +- rewrite H2. rewrite PMap.gi. split; intros. apply AE.ge_top. eapply mmatch_top'; eauto. +Qed. + +Lemma analyze_succ: + forall e m rm f n ae am instr s ae' am' bc, + (analyze rm f)!!n = VA.State ae am -> + f.(fn_code)!n = Some instr -> + In s (successors_instr instr) -> + transfer f rm n ae am = VA.State ae' am' -> + ematch bc e ae' -> + mmatch bc m am' -> + exists ae'' am'', + (analyze rm f)!!s = VA.State ae'' am'' + /\ ematch bc e ae'' + /\ mmatch bc m am''. +Proof. + intros. exploit analyze_successor; eauto. rewrite H2. + destruct (analyze rm f)#s as [ | ae'' am'']; simpl; try tauto. intros [A B]. + exists ae'', am''. + split. auto. + split. eapply ematch_ge; eauto. eauto. +Qed. + +(** Classification of builtin functions *) + +Lemma classify_builtin_sound: + forall bc e ae ef (ge: genv) args m t res m', + ematch bc e ae -> + genv_match bc ge -> + external_call ef ge e##args m t res m' -> + match classify_builtin ef args ae with + | Builtin_vload chunk aaddr => + exists addr, + volatile_load_sem chunk ge (addr::nil) m t res m' /\ vmatch bc addr aaddr + | Builtin_vstore chunk aaddr av => + exists addr v, + volatile_store_sem chunk ge (addr::v::nil) m t res m' + /\ vmatch bc addr aaddr /\ vmatch bc v av + | Builtin_memcpy sz al adst asrc => + exists dst, exists src, + extcall_memcpy_sem sz al ge (dst::src::nil) m t res m' + /\ vmatch bc dst adst /\ vmatch bc src asrc + | Builtin_annot => m' = m /\ res = Vundef + | Builtin_annot_val av => m' = m /\ vmatch bc res av + | Builtin_default => True + end. +Proof. + intros. unfold classify_builtin; destruct ef; auto. +- (* vload *) + destruct args; auto. destruct args; auto. + exists (e#p); split; eauto. +- (* vstore *) + destruct args; auto. destruct args; auto. destruct args; auto. + exists (e#p), (e#p0); eauto. +- (* vload global *) + destruct args; auto. simpl in H1. + rewrite volatile_load_global_charact in H1. destruct H1 as (b & A & B). + exists (Vptr b ofs); split; auto. constructor. constructor. eapply H0; eauto. +- (* vstore global *) + destruct args; auto. destruct args; auto. simpl in H1. + rewrite volatile_store_global_charact in H1. destruct H1 as (b & A & B). + exists (Vptr b ofs), (e#p); split; auto. split; eauto. constructor. constructor. eapply H0; eauto. +- (* memcpy *) + destruct args; auto. destruct args; auto. destruct args; auto. + exists (e#p), (e#p0); eauto. +- (* annot *) + simpl in H1. inv H1. auto. +- (* annot val *) + destruct args; auto. destruct args; auto. + simpl in H1. inv H1. eauto. +Qed. + +(** ** Constructing block classifications *) + +Definition bc_nostack (bc: block_classification) : Prop := + forall b, bc b <> BCstack. + +Section NOSTACK. + +Variable bc: block_classification. +Hypothesis NOSTACK: bc_nostack bc. + +Lemma pmatch_no_stack: forall b ofs p, pmatch bc b ofs p -> pmatch bc b ofs Nonstack. +Proof. + intros. inv H; constructor; congruence. +Qed. + +Lemma vmatch_no_stack: forall v x, vmatch bc v x -> vmatch bc v (Ifptr Nonstack). +Proof. + induction 1; constructor; auto; eapply pmatch_no_stack; eauto. +Qed. + +Lemma smatch_no_stack: forall m b p, smatch bc m b p -> smatch bc m b Nonstack. +Proof. + intros. destruct H as [A B]. split; intros. + eapply vmatch_no_stack; eauto. + eapply pmatch_no_stack; eauto. +Qed. + +Lemma mmatch_no_stack: forall m am astk, + mmatch bc m am -> mmatch bc m {| am_stack := astk; am_glob := PTree.empty _; am_nonstack := Nonstack; am_top := Nonstack |}. +Proof. + intros. destruct H. constructor; simpl; intros. +- elim (NOSTACK b); auto. +- rewrite PTree.gempty in H0; discriminate. +- eapply smatch_no_stack; eauto. +- eapply smatch_no_stack; eauto. +- auto. +Qed. + +End NOSTACK. + +(** ** Construction 1: allocating the stack frame at function entry *) + +Ltac splitall := repeat (match goal with |- _ /\ _ => split end). + +Theorem allocate_stack: + forall m sz m' sp bc ge rm am, + Mem.alloc m 0 sz = (m', sp) -> + genv_match bc ge -> + romatch bc m rm -> + mmatch bc m am -> + bc_nostack bc -> + exists bc', + bc_incr bc bc' + /\ bc' sp = BCstack + /\ genv_match bc' ge + /\ romatch bc' m' rm + /\ mmatch bc' m' mfunction_entry + /\ (forall b, Plt b sp -> bc' b = bc b) + /\ (forall v x, vmatch bc v x -> vmatch bc' v (Ifptr Nonstack)). +Proof. + intros until am; intros ALLOC GENV RO MM NOSTACK. + exploit Mem.nextblock_alloc; eauto. intros NB. + exploit Mem.alloc_result; eauto. intros SP. + assert (SPINVALID: bc sp = BCinvalid). + { rewrite SP. eapply bc_below_invalid. apply Plt_strict. eapply mmatch_below; eauto. } +(* Part 1: constructing bc' *) + set (f := fun b => if eq_block b sp then BCstack else bc b). + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + assert (forall b, f b = BCstack -> b = sp). + { unfold f; intros. destruct (eq_block b sp); auto. eelim NOSTACK; eauto. } + intros. transitivity sp; auto. symmetry; auto. + } + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + assert (forall b id, f b = BCglob id -> bc b = BCglob id). + { unfold f; intros. destruct (eq_block b sp). congruence. auto. } + intros. eapply (bc_glob bc); eauto. + } + set (bc' := BC f F_stack F_glob). unfold f in bc'. + assert (BC'EQ: forall b, bc b <> BCinvalid -> bc' b = bc b). + { intros; simpl. apply dec_eq_false. congruence. } + assert (INCR: bc_incr bc bc'). + { red; simpl; intros. apply BC'EQ; auto. } +(* Part 2: invariance properties *) + assert (SM: forall b p, bc b <> BCinvalid -> smatch bc m b p -> smatch bc' m' b Nonstack). + { + intros. + apply smatch_incr with bc; auto. + apply smatch_inv with m. + apply smatch_no_stack with p; auto. + intros. eapply Mem.loadbytes_alloc_unchanged; eauto. eapply mmatch_below; eauto. + } + assert (SMSTACK: smatch bc' m' sp Pbot). + { + split; intros. + exploit Mem.load_alloc_same; eauto. intros EQ. subst v. constructor. + exploit Mem.loadbytes_alloc_same; eauto with coqlib. congruence. + } +(* Conclusions *) + exists bc'; splitall. +- (* incr *) + assumption. +- (* sp is BCstack *) + simpl; apply dec_eq_true. +- (* genv match *) + eapply genv_match_exten; eauto. + simpl; intros. destruct (eq_block b sp); intuition congruence. + simpl; intros. destruct (eq_block b sp); congruence. +- (* romatch *) + apply romatch_exten with bc. + eapply romatch_alloc; eauto. eapply mmatch_below; eauto. + simpl; intros. destruct (eq_block b sp); intuition. +- (* mmatch *) + constructor; simpl; intros. + + (* stack *) + apply ablock_init_sound. destruct (eq_block b sp). + subst b. apply SMSTACK. + elim (NOSTACK b); auto. + + (* globals *) + rewrite PTree.gempty in H0; discriminate. + + (* nonstack *) + destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto. + + (* top *) + destruct (eq_block b sp). + subst b. apply smatch_ge with Pbot. apply SMSTACK. constructor. + eapply SM; auto. eapply mmatch_top; eauto. + + (* below *) + red; simpl; intros. rewrite NB. destruct (eq_block b sp). + subst b; rewrite SP; xomega. + exploit mmatch_below; eauto. xomega. +- (* unchanged *) + simpl; intros. apply dec_eq_false. apply Plt_ne. auto. +- (* values *) + intros. apply vmatch_incr with bc; auto. eapply vmatch_no_stack; eauto. +Qed. + +(** Construction 2: turn the stack into an "other" block, at public calls or function returns *) + +Theorem anonymize_stack: + forall m sp bc ge rm am, + genv_match bc ge -> + romatch bc m rm -> + mmatch bc m am -> + bc sp = BCstack -> + exists bc', + bc_nostack bc' + /\ bc' sp = BCother + /\ (forall b, b <> sp -> bc' b = bc b) + /\ (forall v x, vmatch bc v x -> vmatch bc' v Vtop) + /\ genv_match bc' ge + /\ romatch bc' m rm + /\ mmatch bc' m mtop. +Proof. + intros until am; intros GENV RO MM SP. +(* Part 1: constructing bc' *) + set (f := fun b => if eq_block b sp then BCother else bc b). + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + unfold f; intros. + destruct (eq_block b1 sp); try discriminate. + destruct (eq_block b2 sp); try discriminate. + eapply bc_stack; eauto. + } + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + unfold f; intros. + destruct (eq_block b1 sp); try discriminate. + destruct (eq_block b2 sp); try discriminate. + eapply bc_glob; eauto. + } + set (bc' := BC f F_stack F_glob). unfold f in bc'. + +(* Part 2: matching wrt bc' *) + assert (PM: forall b ofs p, pmatch bc b ofs p -> pmatch bc' b ofs Ptop). + { + intros. assert (pmatch bc b ofs Ptop) by (eapply pmatch_top'; eauto). + inv H0. constructor; simpl. destruct (eq_block b sp); congruence. + } + assert (VM: forall v x, vmatch bc v x -> vmatch bc' v Vtop). + { + induction 1; constructor; eauto. + } + assert (SM: forall b p, smatch bc m b p -> smatch bc' m b Ptop). + { + intros. destruct H as [S1 S2]. split; intros. + eapply VM. eapply S1; eauto. + eapply PM. eapply S2; eauto. + } +(* Conclusions *) + exists bc'; splitall. +- (* nostack *) + red; simpl; intros. destruct (eq_block b sp). congruence. + red; intros. elim n. eapply bc_stack; eauto. +- (* bc' sp is BCother *) + simpl; apply dec_eq_true. +- (* other blocks *) + intros; simpl; apply dec_eq_false; auto. +- (* values *) + auto. +- (* genv *) + apply genv_match_exten with bc; auto. + simpl; intros. destruct (eq_block b sp); intuition congruence. + simpl; intros. destruct (eq_block b sp); auto. +- (* romatch *) + apply romatch_exten with bc; auto. + simpl; intros. destruct (eq_block b sp); intuition. +- (* mmatch top *) + constructor; simpl; intros. + + destruct (eq_block b sp). congruence. elim n. eapply bc_stack; eauto. + + rewrite PTree.gempty in H0; discriminate. + + destruct (eq_block b sp). + subst b. eapply SM. eapply mmatch_stack; eauto. + eapply SM. eapply mmatch_nonstack; eauto. + + destruct (eq_block b sp). + subst b. eapply SM. eapply mmatch_stack; eauto. + eapply SM. eapply mmatch_top; eauto. + + red; simpl; intros. destruct (eq_block b sp). + subst b. eapply mmatch_below; eauto. congruence. + eapply mmatch_below; eauto. +Qed. + +(** Construction 3: turn the stack into an invalid block, at private calls *) + +Theorem hide_stack: + forall m sp bc ge rm am, + genv_match bc ge -> + romatch bc m rm -> + mmatch bc m am -> + bc sp = BCstack -> + pge Nonstack am.(am_nonstack) -> + exists bc', + bc_nostack bc' + /\ bc' sp = BCinvalid + /\ (forall b, b <> sp -> bc' b = bc b) + /\ (forall v x, vge (Ifptr Nonstack) x -> vmatch bc v x -> vmatch bc' v Vtop) + /\ genv_match bc' ge + /\ romatch bc' m rm + /\ mmatch bc' m mtop. +Proof. + intros until am; intros GENV RO MM SP NOLEAK. +(* Part 1: constructing bc' *) + set (f := fun b => if eq_block b sp then BCinvalid else bc b). + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + unfold f; intros. + destruct (eq_block b1 sp); try discriminate. + destruct (eq_block b2 sp); try discriminate. + eapply bc_stack; eauto. + } + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + unfold f; intros. + destruct (eq_block b1 sp); try discriminate. + destruct (eq_block b2 sp); try discriminate. + eapply bc_glob; eauto. + } + set (bc' := BC f F_stack F_glob). unfold f in bc'. + +(* Part 2: matching wrt bc' *) + assert (PM: forall b ofs p, pge Nonstack p -> pmatch bc b ofs p -> pmatch bc' b ofs Ptop). + { + intros. assert (pmatch bc b ofs Nonstack) by (eapply pmatch_ge; eauto). + inv H1. constructor; simpl; destruct (eq_block b sp); congruence. + } + assert (VM: forall v x, vge (Ifptr Nonstack) x -> vmatch bc v x -> vmatch bc' v Vtop). + { + intros. apply vmatch_ifptr; intros. subst v. + inv H0; inv H; eapply PM; eauto. + } + assert (SM: forall b p, pge Nonstack p -> smatch bc m b p -> smatch bc' m b Ptop). + { + intros. destruct H0 as [S1 S2]. split; intros. + eapply VM with (x := Ifptr p). constructor; auto. eapply S1; eauto. + eapply PM. eauto. eapply S2; eauto. + } +(* Conclusions *) + exists bc'; splitall. +- (* nostack *) + red; simpl; intros. destruct (eq_block b sp). congruence. + red; intros. elim n. eapply bc_stack; eauto. +- (* bc' sp is BCinvalid *) + simpl; apply dec_eq_true. +- (* other blocks *) + intros; simpl; apply dec_eq_false; auto. +- (* values *) + auto. +- (* genv *) + apply genv_match_exten with bc; auto. + simpl; intros. destruct (eq_block b sp); intuition congruence. + simpl; intros. destruct (eq_block b sp); congruence. +- (* romatch *) + apply romatch_exten with bc; auto. + simpl; intros. destruct (eq_block b sp); intuition. +- (* mmatch top *) + constructor; simpl; intros. + + destruct (eq_block b sp). congruence. elim n. eapply bc_stack; eauto. + + rewrite PTree.gempty in H0; discriminate. + + destruct (eq_block b sp). congruence. + eapply SM. eauto. eapply mmatch_nonstack; eauto. + + destruct (eq_block b sp). congruence. + eapply SM. eauto. eapply mmatch_nonstack; eauto. + red; intros; elim n. eapply bc_stack; eauto. + + red; simpl; intros. destruct (eq_block b sp). congruence. + eapply mmatch_below; eauto. +Qed. + +(** Construction 4: restore the stack after a public call *) + +Theorem return_from_public_call: + forall (caller callee: block_classification) bound sp ge e ae v m rm, + bc_below caller bound -> + callee sp = BCother -> + caller sp = BCstack -> + (forall b, Plt b bound -> b <> sp -> caller b = callee b) -> + genv_match caller ge -> + ematch caller e ae -> + Ple bound (Mem.nextblock m) -> + vmatch callee v Vtop -> + romatch callee m rm -> + mmatch callee m mtop -> + genv_match callee ge -> + bc_nostack callee -> + exists bc, + vmatch bc v Vtop + /\ ematch bc e ae + /\ romatch bc m rm + /\ mmatch bc m mafter_public_call + /\ genv_match bc ge + /\ bc sp = BCstack + /\ (forall b, Plt b sp -> bc b = caller b). +Proof. + intros until rm; intros BELOW SP1 SP2 SAME GE1 EM BOUND RESM RM MM GE2 NOSTACK. +(* Constructing bc *) + set (f := fun b => if eq_block b sp then BCstack else callee b). + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + assert (forall b, f b = BCstack -> b = sp). + { unfold f; intros. destruct (eq_block b sp); auto. eelim NOSTACK; eauto. } + intros. transitivity sp; auto. symmetry; auto. + } + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + assert (forall b id, f b = BCglob id -> callee b = BCglob id). + { unfold f; intros. destruct (eq_block b sp). congruence. auto. } + intros. eapply (bc_glob callee); eauto. + } + set (bc := BC f F_stack F_glob). unfold f in bc. + assert (INCR: bc_incr caller bc). + { + red; simpl; intros. destruct (eq_block b sp). congruence. + symmetry; apply SAME; auto. + } +(* Invariance properties *) + assert (PM: forall b ofs p, pmatch callee b ofs p -> pmatch bc b ofs Ptop). + { + intros. assert (pmatch callee b ofs Ptop) by (eapply pmatch_top'; eauto). + inv H0. constructor; simpl. destruct (eq_block b sp); congruence. + } + assert (VM: forall v x, vmatch callee v x -> vmatch bc v Vtop). + { + intros. assert (vmatch callee v0 Vtop) by (eapply vmatch_top; eauto). + inv H0; constructor; eauto. + } + assert (SM: forall b p, smatch callee m b p -> smatch bc m b Ptop). + { + intros. destruct H; split; intros. eapply VM; eauto. eapply PM; eauto. + } +(* Conclusions *) + exists bc; splitall. +- (* result value *) + eapply VM; eauto. +- (* environment *) + eapply ematch_incr; eauto. +- (* romem *) + apply romatch_exten with callee; auto. + intros; simpl. destruct (eq_block b sp); intuition. +- (* mmatch *) + constructor; simpl; intros. + + (* stack *) + apply ablock_init_sound. destruct (eq_block b sp). + subst b. eapply SM. eapply mmatch_nonstack; eauto. congruence. + elim (NOSTACK b); auto. + + (* globals *) + rewrite PTree.gempty in H0; discriminate. + + (* nonstack *) + destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto. + + (* top *) + eapply SM. eapply mmatch_top; eauto. + destruct (eq_block b sp); congruence. + + (* below *) + red; simpl; intros. destruct (eq_block b sp). + subst b. eapply mmatch_below; eauto. congruence. + eapply mmatch_below; eauto. +- (* genv *) + eapply genv_match_exten with caller; eauto. + simpl; intros. destruct (eq_block b sp). intuition congruence. + split; intros. rewrite SAME in H by eauto with va. auto. + apply <- (proj1 GE2) in H. apply (proj1 GE1) in H. auto. + simpl; intros. destruct (eq_block b sp). congruence. + rewrite <- SAME; eauto with va. +- (* sp *) + simpl. apply dec_eq_true. +- (* unchanged *) + simpl; intros. destruct (eq_block b sp). congruence. + symmetry. apply SAME; auto. eapply Plt_trans. eauto. apply BELOW. congruence. +Qed. + +(** Construction 5: restore the stack after a private call *) + +Theorem return_from_private_call: + forall (caller callee: block_classification) bound sp ge e ae v m rm am, + bc_below caller bound -> + callee sp = BCinvalid -> + caller sp = BCstack -> + (forall b, Plt b bound -> b <> sp -> caller b = callee b) -> + genv_match caller ge -> + ematch caller e ae -> + bmatch caller m sp am.(am_stack) -> + Ple bound (Mem.nextblock m) -> + vmatch callee v Vtop -> + romatch callee m rm -> + mmatch callee m mtop -> + genv_match callee ge -> + bc_nostack callee -> + exists bc, + vmatch bc v (Ifptr Nonstack) + /\ ematch bc e ae + /\ romatch bc m rm + /\ mmatch bc m (mafter_private_call am) + /\ genv_match bc ge + /\ bc sp = BCstack + /\ (forall b, Plt b sp -> bc b = caller b). +Proof. + intros until am; intros BELOW SP1 SP2 SAME GE1 EM CONTENTS BOUND RESM RM MM GE2 NOSTACK. +(* Constructing bc *) + set (f := fun b => if eq_block b sp then BCstack else callee b). + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + assert (forall b, f b = BCstack -> b = sp). + { unfold f; intros. destruct (eq_block b sp); auto. eelim NOSTACK; eauto. } + intros. transitivity sp; auto. symmetry; auto. + } + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + assert (forall b id, f b = BCglob id -> callee b = BCglob id). + { unfold f; intros. destruct (eq_block b sp). congruence. auto. } + intros. eapply (bc_glob callee); eauto. + } + set (bc := BC f F_stack F_glob). unfold f in bc. + assert (INCR1: bc_incr caller bc). + { + red; simpl; intros. destruct (eq_block b sp). congruence. + symmetry; apply SAME; auto. + } + assert (INCR2: bc_incr callee bc). + { + red; simpl; intros. destruct (eq_block b sp). congruence. auto. + } + +(* Invariance properties *) + assert (PM: forall b ofs p, pmatch callee b ofs p -> pmatch bc b ofs Nonstack). + { + intros. assert (pmatch callee b ofs Ptop) by (eapply pmatch_top'; eauto). + inv H0. constructor; simpl; destruct (eq_block b sp); congruence. + } + assert (VM: forall v x, vmatch callee v x -> vmatch bc v (Ifptr Nonstack)). + { + intros. assert (vmatch callee v0 Vtop) by (eapply vmatch_top; eauto). + inv H0; constructor; eauto. + } + assert (SM: forall b p, smatch callee m b p -> smatch bc m b Nonstack). + { + intros. destruct H; split; intros. eapply VM; eauto. eapply PM; eauto. + } + assert (BSTK: bmatch bc m sp (am_stack am)). + { + apply bmatch_incr with caller; eauto. + } +(* Conclusions *) + exists bc; splitall. +- (* result value *) + eapply VM; eauto. +- (* environment *) + eapply ematch_incr; eauto. +- (* romem *) + apply romatch_exten with callee; auto. + intros; simpl. destruct (eq_block b sp); intuition. +- (* mmatch *) + constructor; simpl; intros. + + (* stack *) + destruct (eq_block b sp). + subst b. exact BSTK. + elim (NOSTACK b); auto. + + (* globals *) + rewrite PTree.gempty in H0; discriminate. + + (* nonstack *) + destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto. + + (* top *) + destruct (eq_block b sp). + subst. apply smatch_ge with (ab_summary (am_stack am)). apply BSTK. apply pge_lub_l. + apply smatch_ge with Nonstack. eapply SM. eapply mmatch_top; eauto. apply pge_lub_r. + + (* below *) + red; simpl; intros. destruct (eq_block b sp). + subst b. apply Plt_le_trans with bound. apply BELOW. congruence. auto. + eapply mmatch_below; eauto. +- (* genv *) + eapply genv_match_exten; eauto. + simpl; intros. destruct (eq_block b sp); intuition congruence. + simpl; intros. destruct (eq_block b sp); congruence. +- (* sp *) + simpl. apply dec_eq_true. +- (* unchanged *) + simpl; intros. destruct (eq_block b sp). congruence. + symmetry. apply SAME; auto. eapply Plt_trans. eauto. apply BELOW. congruence. +Qed. + +(** Construction 6: external call *) + +Theorem external_call_match: + forall ef (ge: genv) vargs m t vres m' bc rm am, + external_call ef ge vargs m t vres m' -> + genv_match bc ge -> + (forall v, In v vargs -> vmatch bc v Vtop) -> + romatch bc m rm -> + mmatch bc m am -> + bc_nostack bc -> + exists bc', + bc_incr bc bc' + /\ (forall b, Plt b (Mem.nextblock m) -> bc' b = bc b) + /\ vmatch bc' vres Vtop + /\ genv_match bc' ge + /\ romatch bc' m' rm + /\ mmatch bc' m' mtop + /\ bc_nostack bc' + /\ (forall b ofs n, Mem.valid_block m b -> bc b = BCinvalid -> Mem.loadbytes m' b ofs n = Mem.loadbytes m b ofs n). +Proof. + intros until am; intros EC GENV ARGS RO MM NOSTACK. + (* Part 1: using ec_mem_inject *) + exploit (@external_call_mem_inject ef _ _ ge vargs m t vres m' (inj_of_bc bc) m vargs). + apply inj_of_bc_preserves_globals; auto. + exact EC. + eapply mmatch_inj; eauto. eapply mmatch_below; eauto. + revert ARGS. generalize vargs. + induction vargs0; simpl; intros; constructor. + eapply vmatch_inj; eauto. auto. + intros (j' & vres' & m'' & EC' & IRES & IMEM & UNCH1 & UNCH2 & IINCR & ISEP). + assert (JBELOW: forall b, Plt b (Mem.nextblock m) -> j' b = inj_of_bc bc b). + { + intros. destruct (inj_of_bc bc b) as [[b' delta] | ] eqn:EQ. + eapply IINCR; eauto. + destruct (j' b) as [[b'' delta'] | ] eqn:EQ'; auto. + exploit ISEP; eauto. tauto. + } + (* Part 2: constructing bc' from j' *) + set (f := fun b => if plt b (Mem.nextblock m) + then bc b + else match j' b with None => BCinvalid | Some _ => BCother end). + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + assert (forall b, f b = BCstack -> bc b = BCstack). + { unfold f; intros. destruct (plt b (Mem.nextblock m)); auto. destruct (j' b); discriminate. } + intros. apply (bc_stack bc); auto. + } + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + assert (forall b id, f b = BCglob id -> bc b = BCglob id). + { unfold f; intros. destruct (plt b (Mem.nextblock m)); auto. destruct (j' b); discriminate. } + intros. eapply (bc_glob bc); eauto. + } + set (bc' := BC f F_stack F_glob). unfold f in bc'. + assert (INCR: bc_incr bc bc'). + { + red; simpl; intros. apply pred_dec_true. eapply mmatch_below; eauto. + } + assert (BC'INV: forall b, bc' b <> BCinvalid -> exists b' delta, j' b = Some(b', delta)). + { + simpl; intros. destruct (plt b (Mem.nextblock m)). + exists b, 0. rewrite JBELOW by auto. apply inj_of_bc_valid; auto. + destruct (j' b) as [[b' delta] | ]. + exists b', delta; auto. + congruence. + } + + (* Part 3: injection wrt j' implies matching with top wrt bc' *) + assert (PMTOP: forall b b' delta ofs, j' b = Some (b', delta) -> pmatch bc' b ofs Ptop). + { + intros. constructor. simpl; unfold f. + destruct (plt b (Mem.nextblock m)). + rewrite JBELOW in H by auto. eapply inj_of_bc_inv; eauto. + rewrite H; congruence. + } + assert (VMTOP: forall v v', val_inject j' v v' -> vmatch bc' v Vtop). + { + intros. inv H; constructor. eapply PMTOP; eauto. + } + assert (SMTOP: forall b, bc' b <> BCinvalid -> smatch bc' m' b Ptop). + { + intros; split; intros. + - exploit BC'INV; eauto. intros (b' & delta & J'). + exploit Mem.load_inject. eexact IMEM. eauto. eauto. intros (v' & A & B). + eapply VMTOP; eauto. + - exploit BC'INV; eauto. intros (b'' & delta & J'). + exploit Mem.loadbytes_inject. eexact IMEM. eauto. eauto. intros (bytes & A & B). + inv B. inv H3. eapply PMTOP; eauto. + } + (* Conclusions *) + exists bc'; splitall. +- (* incr *) + exact INCR. +- (* unchanged *) + simpl; intros. apply pred_dec_true; auto. +- (* vmatch res *) + eapply VMTOP; eauto. +- (* genv match *) + apply genv_match_exten with bc; auto. + simpl; intros; split; intros. + rewrite pred_dec_true by (eapply mmatch_below; eauto with va). auto. + destruct (plt b (Mem.nextblock m)). auto. destruct (j' b); congruence. + simpl; intros. rewrite pred_dec_true by (eapply mmatch_below; eauto with va). auto. +- (* romatch m' *) + red; simpl; intros. destruct (plt b (Mem.nextblock m)). + exploit RO; eauto. intros (R & P & Q). + split; auto. + split. apply bmatch_incr with bc; auto. apply bmatch_inv with m; auto. + intros. eapply Mem.loadbytes_unchanged_on_1. eapply external_call_readonly; eauto. + auto. intros; red. apply Q. + intros; red; intros; elim (Q ofs). + eapply external_call_max_perm with (m2 := m'); eauto. + destruct (j' b); congruence. +- (* mmatch top *) + constructor; simpl; intros. + + apply ablock_init_sound. apply SMTOP. simpl; congruence. + + rewrite PTree.gempty in H0; discriminate. + + apply SMTOP; auto. + + apply SMTOP; auto. + + red; simpl; intros. destruct (plt b (Mem.nextblock m)). + eapply Plt_le_trans. eauto. eapply external_call_nextblock; eauto. + destruct (j' b) as [[bx deltax] | ] eqn:J'. + eapply Mem.valid_block_inject_1; eauto. + congruence. +- (* nostack *) + red; simpl; intros. destruct (plt b (Mem.nextblock m)). + apply NOSTACK; auto. + destruct (j' b); congruence. +- (* unmapped blocks are invariant *) + intros. eapply Mem.loadbytes_unchanged_on_1; auto. + apply UNCH1; auto. intros; red. unfold inj_of_bc; rewrite H0; auto. +Qed. + +(** ** Semantic invariant *) + +Section SOUNDNESS. + +Variable prog: program. + +Let ge : genv := Genv.globalenv prog. + +Let rm := romem_for_program prog. + +Inductive sound_stack: block_classification -> list stackframe -> mem -> block -> Prop := + | sound_stack_nil: forall bc m bound, + sound_stack bc nil m bound + | sound_stack_public_call: + forall (bc: block_classification) res f sp pc e stk m bound bc' bound' ae + (STK: sound_stack bc' stk m sp) + (INCR: Ple bound' bound) + (BELOW: bc_below bc' bound') + (SP: bc sp = BCother) + (SP': bc' sp = BCstack) + (SAME: forall b, Plt b bound' -> b <> sp -> bc b = bc' b) + (GE: genv_match bc' ge) + (AN: VA.ge (analyze rm f)!!pc (VA.State (AE.set res Vtop ae) mafter_public_call)) + (EM: ematch bc' e ae), + sound_stack bc (Stackframe res f (Vptr sp Int.zero) pc e :: stk) m bound + | sound_stack_private_call: + forall (bc: block_classification) res f sp pc e stk m bound bc' bound' ae am + (STK: sound_stack bc' stk m sp) + (INCR: Ple bound' bound) + (BELOW: bc_below bc' bound') + (SP: bc sp = BCinvalid) + (SP': bc' sp = BCstack) + (SAME: forall b, Plt b bound' -> b <> sp -> bc b = bc' b) + (GE: genv_match bc' ge) + (AN: VA.ge (analyze rm f)!!pc (VA.State (AE.set res (Ifptr Nonstack) ae) (mafter_private_call am))) + (EM: ematch bc' e ae) + (CONTENTS: bmatch bc' m sp am.(am_stack)), + sound_stack bc (Stackframe res f (Vptr sp Int.zero) pc e :: stk) m bound. + +Inductive sound_state: state -> Prop := + | sound_regular_state: + forall s f sp pc e m ae am bc + (STK: sound_stack bc s m sp) + (AN: (analyze rm f)!!pc = VA.State ae am) + (EM: ematch bc e ae) + (RO: romatch bc m rm) + (MM: mmatch bc m am) + (GE: genv_match bc ge) + (SP: bc sp = BCstack), + sound_state (State s f (Vptr sp Int.zero) pc e m) + | sound_call_state: + forall s fd args m bc + (STK: sound_stack bc s m (Mem.nextblock m)) + (ARGS: forall v, In v args -> vmatch bc v Vtop) + (RO: romatch bc m rm) + (MM: mmatch bc m mtop) + (GE: genv_match bc ge) + (NOSTK: bc_nostack bc), + sound_state (Callstate s fd args m) + | sound_return_state: + forall s v m bc + (STK: sound_stack bc s m (Mem.nextblock m)) + (RES: vmatch bc v Vtop) + (RO: romatch bc m rm) + (MM: mmatch bc m mtop) + (GE: genv_match bc ge) + (NOSTK: bc_nostack bc), + sound_state (Returnstate s v m). + +(** Properties of the [sound_stack] invariant on call stacks. *) + +Lemma sound_stack_ext: + forall m' bc stk m bound, + sound_stack bc stk m bound -> + (forall b ofs n bytes, + Plt b bound -> bc b = BCinvalid -> n >= 0 -> + Mem.loadbytes m' b ofs n = Some bytes -> + Mem.loadbytes m b ofs n = Some bytes) -> + sound_stack bc stk m' bound. +Proof. + induction 1; intros INV. +- constructor. +- assert (Plt sp bound') by eauto with va. + eapply sound_stack_public_call; eauto. apply IHsound_stack; intros. + apply INV. xomega. rewrite SAME; auto. xomega. auto. auto. +- assert (Plt sp bound') by eauto with va. + eapply sound_stack_private_call; eauto. apply IHsound_stack; intros. + apply INV. xomega. rewrite SAME; auto. xomega. auto. auto. + apply bmatch_ext with m; auto. intros. apply INV. xomega. auto. auto. auto. +Qed. + +Lemma sound_stack_inv: + forall m' bc stk m bound, + sound_stack bc stk m bound -> + (forall b ofs n, Plt b bound -> bc b = BCinvalid -> n >= 0 -> Mem.loadbytes m' b ofs n = Mem.loadbytes m b ofs n) -> + sound_stack bc stk m' bound. +Proof. + intros. eapply sound_stack_ext; eauto. intros. rewrite <- H0; auto. +Qed. + +Lemma sound_stack_storev: + forall chunk m addr v m' bc aaddr stk bound, + Mem.storev chunk m addr v = Some m' -> + vmatch bc addr aaddr -> + sound_stack bc stk m bound -> + sound_stack bc stk m' bound. +Proof. + intros. apply sound_stack_inv with m; auto. + destruct addr; simpl in H; try discriminate. + assert (A: pmatch bc b i Ptop). + { inv H0; eapply pmatch_top'; eauto. } + inv A. + intros. eapply Mem.loadbytes_store_other; eauto. left; congruence. +Qed. + +Lemma sound_stack_storebytes: + forall m b ofs bytes m' bc aaddr stk bound, + Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' -> + vmatch bc (Vptr b ofs) aaddr -> + sound_stack bc stk m bound -> + sound_stack bc stk m' bound. +Proof. + intros. apply sound_stack_inv with m; auto. + assert (A: pmatch bc b ofs Ptop). + { inv H0; eapply pmatch_top'; eauto. } + inv A. + intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence. +Qed. + +Lemma sound_stack_free: + forall m b lo hi m' bc stk bound, + Mem.free m b lo hi = Some m' -> + sound_stack bc stk m bound -> + sound_stack bc stk m' bound. +Proof. + intros. eapply sound_stack_ext; eauto. intros. + eapply Mem.loadbytes_free_2; eauto. +Qed. + +Lemma sound_stack_new_bound: + forall bc stk m bound bound', + sound_stack bc stk m bound -> + Ple bound bound' -> + sound_stack bc stk m bound'. +Proof. + intros. inv H. +- constructor. +- eapply sound_stack_public_call with (bound' := bound'0); eauto. xomega. +- eapply sound_stack_private_call with (bound' := bound'0); eauto. xomega. +Qed. + +Lemma sound_stack_exten: + forall bc stk m bound (bc1: block_classification), + sound_stack bc stk m bound -> + (forall b, Plt b bound -> bc1 b = bc b) -> + sound_stack bc1 stk m bound. +Proof. + intros. inv H. +- constructor. +- assert (Plt sp bound') by eauto with va. + eapply sound_stack_public_call; eauto. + rewrite H0; auto. xomega. + intros. rewrite H0; auto. xomega. +- assert (Plt sp bound') by eauto with va. + eapply sound_stack_private_call; eauto. + rewrite H0; auto. xomega. + intros. rewrite H0; auto. xomega. +Qed. + +(** ** Preservation of the semantic invariant by one step of execution *) + +Lemma sound_succ_state: + forall bc pc ae am instr ae' am' s f sp pc' e' m', + (analyze rm f)!!pc = VA.State ae am -> + f.(fn_code)!pc = Some instr -> + In pc' (successors_instr instr) -> + transfer f rm pc ae am = VA.State ae' am' -> + ematch bc e' ae' -> + mmatch bc m' am' -> + romatch bc m' rm -> + genv_match bc ge -> + bc sp = BCstack -> + sound_stack bc s m' sp -> + sound_state (State s f (Vptr sp Int.zero) pc' e' m'). +Proof. + intros. exploit analyze_succ; eauto. intros (ae'' & am'' & AN & EM & MM). + econstructor; eauto. +Qed. + +Lemma areg_sound: + forall bc e ae r, ematch bc e ae -> vmatch bc (e#r) (areg ae r). +Proof. + intros. apply H. +Qed. + +Lemma aregs_sound: + forall bc e ae rl, ematch bc e ae -> list_forall2 (vmatch bc) (e##rl) (aregs ae rl). +Proof. + induction rl; simpl; intros. constructor. constructor; auto. apply areg_sound; auto. +Qed. + +Hint Resolve areg_sound aregs_sound: va. + +Theorem sound_step: + forall st t st', RTL.step ge st t st' -> sound_state st -> sound_state st'. +Proof. + induction 1; intros SOUND; inv SOUND. + +- (* nop *) + eapply sound_succ_state; eauto. simpl; auto. + unfold transfer; rewrite H. auto. + +- (* op *) + eapply sound_succ_state; eauto. simpl; auto. + unfold transfer; rewrite H. eauto. + apply ematch_update; auto. eapply eval_static_operation_sound; eauto with va. + +- (* load *) + eapply sound_succ_state; eauto. simpl; auto. + unfold transfer; rewrite H. eauto. + apply ematch_update; auto. eapply loadv_sound; eauto with va. + eapply eval_static_addressing_sound; eauto with va. + +- (* store *) + exploit eval_static_addressing_sound; eauto with va. intros VMADDR. + eapply sound_succ_state; eauto. simpl; auto. + unfold transfer; rewrite H. eauto. + eapply storev_sound; eauto. + destruct a; simpl in H1; try discriminate. eapply romatch_store; eauto. + eapply sound_stack_storev; eauto. + +- (* call *) + assert (TR: transfer f rm pc ae am = transfer_call ae am args res). + { unfold transfer; rewrite H; auto. } + unfold transfer_call in TR. + destruct (pincl (am_nonstack am) Nonstack && + forallb (fun r : reg => vpincl (areg ae r) Nonstack) args) eqn:NOLEAK. ++ (* private call *) + InvBooleans. + exploit analyze_successor; eauto. simpl; eauto. rewrite TR. intros SUCC. + exploit hide_stack; eauto. apply pincl_ge; auto. + intros (bc' & A & B & C & D & E & F & G). + apply sound_call_state with bc'; auto. + * eapply sound_stack_private_call with (bound' := Mem.nextblock m) (bc' := bc); eauto. + apply Ple_refl. + eapply mmatch_below; eauto. + eapply mmatch_stack; eauto. + * intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v. + apply D with (areg ae r). + rewrite forallb_forall in H2. apply vpincl_ge. apply H2; auto. auto with va. ++ (* public call *) + exploit analyze_successor; eauto. simpl; eauto. rewrite TR. intros SUCC. + exploit anonymize_stack; eauto. intros (bc' & A & B & C & D & E & F & G). + apply sound_call_state with bc'; auto. + * eapply sound_stack_public_call with (bound' := Mem.nextblock m) (bc' := bc); eauto. + apply Ple_refl. + eapply mmatch_below; eauto. + * intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v. + apply D with (areg ae r). auto with va. + +- (* tailcall *) + exploit anonymize_stack; eauto. intros (bc' & A & B & C & D & E & F & G). + apply sound_call_state with bc'; auto. + erewrite Mem.nextblock_free by eauto. + apply sound_stack_new_bound with stk. + apply sound_stack_exten with bc. + eapply sound_stack_free; eauto. + intros. apply C. apply Plt_ne; auto. + apply Plt_Ple. eapply mmatch_below; eauto. congruence. + intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v. + apply D with (areg ae r). auto with va. + eapply romatch_free; eauto. + eapply mmatch_free; eauto. + +- (* builtin *) + assert (SPVALID: Plt sp0 (Mem.nextblock m)) by (eapply mmatch_below; eauto with va). + assert (TR: transfer f rm pc ae am = transfer_builtin ae am rm ef args res). + { unfold transfer; rewrite H; auto. } + unfold transfer_builtin in TR. + exploit classify_builtin_sound; eauto. destruct (classify_builtin ef args ae). ++ (* volatile load *) + intros (addr & VLOAD & VADDR). inv VLOAD. + eapply sound_succ_state; eauto. simpl; auto. + apply ematch_update; auto. + inv H2. + * (* true volatile access *) + assert (V: vmatch bc v0 (Ifptr Glob)). + { inv H4; constructor. econstructor. eapply GE; eauto. } + destruct strict. apply vmatch_lub_r. apply vnormalize_sound. auto. + apply vnormalize_sound. eapply vmatch_ge; eauto. constructor. constructor. + * (* normal memory access *) + exploit loadv_sound; eauto. simpl; eauto. intros V. + destruct strict. + apply vmatch_lub_l. auto. + eapply vnormalize_cast; eauto. eapply vmatch_top; eauto. ++ (* volatile store *) + intros (addr & src & VSTORE & VADDR & VSRC). inv VSTORE. inv H7. + * (* true volatile access *) + eapply sound_succ_state; eauto. simpl; auto. + apply ematch_update; auto. constructor. + apply mmatch_lub_l; auto. + * (* normal memory access *) + eapply sound_succ_state; eauto. simpl; auto. + apply ematch_update; auto. constructor. + apply mmatch_lub_r. eapply storev_sound; eauto. auto. + eapply romatch_store; eauto. + eapply sound_stack_storev; eauto. simpl; eauto. ++ (* memcpy *) + intros (dst & src & MEMCPY & VDST & VSRC). inv MEMCPY. + eapply sound_succ_state; eauto. simpl; auto. + apply ematch_update; auto. constructor. + eapply storebytes_sound; eauto. + apply match_aptr_of_aval; auto. + eapply Mem.loadbytes_length; eauto. + intros. eapply loadbytes_sound; eauto. apply match_aptr_of_aval; auto. + eapply romatch_storebytes; eauto. + exploit Mem.loadbytes_length; eauto. + intros. exploit (nat_of_Z_eq sz). omega. rewrite <- H1; intros. + destruct bytes. simpl in H2. omegaContradiction. congruence. + eapply sound_stack_storebytes; eauto. ++ (* annot *) + intros (A & B); subst. + eapply sound_succ_state; eauto. simpl; auto. + apply ematch_update; auto. constructor. ++ (* annot val *) + intros (A & B); subst. + eapply sound_succ_state; eauto. simpl; auto. + apply ematch_update; auto. ++ (* general case *) + intros _. + unfold transfer_call in TR. + destruct (pincl (am_nonstack am) Nonstack && + forallb (fun r : reg => vpincl (areg ae r) Nonstack) args) eqn:NOLEAK. +* (* private builtin call *) + InvBooleans. rewrite forallb_forall in H2. + exploit hide_stack; eauto. apply pincl_ge; auto. + intros (bc1 & A & B & C & D & E & F & G). + exploit external_call_match; eauto. + intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v0. + eapply D; eauto with va. apply vpincl_ge. apply H2; auto. + intros (bc2 & J & K & L & M & N & O & P & Q). + exploit (return_from_private_call bc bc2); eauto. + eapply mmatch_below; eauto. + rewrite K; auto. + intros. rewrite K; auto. rewrite C; auto. + apply bmatch_inv with m. eapply mmatch_stack; eauto. + intros. apply Q; auto. + eapply external_call_nextblock; eauto. + intros (bc3 & U & V & W & X & Y & Z & AA). + eapply sound_succ_state with (bc := bc3); eauto. simpl; auto. + apply ematch_update; auto. + apply sound_stack_exten with bc. + apply sound_stack_inv with m. auto. + intros. apply Q. red. eapply Plt_trans; eauto. + rewrite C; auto. + exact AA. +* (* public builtin call *) + exploit anonymize_stack; eauto. + intros (bc1 & A & B & C & D & E & F & G). + exploit external_call_match; eauto. + intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v0. eapply D; eauto with va. + intros (bc2 & J & K & L & M & N & O & P & Q). + exploit (return_from_public_call bc bc2); eauto. + eapply mmatch_below; eauto. + rewrite K; auto. + intros. rewrite K; auto. rewrite C; auto. + eapply external_call_nextblock; eauto. + intros (bc3 & U & V & W & X & Y & Z & AA). + eapply sound_succ_state with (bc := bc3); eauto. simpl; auto. + apply ematch_update; auto. + apply sound_stack_exten with bc. + apply sound_stack_inv with m. auto. + intros. apply Q. red. eapply Plt_trans; eauto. + rewrite C; auto. + exact AA. + +- (* cond *) + eapply sound_succ_state; eauto. + simpl. destruct b; auto. + unfold transfer; rewrite H; auto. + +- (* jumptable *) + eapply sound_succ_state; eauto. + simpl. eapply list_nth_z_in; eauto. + unfold transfer; rewrite H; auto. + +- (* return *) + exploit anonymize_stack; eauto. intros (bc' & A & B & C & D & E & F & G). + apply sound_return_state with bc'; auto. + erewrite Mem.nextblock_free by eauto. + apply sound_stack_new_bound with stk. + apply sound_stack_exten with bc. + eapply sound_stack_free; eauto. + intros. apply C. apply Plt_ne; auto. + apply Plt_Ple. eapply mmatch_below; eauto with va. + destruct or; simpl. eapply D; eauto. constructor. + eapply romatch_free; eauto. + eapply mmatch_free; eauto. + +- (* internal function *) + exploit allocate_stack; eauto. + intros (bc' & A & B & C & D & E & F & G). + exploit (analyze_entrypoint rm f args m' bc'); eauto. + intros (ae & am & AN & EM & MM'). + econstructor; eauto. + erewrite Mem.alloc_result by eauto. + apply sound_stack_exten with bc; auto. + apply sound_stack_inv with m; auto. + intros. eapply Mem.loadbytes_alloc_unchanged; eauto. + intros. apply F. erewrite Mem.alloc_result by eauto. auto. + +- (* external function *) + exploit external_call_match; eauto with va. + intros (bc' & A & B & C & D & E & F & G & K). + econstructor; eauto. + apply sound_stack_new_bound with (Mem.nextblock m). + apply sound_stack_exten with bc; auto. + apply sound_stack_inv with m; auto. + eapply external_call_nextblock; eauto. + +- (* return *) + inv STK. + + (* from public call *) + exploit return_from_public_call; eauto. + intros; rewrite SAME; auto. + intros (bc1 & A & B & C & D & E & F & G). + destruct (analyze rm f)#pc as [ |ae' am'] eqn:EQ; simpl in AN; try contradiction. destruct AN as [A1 A2]. + eapply sound_regular_state with (bc := bc1); eauto. + apply sound_stack_exten with bc'; auto. + eapply ematch_ge; eauto. apply ematch_update. auto. auto. + + (* from private call *) + exploit return_from_private_call; eauto. + intros; rewrite SAME; auto. + intros (bc1 & A & B & C & D & E & F & G). + destruct (analyze rm f)#pc as [ |ae' am'] eqn:EQ; simpl in AN; try contradiction. destruct AN as [A1 A2]. + eapply sound_regular_state with (bc := bc1); eauto. + apply sound_stack_exten with bc'; auto. + eapply ematch_ge; eauto. apply ematch_update. auto. auto. +Qed. + +End SOUNDNESS. + +(** ** Soundness of the initial memory abstraction *) + +Section INITIAL. + +Variable prog: program. + +Let ge := Genv.globalenv prog. + +Lemma initial_block_classification: + forall m, + Genv.init_mem prog = Some m -> + exists bc, + genv_match bc ge + /\ bc_below bc (Mem.nextblock m) + /\ bc_nostack bc + /\ (forall b id, bc b = BCglob id -> Genv.find_symbol ge id = Some b) + /\ (forall b, Mem.valid_block m b -> bc b <> BCinvalid). +Proof. + intros. + set (f := fun b => + if plt b (Genv.genv_next ge) then + match Genv.invert_symbol ge b with None => BCother | Some id => BCglob id end + else + BCinvalid). + assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2). + { + unfold f; intros. + destruct (plt b1 (Genv.genv_next ge)); try discriminate. + destruct (Genv.invert_symbol ge b1) as [id1|] eqn:I1; inv H0. + destruct (plt b2 (Genv.genv_next ge)); try discriminate. + destruct (Genv.invert_symbol ge b2) as [id2|] eqn:I2; inv H1. + exploit Genv.invert_find_symbol. eexact I1. + exploit Genv.invert_find_symbol. eexact I2. + congruence. + } + assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2). + { + unfold f; intros. + destruct (plt b1 (Genv.genv_next ge)); try discriminate. + destruct (Genv.invert_symbol ge b1); discriminate. + } + set (bc := BC f F_stack F_glob). unfold f in bc. + exists bc; splitall. +- split; simpl; intros. + + split; intros. + * rewrite pred_dec_true by (eapply Genv.genv_symb_range; eauto). + erewrite Genv.find_invert_symbol; eauto. + * apply Genv.invert_find_symbol. + destruct (plt b (Genv.genv_next ge)); try discriminate. + destruct (Genv.invert_symbol ge b); congruence. + + rewrite ! pred_dec_true by assumption. + destruct (Genv.invert_symbol); split; congruence. +- red; simpl; intros. destruct (plt b (Genv.genv_next ge)); try congruence. + erewrite <- Genv.init_mem_genv_next by eauto. auto. +- red; simpl; intros. + destruct (plt b (Genv.genv_next ge)). + destruct (Genv.invert_symbol ge b); congruence. + congruence. +- simpl; intros. destruct (plt b (Genv.genv_next ge)); try discriminate. + destruct (Genv.invert_symbol ge b) as [id' | ] eqn:IS; inv H0. + apply Genv.invert_find_symbol; auto. +- intros; simpl. unfold ge; erewrite Genv.init_mem_genv_next by eauto. + rewrite pred_dec_true by assumption. + destruct (Genv.invert_symbol (Genv.globalenv prog) b); congruence. +Qed. + +Section INIT. + +Variable bc: block_classification. +Hypothesis GMATCH: genv_match bc ge. + +Lemma store_init_data_summary: + forall ab p id, + pge Glob (ab_summary ab) -> + pge Glob (ab_summary (store_init_data ab p id)). +Proof. + intros. + assert (DFL: forall chunk av, + vge (Ifptr Glob) av -> + pge Glob (ab_summary (ablock_store chunk ab p av))). + { + intros. simpl. unfold vplub; destruct av; auto. + inv H0. apply plub_least; auto. + inv H0. apply plub_least; auto. + } + destruct id; auto. + simpl. destruct (propagate_float_constants tt); auto. + simpl. destruct (propagate_float_constants tt); auto. + apply DFL. constructor. constructor. +Qed. + +Lemma store_init_data_list_summary: + forall idl ab p, + pge Glob (ab_summary ab) -> + pge Glob (ab_summary (store_init_data_list ab p idl)). +Proof. + induction idl; simpl; intros. auto. apply IHidl. apply store_init_data_summary; auto. +Qed. + +Lemma store_init_data_sound: + forall m b p id m' ab, + Genv.store_init_data ge m b p id = Some m' -> + bmatch bc m b ab -> + bmatch bc m' b (store_init_data ab p id). +Proof. + intros. destruct id; try (eapply ablock_store_sound; eauto; constructor). + simpl. destruct (propagate_float_constants tt); eapply ablock_store_sound; eauto; constructor. + simpl. destruct (propagate_float_constants tt); eapply ablock_store_sound; eauto; constructor. + simpl in H. inv H. auto. + simpl in H. destruct (Genv.find_symbol ge i) as [b'|] eqn:FS; try discriminate. + eapply ablock_store_sound; eauto. constructor. constructor. apply GMATCH; auto. +Qed. + +Lemma store_init_data_list_sound: + forall idl m b p m' ab, + Genv.store_init_data_list ge m b p idl = Some m' -> + bmatch bc m b ab -> + bmatch bc m' b (store_init_data_list ab p idl). +Proof. + induction idl; simpl; intros. +- inv H; auto. +- destruct (Genv.store_init_data ge m b p a) as [m1|] eqn:SI; try discriminate. + eapply IHidl; eauto. eapply store_init_data_sound; eauto. +Qed. + +Lemma store_init_data_other: + forall m b p id m' ab b', + Genv.store_init_data ge m b p id = Some m' -> + b' <> b -> + bmatch bc m b' ab -> + bmatch bc m' b' ab. +Proof. + intros. eapply bmatch_inv; eauto. + intros. destruct id; try (eapply Mem.loadbytes_store_other; eauto; fail); simpl in H. + inv H; auto. + destruct (Genv.find_symbol ge i); try discriminate. + eapply Mem.loadbytes_store_other; eauto. +Qed. + +Lemma store_init_data_list_other: + forall b b' ab idl m p m', + Genv.store_init_data_list ge m b p idl = Some m' -> + b' <> b -> + bmatch bc m b' ab -> + bmatch bc m' b' ab. +Proof. + induction idl; simpl; intros. + inv H; auto. + destruct (Genv.store_init_data ge m b p a) as [m1|] eqn:SI; try discriminate. + eapply IHidl; eauto. eapply store_init_data_other; eauto. +Qed. + +Lemma store_zeros_same: + forall p m b pos n m', + store_zeros m b pos n = Some m' -> + smatch bc m b p -> + smatch bc m' b p. +Proof. + intros until n. functional induction (store_zeros m b pos n); intros. +- inv H. auto. +- eapply IHo; eauto. change p with (vplub (I Int.zero) p). + eapply smatch_store; eauto. constructor. +- discriminate. +Qed. + +Lemma store_zeros_other: + forall b' ab m b p n m', + store_zeros m b p n = Some m' -> + b' <> b -> + bmatch bc m b' ab -> + bmatch bc m' b' ab. +Proof. + intros until n. functional induction (store_zeros m b p n); intros. +- inv H. auto. +- eapply IHo; eauto. eapply bmatch_inv; eauto. + intros. eapply Mem.loadbytes_store_other; eauto. +- discriminate. +Qed. + +Definition initial_mem_match (bc: block_classification) (m: mem) (g: genv) := + forall b v, + Genv.find_var_info g b = Some v -> + v.(gvar_volatile) = false -> v.(gvar_readonly) = true -> + bmatch bc m b (store_init_data_list (ablock_init Pbot) 0 v.(gvar_init)). + +Lemma alloc_global_match: + forall m g idg m', + Genv.genv_next g = Mem.nextblock m -> + initial_mem_match bc m g -> + Genv.alloc_global ge m idg = Some m' -> + initial_mem_match bc m' (Genv.add_global g idg). +Proof. + intros; red; intros. destruct idg as [id [fd | gv]]; simpl in *. +- destruct (Mem.alloc m 0 1) as [m1 b1] eqn:ALLOC. + unfold Genv.find_var_info, Genv.add_global in H2; simpl in H2. + assert (Plt b (Mem.nextblock m)). + { rewrite <- H. eapply Genv.genv_vars_range; eauto. } + assert (b <> b1). + { apply Plt_ne. erewrite Mem.alloc_result by eauto. auto. } + apply bmatch_inv with m. + eapply H0; eauto. + intros. transitivity (Mem.loadbytes m1 b ofs n). + eapply Mem.loadbytes_drop; eauto. + eapply Mem.loadbytes_alloc_unchanged; eauto. +- set (sz := Genv.init_data_list_size (gvar_init gv)) in *. + destruct (Mem.alloc m 0 sz) as [m1 b1] eqn:ALLOC. + destruct (store_zeros m1 b1 0 sz) as [m2 | ] eqn:STZ; try discriminate. + destruct (Genv.store_init_data_list ge m2 b1 0 (gvar_init gv)) as [m3 | ] eqn:SIDL; try discriminate. + unfold Genv.find_var_info, Genv.add_global in H2; simpl in H2. + rewrite PTree.gsspec in H2. destruct (peq b (Genv.genv_next g)). ++ inversion H2; clear H2; subst v. + assert (b = b1). { erewrite Mem.alloc_result by eauto. congruence. } + clear e. subst b. + apply bmatch_inv with m3. + eapply store_init_data_list_sound; eauto. + apply ablock_init_sound. + eapply store_zeros_same; eauto. + split; intros. + exploit Mem.load_alloc_same; eauto. intros EQ; subst v; constructor. + exploit Mem.loadbytes_alloc_same; eauto with coqlib. congruence. + intros. eapply Mem.loadbytes_drop; eauto. + right; right; right. unfold Genv.perm_globvar. rewrite H3, H4. constructor. ++ assert (Plt b (Mem.nextblock m)). + { rewrite <- H. eapply Genv.genv_vars_range; eauto. } + assert (b <> b1). + { apply Plt_ne. erewrite Mem.alloc_result by eauto. auto. } + apply bmatch_inv with m3. + eapply store_init_data_list_other; eauto. + eapply store_zeros_other; eauto. + apply bmatch_inv with m. + eapply H0; eauto. + intros. eapply Mem.loadbytes_alloc_unchanged; eauto. + intros. eapply Mem.loadbytes_drop; eauto. +Qed. + +Lemma alloc_globals_match: + forall gl m g m', + Genv.genv_next g = Mem.nextblock m -> + initial_mem_match bc m g -> + Genv.alloc_globals ge m gl = Some m' -> + initial_mem_match bc m' (Genv.add_globals g gl). +Proof. + induction gl; simpl; intros. +- inv H1; auto. +- destruct (Genv.alloc_global ge m a) as [m1|] eqn:AG; try discriminate. + eapply IHgl; eauto. + erewrite Genv.alloc_global_nextblock; eauto. simpl. congruence. + eapply alloc_global_match; eauto. +Qed. + +Definition romem_consistent (g: genv) (rm: romem) := + forall id b ab, + Genv.find_symbol g id = Some b -> rm!id = Some ab -> + exists v, + Genv.find_var_info g b = Some v + /\ v.(gvar_readonly) = true + /\ v.(gvar_volatile) = false + /\ ab = store_init_data_list (ablock_init Pbot) 0 v.(gvar_init). + +Lemma alloc_global_consistent: + forall g rm idg, + romem_consistent g rm -> + romem_consistent (Genv.add_global g idg) (alloc_global rm idg). +Proof. + intros; red; intros. destruct idg as [id1 [fd1 | v1]]; + unfold Genv.add_global, Genv.find_symbol, Genv.find_var_info, alloc_global in *; simpl in *. +- rewrite PTree.gsspec in H0. rewrite PTree.grspec in H1. unfold PTree.elt_eq in *. + destruct (peq id id1). congruence. eapply H; eauto. +- rewrite PTree.gsspec in H0. destruct (peq id id1). ++ inv H0. rewrite PTree.gss. + destruct (gvar_readonly v1 && negb (gvar_volatile v1)) eqn:RO. + InvBooleans. rewrite negb_true_iff in H2. + rewrite PTree.gss in H1. + exists v1. intuition congruence. + rewrite PTree.grs in H1. discriminate. ++ rewrite PTree.gso. eapply H; eauto. + destruct (gvar_readonly v1 && negb (gvar_volatile v1)). + rewrite PTree.gso in H1; auto. + rewrite PTree.gro in H1; auto. + apply Plt_ne. eapply Genv.genv_symb_range; eauto. +Qed. + +Lemma alloc_globals_consistent: + forall gl g rm, + romem_consistent g rm -> + romem_consistent (Genv.add_globals g gl) (List.fold_left alloc_global gl rm). +Proof. + induction gl; simpl; intros. auto. apply IHgl. apply alloc_global_consistent; auto. +Qed. + +End INIT. + +Theorem initial_mem_matches: + forall m, + Genv.init_mem prog = Some m -> + exists bc, + genv_match bc ge + /\ bc_below bc (Mem.nextblock m) + /\ bc_nostack bc + /\ romatch bc m (romem_for_program prog) + /\ (forall b, Mem.valid_block m b -> bc b <> BCinvalid). +Proof. + intros. + exploit initial_block_classification; eauto. intros (bc & GE & BELOW & NOSTACK & INV & VALID). + exists bc; splitall; auto. + assert (A: initial_mem_match bc m ge). + { + apply alloc_globals_match with (m := Mem.empty); auto. + red. unfold Genv.find_var_info; simpl. intros. rewrite PTree.gempty in H0; discriminate. + } + assert (B: romem_consistent ge (romem_for_program prog)). + { + apply alloc_globals_consistent. + red; intros. rewrite PTree.gempty in H1; discriminate. + } + red; intros. + exploit B; eauto. intros (v & FV & RO & NVOL & EQ). + split. subst ab. apply store_init_data_list_summary. constructor. + split. subst ab. eapply A; eauto. + unfold ge in FV; exploit Genv.init_mem_characterization; eauto. + intros (P & Q & R). + intros; red; intros. exploit Q; eauto. intros [U V]. + unfold Genv.perm_globvar in V; rewrite RO, NVOL in V. inv V. +Qed. + +End INITIAL. + +Require Import Axioms. + +Theorem sound_initial: + forall prog st, initial_state prog st -> sound_state prog st. +Proof. + destruct 1. + exploit initial_mem_matches; eauto. intros (bc & GE & BELOW & NOSTACK & RM & VALID). + apply sound_call_state with bc. +- constructor. +- simpl; tauto. +- exact RM. +- apply mmatch_inj_top with m0. + replace (inj_of_bc bc) with (Mem.flat_inj (Mem.nextblock m0)). + eapply Genv.initmem_inject; eauto. + symmetry; apply extensionality; unfold Mem.flat_inj; intros x. + destruct (plt x (Mem.nextblock m0)). + apply inj_of_bc_valid; auto. + unfold inj_of_bc. erewrite bc_below_invalid; eauto. +- exact GE. +- exact NOSTACK. +Qed. + +Hint Resolve areg_sound aregs_sound: va. + +(** * Interface with other optimizations *) + +Definition avalue (a: VA.t) (r: reg) : aval := + match a with + | VA.Bot => Vbot + | VA.State ae am => AE.get r ae + end. + +Lemma avalue_sound: + forall prog s f sp pc e m r, + sound_state prog (State s f (Vptr sp Int.zero) pc e m) -> + exists bc, + vmatch bc e#r (avalue (analyze (romem_for_program prog) f)!!pc r) + /\ genv_match bc (Genv.globalenv prog) + /\ bc sp = BCstack. +Proof. + intros. inv H. exists bc; split; auto. rewrite AN. apply EM. +Qed. + +Definition aaddr (a: VA.t) (r: reg) : aptr := + match a with + | VA.Bot => Pbot + | VA.State ae am => aptr_of_aval (AE.get r ae) + end. + +Lemma aaddr_sound: + forall prog s f sp pc e m r b ofs, + sound_state prog (State s f (Vptr sp Int.zero) pc e m) -> + e#r = Vptr b ofs -> + exists bc, + pmatch bc b ofs (aaddr (analyze (romem_for_program prog) f)!!pc r) + /\ genv_match bc (Genv.globalenv prog) + /\ bc sp = BCstack. +Proof. + intros. inv H. exists bc; split; auto. + unfold aaddr; rewrite AN. apply match_aptr_of_aval. rewrite <- H0. apply EM. +Qed. + +Definition aaddressing (a: VA.t) (addr: addressing) (args: list reg) : aptr := + match a with + | VA.Bot => Pbot + | VA.State ae am => aptr_of_aval (eval_static_addressing addr (aregs ae args)) + end. + +Lemma aaddressing_sound: + forall prog s f sp pc e m addr args b ofs, + sound_state prog (State s f (Vptr sp Int.zero) pc e m) -> + eval_addressing (Genv.globalenv prog) (Vptr sp Int.zero) addr e##args = Some (Vptr b ofs) -> + exists bc, + pmatch bc b ofs (aaddressing (analyze (romem_for_program prog) f)!!pc addr args) + /\ genv_match bc (Genv.globalenv prog) + /\ bc sp = BCstack. +Proof. + intros. inv H. exists bc; split; auto. + unfold aaddressing. rewrite AN. apply match_aptr_of_aval. + eapply eval_static_addressing_sound; eauto with va. +Qed. + + + + + diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v new file mode 100644 index 0000000..2c728d3 --- /dev/null +++ b/backend/ValueDomain.v @@ -0,0 +1,3692 @@ +Require Import Coqlib. +Require Import Zwf. +Require Import Maps. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Globalenvs. +Require Import Events. +Require Import Lattice. +Require Import Kildall. +Require Import Registers. +Require Import RTL. + +Parameter strict: bool. +Parameter propagate_float_constants: unit -> bool. +Parameter generate_float_constants : unit -> bool. + +Inductive block_class : Type := + | BCinvalid + | BCglob (id: ident) + | BCstack + | BCother. + +Definition block_class_eq: forall (x y: block_class), {x=y} + {x<>y}. +Proof. decide equality. apply peq. Defined. + +Record block_classification : Type := BC { + bc_img :> block -> block_class; + bc_stack: forall b1 b2, bc_img b1 = BCstack -> bc_img b2 = BCstack -> b1 = b2; + bc_glob: forall b1 b2 id, bc_img b1 = BCglob id -> bc_img b2 = BCglob id -> b1 = b2 +}. + +Definition bc_below (bc: block_classification) (bound: block) : Prop := + forall b, bc b <> BCinvalid -> Plt b bound. + +Lemma bc_below_invalid: + forall b bc bound, ~Plt b bound -> bc_below bc bound -> bc b = BCinvalid. +Proof. + intros. destruct (block_class_eq (bc b) BCinvalid); auto. + elim H. apply H0; auto. +Qed. + +Hint Extern 2 (_ = _) => congruence : va. +Hint Extern 2 (_ <> _) => congruence : va. +Hint Extern 2 (_ < _) => xomega : va. +Hint Extern 2 (_ <= _) => xomega : va. +Hint Extern 2 (_ > _) => xomega : va. +Hint Extern 2 (_ >= _) => xomega : va. + +Section MATCH. + +Variable bc: block_classification. + +(** * Abstracting the result of conditions (type [option bool]) *) + +Inductive abool := Bnone | Just (b: bool) | Maybe (b: bool) | Btop. + +Inductive cmatch: option bool -> abool -> Prop := + | cmatch_none: cmatch None Bnone + | cmatch_just: forall b, cmatch (Some b) (Just b) + | cmatch_maybe_none: forall b, cmatch None (Maybe b) + | cmatch_maybe_some: forall b, cmatch (Some b) (Maybe b) + | cmatch_top: forall ob, cmatch ob Btop. + +Hint Constructors cmatch : va. + +Definition club (x y: abool) : abool := + match x, y with + | Bnone, Bnone => Bnone + | Bnone, (Just b | Maybe b) => Maybe b + | (Just b | Maybe b), Bnone => Maybe b + | Just b1, Just b2 => if eqb b1 b2 then x else Btop + | Maybe b1, Maybe b2 => if eqb b1 b2 then x else Btop + | Maybe b1, Just b2 => if eqb b1 b2 then x else Btop + | Just b1, Maybe b2 => if eqb b1 b2 then y else Btop + | _, _ => Btop + end. + +Lemma cmatch_lub_l: + forall ob x y, cmatch ob x -> cmatch ob (club x y). +Proof. + intros. unfold club; inv H; destruct y; try constructor; + destruct (eqb b b0) eqn:EQ; try constructor. + replace b0 with b by (apply eqb_prop; auto). constructor. +Qed. + +Lemma cmatch_lub_r: + forall ob x y, cmatch ob y -> cmatch ob (club x y). +Proof. + intros. unfold club; inv H; destruct x; try constructor; + destruct (eqb b0 b) eqn:EQ; try constructor. + replace b with b0 by (apply eqb_prop; auto). constructor. + replace b with b0 by (apply eqb_prop; auto). constructor. + replace b with b0 by (apply eqb_prop; auto). constructor. +Qed. + +Definition cnot (x: abool) : abool := + match x with + | Just b => Just (negb b) + | Maybe b => Maybe (negb b) + | _ => x + end. + +Lemma cnot_sound: + forall ob x, cmatch ob x -> cmatch (option_map negb ob) (cnot x). +Proof. + destruct 1; constructor. +Qed. + +(** * Abstracting pointers *) + +Inductive aptr : Type := + | Pbot + | Gl (id: ident) (ofs: int) + | Glo (id: ident) + | Glob + | Stk (ofs: int) + | Stack + | Nonstack + | Ptop. + +Definition eq_aptr: forall (p1 p2: aptr), {p1=p2} + {p1<>p2}. +Proof. + intros. generalize ident_eq, Int.eq_dec; intros. decide equality. +Defined. + +Inductive pmatch (b: block) (ofs: int): aptr -> Prop := + | pmatch_gl: forall id, + bc b = BCglob id -> + pmatch b ofs (Gl id ofs) + | pmatch_glo: forall id, + bc b = BCglob id -> + pmatch b ofs (Glo id) + | pmatch_glob: forall id, + bc b = BCglob id -> + pmatch b ofs Glob + | pmatch_stk: + bc b = BCstack -> + pmatch b ofs (Stk ofs) + | pmatch_stack: + bc b = BCstack -> + pmatch b ofs Stack + | pmatch_nonstack: + bc b <> BCstack -> bc b <> BCinvalid -> + pmatch b ofs Nonstack + | pmatch_top: + bc b <> BCinvalid -> + pmatch b ofs Ptop. + +Hint Constructors pmatch: va. + +Inductive pge: aptr -> aptr -> Prop := + | pge_top: forall p, pge Ptop p + | pge_bot: forall p, pge p Pbot + | pge_refl: forall p, pge p p + | pge_glo_gl: forall id ofs, pge (Glo id) (Gl id ofs) + | pge_glob_gl: forall id ofs, pge Glob (Gl id ofs) + | pge_glob_glo: forall id, pge Glob (Glo id) + | pge_ns_gl: forall id ofs, pge Nonstack (Gl id ofs) + | pge_ns_glo: forall id, pge Nonstack (Glo id) + | pge_ns_glob: pge Nonstack Glob + | pge_stack_stk: forall ofs, pge Stack (Stk ofs). + +Hint Constructors pge: va. + +Lemma pge_trans: + forall p q, pge p q -> forall r, pge q r -> pge p r. +Proof. + induction 1; intros r PM; inv PM; auto with va. +Qed. + +Lemma pmatch_ge: + forall b ofs p q, pge p q -> pmatch b ofs q -> pmatch b ofs p. +Proof. + induction 1; intros PM; inv PM; eauto with va. +Qed. + +Lemma pmatch_top': forall b ofs p, pmatch b ofs p -> pmatch b ofs Ptop. +Proof. + intros. apply pmatch_ge with p; auto with va. +Qed. + +Definition plub (p q: aptr) : aptr := + match p, q with + | Pbot, _ => q + | _, Pbot => p + | Gl id1 ofs1, Gl id2 ofs2 => + if ident_eq id1 id2 then if Int.eq_dec ofs1 ofs2 then p else Glo id1 else Glob + | Gl id1 ofs1, Glo id2 => + if ident_eq id1 id2 then q else Glob + | Glo id1, Gl id2 ofs2 => + if ident_eq id1 id2 then p else Glob + | Glo id1, Glo id2 => + if ident_eq id1 id2 then p else Glob + | (Gl _ _ | Glo _ | Glob), Glob => Glob + | Glob, (Gl _ _ | Glo _) => Glob + | (Gl _ _ | Glo _ | Glob | Nonstack), Nonstack => + Nonstack + | Nonstack, (Gl _ _ | Glo _ | Glob) => + Nonstack + | Stk ofs1, Stk ofs2 => + if Int.eq_dec ofs1 ofs2 then p else Stack + | (Stk _ | Stack), Stack => + Stack + | Stack, Stk _ => + Stack + | _, _ => Ptop + end. + +Lemma plub_comm: + forall p q, plub p q = plub q p. +Proof. + intros; unfold plub; destruct p; destruct q; auto. + destruct (ident_eq id id0). subst id0. + rewrite dec_eq_true. + destruct (Int.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true. auto. + rewrite dec_eq_false by auto. auto. + rewrite dec_eq_false by auto. auto. + destruct (ident_eq id id0). subst id0. + rewrite dec_eq_true; auto. + rewrite dec_eq_false; auto. + destruct (ident_eq id id0). subst id0. + rewrite dec_eq_true; auto. + rewrite dec_eq_false; auto. + destruct (ident_eq id id0). subst id0. + rewrite dec_eq_true; auto. + rewrite dec_eq_false; auto. + destruct (Int.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true; auto. + rewrite dec_eq_false; auto. +Qed. + +Lemma pge_lub_l: + forall p q, pge (plub p q) p. +Proof. + unfold plub; destruct p, q; auto with va. +- destruct (ident_eq id id0). + destruct (Int.eq_dec ofs ofs0); subst; constructor. + constructor. +- destruct (ident_eq id id0); subst; constructor. +- destruct (ident_eq id id0); subst; constructor. +- destruct (ident_eq id id0); subst; constructor. +- destruct (Int.eq_dec ofs ofs0); subst; constructor. +Qed. + +Lemma pge_lub_r: + forall p q, pge (plub p q) q. +Proof. + intros. rewrite plub_comm. apply pge_lub_l. +Qed. + +Lemma pmatch_lub_l: + forall b ofs p q, pmatch b ofs p -> pmatch b ofs (plub p q). +Proof. + intros. eapply pmatch_ge; eauto. apply pge_lub_l. +Qed. + +Lemma pmatch_lub_r: + forall b ofs p q, pmatch b ofs q -> pmatch b ofs (plub p q). +Proof. + intros. eapply pmatch_ge; eauto. apply pge_lub_r. +Qed. + +Lemma plub_least: + forall r p q, pge r p -> pge r q -> pge r (plub p q). +Proof. + intros. inv H; inv H0; simpl; try constructor. +- destruct p; constructor. +- unfold plub; destruct q; repeat rewrite dec_eq_true; constructor. +- rewrite dec_eq_true; constructor. +- rewrite dec_eq_true; constructor. +- rewrite dec_eq_true. destruct (Int.eq_dec ofs ofs0); constructor. +- destruct (ident_eq id id0). destruct (Int.eq_dec ofs ofs0); constructor. constructor. +- destruct (ident_eq id id0); constructor. +- destruct (ident_eq id id0); constructor. +- destruct (ident_eq id id0); constructor. +- destruct (ident_eq id id0). destruct (Int.eq_dec ofs ofs0); constructor. constructor. +- destruct (ident_eq id id0); constructor. +- destruct (ident_eq id id0); constructor. +- destruct (ident_eq id id0); constructor. +- destruct (Int.eq_dec ofs ofs0); constructor. +Qed. + +Definition pincl (p q: aptr) : bool := + match p, q with + | Pbot, _ => true + | Gl id1 ofs1, Gl id2 ofs2 => peq id1 id2 && Int.eq_dec ofs1 ofs2 + | Gl id1 ofs1, Glo id2 => peq id1 id2 + | Glo id1, Glo id2 => peq id1 id2 + | (Gl _ _ | Glo _ | Glob), Glob => true + | (Gl _ _ | Glo _ | Glob | Nonstack), Nonstack => true + | Stk ofs1, Stk ofs2 => Int.eq_dec ofs1 ofs2 + | Stk ofs1, Stack => true + | _, Ptop => true + | _, _ => false + end. + +Lemma pincl_ge: forall p q, pincl p q = true -> pge q p. +Proof. + unfold pincl; destruct p, q; intros; try discriminate; auto with va; + InvBooleans; subst; auto with va. +Qed. + +Lemma pincl_sound: + forall b ofs p q, + pincl p q = true -> pmatch b ofs p -> pmatch b ofs q. +Proof. + intros. eapply pmatch_ge; eauto. apply pincl_ge; auto. +Qed. + +Definition padd (p: aptr) (n: int) : aptr := + match p with + | Gl id ofs => Gl id (Int.add ofs n) + | Stk ofs => Stk (Int.add ofs n) + | _ => p + end. + +Lemma padd_sound: + forall b ofs p delta, + pmatch b ofs p -> + pmatch b (Int.add ofs delta) (padd p delta). +Proof. + intros. inv H; simpl padd; eauto with va. +Qed. + +Definition psub (p: aptr) (n: int) : aptr := + match p with + | Gl id ofs => Gl id (Int.sub ofs n) + | Stk ofs => Stk (Int.sub ofs n) + | _ => p + end. + +Lemma psub_sound: + forall b ofs p delta, + pmatch b ofs p -> + pmatch b (Int.sub ofs delta) (psub p delta). +Proof. + intros. inv H; simpl psub; eauto with va. +Qed. + +Definition poffset (p: aptr) : aptr := + match p with + | Gl id ofs => Glo id + | Stk ofs => Stack + | _ => p + end. + +Lemma poffset_sound: + forall b ofs1 ofs2 p, + pmatch b ofs1 p -> + pmatch b ofs2 (poffset p). +Proof. + intros. inv H; simpl poffset; eauto with va. +Qed. + +Definition psub2 (p q: aptr) : option int := + match p, q with + | Gl id1 ofs1, Gl id2 ofs2 => + if peq id1 id2 then Some (Int.sub ofs1 ofs2) else None + | Stk ofs1, Stk ofs2 => + Some (Int.sub ofs1 ofs2) + | _, _ => + None + end. + +Lemma psub2_sound: + forall b1 ofs1 p1 b2 ofs2 p2 delta, + psub2 p1 p2 = Some delta -> + pmatch b1 ofs1 p1 -> + pmatch b2 ofs2 p2 -> + b1 = b2 /\ delta = Int.sub ofs1 ofs2. +Proof. + intros. destruct p1; try discriminate; destruct p2; try discriminate; simpl in H. +- destruct (peq id id0); inv H. inv H0; inv H1. + split. eapply bc_glob; eauto. reflexivity. +- inv H; inv H0; inv H1. split. eapply bc_stack; eauto. reflexivity. +Qed. + +Definition cmp_different_blocks (c: comparison) : abool := + match c with + | Ceq => Maybe false + | Cne => Maybe true + | _ => Bnone + end. + +Lemma cmp_different_blocks_none: + forall c, cmatch None (cmp_different_blocks c). +Proof. + intros; destruct c; constructor. +Qed. + +Lemma cmp_different_blocks_sound: + forall c, cmatch (Val.cmp_different_blocks c) (cmp_different_blocks c). +Proof. + intros; destruct c; constructor. +Qed. + +Definition pcmp (c: comparison) (p1 p2: aptr) : abool := + match p1, p2 with + | Pbot, _ | _, Pbot => Bnone + | Gl id1 ofs1, Gl id2 ofs2 => + if peq id1 id2 then Maybe (Int.cmpu c ofs1 ofs2) + else cmp_different_blocks c + | Gl id1 ofs1, Glo id2 => + if peq id1 id2 then Btop else cmp_different_blocks c + | Glo id1, Gl id2 ofs2 => + if peq id1 id2 then Btop else cmp_different_blocks c + | Glo id1, Glo id2 => + if peq id1 id2 then Btop else cmp_different_blocks c + | Stk ofs1, Stk ofs2 => Maybe (Int.cmpu c ofs1 ofs2) + | (Gl _ _ | Glo _ | Glob | Nonstack), (Stk _ | Stack) => cmp_different_blocks c + | (Stk _ | Stack), (Gl _ _ | Glo _ | Glob | Nonstack) => cmp_different_blocks c + | _, _ => Btop + end. + +Lemma pcmp_sound: + forall valid c b1 ofs1 p1 b2 ofs2 p2, + pmatch b1 ofs1 p1 -> pmatch b2 ofs2 p2 -> + cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2)) (pcmp c p1 p2). +Proof. + intros. + assert (DIFF: b1 <> b2 -> + cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2)) + (cmp_different_blocks c)). + { + intros. simpl. rewrite dec_eq_false by assumption. + destruct (valid b1 (Int.unsigned ofs1) && valid b2 (Int.unsigned ofs2)); simpl. + apply cmp_different_blocks_sound. + apply cmp_different_blocks_none. + } + assert (SAME: b1 = b2 -> + cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2)) + (Maybe (Int.cmpu c ofs1 ofs2))). + { + intros. subst b2. simpl. rewrite dec_eq_true. + destruct ((valid b1 (Int.unsigned ofs1) || valid b1 (Int.unsigned ofs1 - 1)) && + (valid b1 (Int.unsigned ofs2) || valid b1 (Int.unsigned ofs2 - 1))); simpl. + constructor. + constructor. + } + unfold pcmp; inv H; inv H0; (apply cmatch_top || (apply DIFF; congruence) || idtac). + - destruct (peq id id0). subst id0. apply SAME. eapply bc_glob; eauto. + auto with va. + - destruct (peq id id0); auto with va. + - destruct (peq id id0); auto with va. + - destruct (peq id id0); auto with va. + - apply SAME. eapply bc_stack; eauto. +Qed. + +Lemma pcmp_none: + forall c p1 p2, cmatch None (pcmp c p1 p2). +Proof. + intros. + unfold pcmp; destruct p1; try constructor; destruct p2; + try (destruct (peq id id0)); try constructor; try (apply cmp_different_blocks_none). +Qed. + +Definition pdisjoint (p1: aptr) (sz1: Z) (p2: aptr) (sz2: Z) : bool := + match p1, p2 with + | Pbot, _ => true + | _, Pbot => true + | Gl id1 ofs1, Gl id2 ofs2 => + if peq id1 id2 + then zle (Int.unsigned ofs1 + sz1) (Int.unsigned ofs2) + || zle (Int.unsigned ofs2 + sz2) (Int.unsigned ofs1) + else true + | Gl id1 ofs1, Glo id2 => negb(peq id1 id2) + | Glo id1, Gl id2 ofs2 => negb(peq id1 id2) + | Glo id1, Glo id2 => negb(peq id1 id2) + | Stk ofs1, Stk ofs2 => + zle (Int.unsigned ofs1 + sz1) (Int.unsigned ofs2) + || zle (Int.unsigned ofs2 + sz2) (Int.unsigned ofs1) + | (Gl _ _ | Glo _ | Glob | Nonstack), (Stk _ | Stack) => true + | (Stk _ | Stack), (Gl _ _ | Glo _ | Glob | Nonstack) => true + | _, _ => false + end. + +Lemma pdisjoint_sound: + forall sz1 b1 ofs1 p1 sz2 b2 ofs2 p2, + pdisjoint p1 sz1 p2 sz2 = true -> + pmatch b1 ofs1 p1 -> pmatch b2 ofs2 p2 -> + b1 <> b2 \/ Int.unsigned ofs1 + sz1 <= Int.unsigned ofs2 \/ Int.unsigned ofs2 + sz2 <= Int.unsigned ofs1. +Proof. + intros. inv H0; inv H1; simpl in H; try discriminate; try (left; congruence). +- destruct (peq id id0). subst id0. destruct (orb_true_elim _ _ H); InvBooleans; auto. + left; congruence. +- destruct (peq id id0); try discriminate. left; congruence. +- destruct (peq id id0); try discriminate. left; congruence. +- destruct (peq id id0); try discriminate. left; congruence. +- destruct (orb_true_elim _ _ H); InvBooleans; auto. +Qed. + +(** * Abstracting values *) + +Inductive aval : Type := + | Vbot + | I (n: int) + | Uns (n: Z) + | Sgn (n: Z) + | L (n: int64) + | F (f: float) + | Fsingle + | Ptr (p: aptr) + | Ifptr (p: aptr). + +Definition eq_aval: forall (v1 v2: aval), {v1=v2} + {v1<>v2}. +Proof. + intros. generalize zeq Int.eq_dec Int64.eq_dec Float.eq_dec eq_aptr; intros. + decide equality. +Defined. + +Definition Vtop := Ifptr Ptop. +Definition itop := Ifptr Pbot. +Definition ftop := Ifptr Pbot. +Definition ltop := Ifptr Pbot. + +Definition is_uns (n: Z) (i: int) : Prop := + forall m, 0 <= m < Int.zwordsize -> m >= n -> Int.testbit i m = false. +Definition is_sgn (n: Z) (i: int) : Prop := + forall m, 0 <= m < Int.zwordsize -> m >= n - 1 -> Int.testbit i m = Int.testbit i (Int.zwordsize - 1). + +Inductive vmatch : val -> aval -> Prop := + | vmatch_i: forall i, vmatch (Vint i) (I i) + | vmatch_Uns: forall i n, 0 <= n -> is_uns n i -> vmatch (Vint i) (Uns n) + | vmatch_Uns_undef: forall n, vmatch Vundef (Uns n) + | vmatch_Sgn: forall i n, 0 < n -> is_sgn n i -> vmatch (Vint i) (Sgn n) + | vmatch_Sgn_undef: forall n, vmatch Vundef (Sgn n) + | vmatch_l: forall i, vmatch (Vlong i) (L i) + | vmatch_f: forall f, vmatch (Vfloat f) (F f) + | vmatch_single: forall f, Float.is_single f -> vmatch (Vfloat f) Fsingle + | vmatch_single_undef: vmatch Vundef Fsingle + | vmatch_ptr: forall b ofs p, pmatch b ofs p -> vmatch (Vptr b ofs) (Ptr p) + | vmatch_ptr_undef: forall p, vmatch Vundef (Ptr p) + | vmatch_ifptr_undef: forall p, vmatch Vundef (Ifptr p) + | vmatch_ifptr_i: forall i p, vmatch (Vint i) (Ifptr p) + | vmatch_ifptr_l: forall i p, vmatch (Vlong i) (Ifptr p) + | vmatch_ifptr_f: forall f p, vmatch (Vfloat f) (Ifptr p) + | vmatch_ifptr_p: forall b ofs p, pmatch b ofs p -> vmatch (Vptr b ofs) (Ifptr p). + +Lemma vmatch_ifptr: + forall v p, + (forall b ofs, v = Vptr b ofs -> pmatch b ofs p) -> + vmatch v (Ifptr p). +Proof. + intros. destruct v; constructor; auto. +Qed. + +Lemma vmatch_top: forall v x, vmatch v x -> vmatch v Vtop. +Proof. + intros. apply vmatch_ifptr. intros. subst v. inv H; eapply pmatch_top'; eauto. +Qed. + +Lemma vmatch_itop: forall i, vmatch (Vint i) itop. +Proof. intros; constructor. Qed. + +Lemma vmatch_undef_itop: vmatch Vundef itop. +Proof. constructor. Qed. + +Lemma vmatch_ftop: forall f, vmatch (Vfloat f) ftop. +Proof. intros; constructor. Qed. + +Lemma vmatch_undef_ftop: vmatch Vundef ftop. +Proof. constructor. Qed. + +Hint Constructors vmatch : va. +Hint Resolve vmatch_itop vmatch_undef_itop vmatch_ftop vmatch_undef_ftop : va. + +(* Some properties about [is_uns] and [is_sgn]. *) + +Lemma is_uns_mon: forall n1 n2 i, is_uns n1 i -> n1 <= n2 -> is_uns n2 i. +Proof. + intros; red; intros. apply H; omega. +Qed. + +Lemma is_sgn_mon: forall n1 n2 i, is_sgn n1 i -> n1 <= n2 -> is_sgn n2 i. +Proof. + intros; red; intros. apply H; omega. +Qed. + +Lemma is_uns_sgn: forall n1 n2 i, is_uns n1 i -> n1 < n2 -> is_sgn n2 i. +Proof. + intros; red; intros. rewrite ! H by omega. auto. +Qed. + +Definition usize := Int.size. + +Definition ssize (i: int) := Int.size (if Int.lt i Int.zero then Int.not i else i) + 1. + +Lemma is_uns_usize: + forall i, is_uns (usize i) i. +Proof. + unfold usize; intros; red; intros. + apply Int.bits_size_2. omega. +Qed. + +Lemma is_sgn_ssize: + forall i, is_sgn (ssize i) i. +Proof. + unfold ssize; intros; red; intros. + destruct (Int.lt i Int.zero) eqn:LT. +- rewrite <- (negb_involutive (Int.testbit i m)). + rewrite <- (negb_involutive (Int.testbit i (Int.zwordsize - 1))). + f_equal. + generalize (Int.size_range (Int.not i)); intros RANGE. + rewrite <- ! Int.bits_not by omega. + rewrite ! Int.bits_size_2 by omega. + auto. +- rewrite ! Int.bits_size_2 by omega. + auto. +Qed. + +Lemma is_uns_zero_ext: + forall n i, is_uns n i <-> Int.zero_ext n i = i. +Proof. + intros; split; intros. + Int.bit_solve. destruct (zlt i0 n); auto. symmetry; apply H; auto. omega. + rewrite <- H. red; intros. rewrite Int.bits_zero_ext by omega. rewrite zlt_false by omega. auto. +Qed. + +Lemma is_sgn_sign_ext: + forall n i, 0 < n -> (is_sgn n i <-> Int.sign_ext n i = i). +Proof. + intros; split; intros. + Int.bit_solve. destruct (zlt i0 n); auto. + transitivity (Int.testbit i (Int.zwordsize - 1)). + apply H0; omega. symmetry; apply H0; omega. + rewrite <- H0. red; intros. rewrite ! Int.bits_sign_ext by omega. + f_equal. transitivity (n-1). destruct (zlt m n); omega. + destruct (zlt (Int.zwordsize - 1) n); omega. +Qed. + +Lemma is_zero_ext_uns: + forall i n m, + is_uns m i \/ n <= m -> is_uns m (Int.zero_ext n i). +Proof. + intros. red; intros. rewrite Int.bits_zero_ext by omega. + destruct (zlt m0 n); auto. destruct H. apply H; omega. omegaContradiction. +Qed. + +Lemma is_zero_ext_sgn: + forall i n m, + n < m -> + is_sgn m (Int.zero_ext n i). +Proof. + intros. red; intros. rewrite ! Int.bits_zero_ext by omega. + transitivity false. apply zlt_false; omega. + symmetry; apply zlt_false; omega. +Qed. + +Lemma is_sign_ext_uns: + forall i n m, + 0 <= m < n -> + is_uns m i -> + is_uns m (Int.sign_ext n i). +Proof. + intros; red; intros. rewrite Int.bits_sign_ext by omega. + apply H0. destruct (zlt m0 n); omega. destruct (zlt m0 n); omega. +Qed. + +Lemma is_sign_ext_sgn: + forall i n m, + 0 < n -> 0 < m -> + is_sgn m i \/ n <= m -> is_sgn m (Int.sign_ext n i). +Proof. + intros. apply is_sgn_sign_ext; auto. + destruct (zlt m n). destruct H1. apply is_sgn_sign_ext in H1; auto. + rewrite <- H1. rewrite (Int.sign_ext_widen i) by omega. apply Int.sign_ext_idem; auto. + omegaContradiction. + apply Int.sign_ext_widen; omega. +Qed. + +Hint Resolve is_uns_mon is_sgn_mon is_uns_sgn is_uns_usize is_sgn_ssize : va. + +(** Smart constructors for [Uns] and [Sgn]. *) + +Definition uns (n: Z) : aval := + if zle n 1 then Uns 1 + else if zle n 7 then Uns 7 + else if zle n 8 then Uns 8 + else if zle n 15 then Uns 15 + else if zle n 16 then Uns 16 + else itop. + +Definition sgn (n: Z) : aval := + if zle n 8 then Sgn 8 else if zle n 16 then Sgn 16 else itop. + +Lemma vmatch_uns': + forall i n, is_uns (Zmax 0 n) i -> vmatch (Vint i) (uns n). +Proof. + intros. + assert (A: forall n', n' >= 0 -> n' >= n -> is_uns n' i) by (eauto with va). + unfold uns. + destruct (zle n 1). auto with va. + destruct (zle n 7). auto with va. + destruct (zle n 8). auto with va. + destruct (zle n 15). auto with va. + destruct (zle n 16). auto with va. + auto with va. +Qed. + +Lemma vmatch_uns: + forall i n, is_uns n i -> vmatch (Vint i) (uns n). +Proof. + intros. apply vmatch_uns'. eauto with va. +Qed. + +Lemma vmatch_uns_undef: forall n, vmatch Vundef (uns n). +Proof. + intros. unfold uns. + destruct (zle n 1). auto with va. + destruct (zle n 7). auto with va. + destruct (zle n 8). auto with va. + destruct (zle n 15). auto with va. + destruct (zle n 16); auto with va. +Qed. + +Lemma vmatch_sgn': + forall i n, is_sgn (Zmax 1 n) i -> vmatch (Vint i) (sgn n). +Proof. + intros. + assert (A: forall n', n' >= 1 -> n' >= n -> is_sgn n' i) by (eauto with va). + unfold sgn. + destruct (zle n 8). auto with va. + destruct (zle n 16); auto with va. +Qed. + +Lemma vmatch_sgn: + forall i n, is_sgn n i -> vmatch (Vint i) (sgn n). +Proof. + intros. apply vmatch_sgn'. eauto with va. +Qed. + +Lemma vmatch_sgn_undef: forall n, vmatch Vundef (sgn n). +Proof. + intros. unfold sgn. + destruct (zle n 8). auto with va. + destruct (zle n 16); auto with va. +Qed. + +Hint Resolve vmatch_uns vmatch_uns_undef vmatch_sgn vmatch_sgn_undef : va. + +(** Ordering *) + +Inductive vge: aval -> aval -> Prop := + | vge_bot: forall v, vge v Vbot + | vge_i: forall i, vge (I i) (I i) + | vge_l: forall i, vge (L i) (L i) + | vge_f: forall f, vge (F f) (F f) + | vge_uns_i: forall n i, 0 <= n -> is_uns n i -> vge (Uns n) (I i) + | vge_uns_uns: forall n1 n2, n1 >= n2 -> vge (Uns n1) (Uns n2) + | vge_sgn_i: forall n i, 0 < n -> is_sgn n i -> vge (Sgn n) (I i) + | vge_sgn_sgn: forall n1 n2, n1 >= n2 -> vge (Sgn n1) (Sgn n2) + | vge_sgn_uns: forall n1 n2, n1 > n2 -> vge (Sgn n1) (Uns n2) + | vge_single_f: forall f, Float.is_single f -> vge Fsingle (F f) + | vge_single: vge Fsingle Fsingle + | vge_p_p: forall p q, pge p q -> vge (Ptr p) (Ptr q) + | vge_ip_p: forall p q, pge p q -> vge (Ifptr p) (Ptr q) + | vge_ip_ip: forall p q, pge p q -> vge (Ifptr p) (Ifptr q) + | vge_ip_i: forall p i, vge (Ifptr p) (I i) + | vge_ip_l: forall p i, vge (Ifptr p) (L i) + | vge_ip_f: forall p f, vge (Ifptr p) (F f) + | vge_ip_uns: forall p n, vge (Ifptr p) (Uns n) + | vge_ip_sgn: forall p n, vge (Ifptr p) (Sgn n) + | vge_ip_single: forall p, vge (Ifptr p) Fsingle. + +Hint Constructors vge : va. + +Lemma vge_top: forall v, vge Vtop v. +Proof. + destruct v; constructor; constructor. +Qed. + +Hint Resolve vge_top : va. + +Lemma vge_refl: forall v, vge v v. +Proof. + destruct v; auto with va. +Qed. + +Lemma vge_trans: forall u v, vge u v -> forall w, vge v w -> vge u w. +Proof. + induction 1; intros w V; inv V; eauto with va; constructor; eapply pge_trans; eauto. +Qed. + +Lemma vmatch_ge: + forall v x y, vge x y -> vmatch v y -> vmatch v x. +Proof. + induction 1; intros V; inv V; eauto with va; constructor; eapply pmatch_ge; eauto. +Qed. + +(** Least upper bound *) + +Definition vlub (v w: aval) : aval := + match v, w with + | Vbot, _ => w + | _, Vbot => v + | I i1, I i2 => + if Int.eq i1 i2 then v else + if Int.lt i1 Int.zero || Int.lt i2 Int.zero + then sgn (Z.max (ssize i1) (ssize i2)) + else uns (Z.max (usize i1) (usize i2)) + | I i, Uns n | Uns n, I i => + if Int.lt i Int.zero + then sgn (Z.max (ssize i) (n + 1)) + else uns (Z.max (usize i) n) + | I i, Sgn n | Sgn n, I i => + sgn (Z.max (ssize i) n) + | I i, (Ptr p | Ifptr p) | (Ptr p | Ifptr p), I i => + if strict || Int.eq i Int.zero then Ifptr p else Vtop + | Uns n1, Uns n2 => Uns (Z.max n1 n2) + | Uns n1, Sgn n2 => sgn (Z.max (n1 + 1) n2) + | Sgn n1, Uns n2 => sgn (Z.max n1 (n2 + 1)) + | Sgn n1, Sgn n2 => sgn (Z.max n1 n2) + | F f1, F f2 => + if Float.eq_dec f1 f2 then v else + if Float.is_single_dec f1 && Float.is_single_dec f2 then Fsingle else ftop + | F f, Fsingle | Fsingle, F f => + if Float.is_single_dec f then Fsingle else ftop + | Fsingle, Fsingle => + Fsingle + | L i1, L i2 => + if Int64.eq i1 i2 then v else ltop + | Ptr p1, Ptr p2 => Ptr(plub p1 p2) + | Ptr p1, Ifptr p2 => Ifptr(plub p1 p2) + | Ifptr p1, Ptr p2 => Ifptr(plub p1 p2) + | Ifptr p1, Ifptr p2 => Ifptr(plub p1 p2) + | _, (Ptr p | Ifptr p) | (Ptr p | Ifptr p), _ => if strict then Ifptr p else Vtop + | _, _ => Vtop + end. + +Lemma vlub_comm: + forall v w, vlub v w = vlub w v. +Proof. + intros. unfold vlub; destruct v; destruct w; auto. +- rewrite Int.eq_sym. predSpec Int.eq Int.eq_spec n0 n. + congruence. + rewrite orb_comm. + destruct (Int.lt n0 Int.zero || Int.lt n Int.zero); f_equal; apply Z.max_comm. +- f_equal; apply Z.max_comm. +- f_equal; apply Z.max_comm. +- f_equal; apply Z.max_comm. +- f_equal; apply Z.max_comm. +- rewrite Int64.eq_sym. predSpec Int64.eq Int64.eq_spec n0 n; congruence. +- rewrite dec_eq_sym. destruct (Float.eq_dec f0 f). congruence. + rewrite andb_comm. auto. +- f_equal; apply plub_comm. +- f_equal; apply plub_comm. +- f_equal; apply plub_comm. +- f_equal; apply plub_comm. +Qed. + +Lemma vge_uns_uns': forall n, vge (uns n) (Uns n). +Proof. + unfold uns, itop; intros. + destruct (zle n 1). auto with va. + destruct (zle n 7). auto with va. + destruct (zle n 8). auto with va. + destruct (zle n 15). auto with va. + destruct (zle n 16); auto with va. +Qed. + +Lemma vge_uns_i': forall n i, 0 <= n -> is_uns n i -> vge (uns n) (I i). +Proof. + intros. apply vge_trans with (Uns n). apply vge_uns_uns'. auto with va. +Qed. + +Lemma vge_sgn_sgn': forall n, vge (sgn n) (Sgn n). +Proof. + unfold sgn, itop; intros. + destruct (zle n 8). auto with va. + destruct (zle n 16); auto with va. +Qed. + +Lemma vge_sgn_i': forall n i, 0 < n -> is_sgn n i -> vge (sgn n) (I i). +Proof. + intros. apply vge_trans with (Sgn n). apply vge_sgn_sgn'. auto with va. +Qed. + +Hint Resolve vge_uns_uns' vge_uns_i' vge_sgn_sgn' vge_sgn_i' : va. + +Lemma usize_pos: forall n, 0 <= usize n. +Proof. + unfold usize; intros. generalize (Int.size_range n); omega. +Qed. + +Lemma ssize_pos: forall n, 0 < ssize n. +Proof. + unfold ssize; intros. + generalize (Int.size_range (if Int.lt n Int.zero then Int.not n else n)); omega. +Qed. + +Lemma vge_lub_l: + forall x y, vge (vlub x y) x. +Proof. + assert (IFSTRICT: forall (cond: bool) x y, vge x y -> vge (if cond then x else Vtop ) y). + { destruct cond; auto with va. } + unfold vlub; destruct x, y; eauto with va. +- predSpec Int.eq Int.eq_spec n n0. auto with va. + destruct (Int.lt n Int.zero || Int.lt n0 Int.zero). + apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va. + apply vge_uns_i'. generalize (usize_pos n); xomega. eauto with va. +- destruct (Int.lt n Int.zero). + apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va. + apply vge_uns_i'. generalize (usize_pos n); xomega. eauto with va. +- apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va. +- destruct (Int.lt n0 Int.zero). + eapply vge_trans. apply vge_sgn_sgn'. + apply vge_trans with (Sgn (n + 1)); eauto with va. + eapply vge_trans. apply vge_uns_uns'. eauto with va. +- eapply vge_trans. apply vge_sgn_sgn'. + apply vge_trans with (Sgn (n + 1)); eauto with va. +- eapply vge_trans. apply vge_sgn_sgn'. eauto with va. +- eapply vge_trans. apply vge_sgn_sgn'. eauto with va. +- eapply vge_trans. apply vge_sgn_sgn'. eauto with va. +- destruct (Int64.eq n n0); constructor. +- destruct (Float.eq_dec f f0). constructor. + destruct (Float.is_single_dec f && Float.is_single_dec f0) eqn:FS. + InvBooleans. auto with va. + constructor. +- destruct (Float.is_single_dec f); constructor; auto. +- destruct (Float.is_single_dec f); constructor; auto. +- constructor; apply pge_lub_l; auto. +- constructor; apply pge_lub_l; auto. +- constructor; apply pge_lub_l; auto. +- constructor; apply pge_lub_l; auto. +Qed. + +Lemma vge_lub_r: + forall x y, vge (vlub x y) y. +Proof. + intros. rewrite vlub_comm. apply vge_lub_l. +Qed. + +Lemma vmatch_lub_l: + forall v x y, vmatch v x -> vmatch v (vlub x y). +Proof. + intros. eapply vmatch_ge; eauto. apply vge_lub_l. +Qed. + +Lemma vmatch_lub_r: + forall v x y, vmatch v y -> vmatch v (vlub x y). +Proof. + intros. rewrite vlub_comm. apply vmatch_lub_l; auto. +Qed. + +Definition aptr_of_aval (v: aval) : aptr := + match v with + | Ptr p => p + | Ifptr p => p + | _ => Pbot + end. + +Lemma match_aptr_of_aval: + forall b ofs av, + vmatch (Vptr b ofs) av <-> pmatch b ofs (aptr_of_aval av). +Proof. + unfold aptr_of_aval; intros; split; intros. +- inv H; auto. +- destruct av; ((constructor; assumption) || inv H). +Qed. + +Definition vplub (v: aval) (p: aptr) : aptr := + match v with + | Ptr q => plub q p + | Ifptr q => plub q p + | _ => p + end. + +Lemma vmatch_vplub_l: + forall v x p, vmatch v x -> vmatch v (Ifptr (vplub x p)). +Proof. + intros. unfold vplub; inv H; auto with va; constructor; eapply pmatch_lub_l; eauto. +Qed. + +Lemma pmatch_vplub: + forall b ofs x p, pmatch b ofs p -> pmatch b ofs (vplub x p). +Proof. + intros. + assert (DFL: pmatch b ofs (if strict then p else Ptop)). + { destruct strict; auto. eapply pmatch_top'; eauto. } + unfold vplub; destruct x; auto; apply pmatch_lub_r; auto. +Qed. + +Lemma vmatch_vplub_r: + forall v x p, vmatch v (Ifptr p) -> vmatch v (Ifptr (vplub x p)). +Proof. + intros. apply vmatch_ifptr; intros; subst v. inv H. apply pmatch_vplub; auto. +Qed. + +(** Inclusion *) + +Definition vpincl (v: aval) (p: aptr) : bool := + match v with + | Ptr q => pincl q p + | Ifptr q => pincl q p + | _ => true + end. + +Lemma vpincl_ge: + forall x p, vpincl x p = true -> vge (Ifptr p) x. +Proof. + unfold vpincl; intros. destruct x; constructor; apply pincl_ge; auto. +Qed. + +Lemma vpincl_sound: + forall v x p, vpincl x p = true -> vmatch v x -> vmatch v (Ifptr p). +Proof. + intros. apply vmatch_ge with x; auto. apply vpincl_ge; auto. +Qed. + +Definition vincl (v w: aval) : bool := + match v, w with + | Vbot, _ => true + | I i, I j => Int.eq_dec i j + | I i, Uns n => Int.eq_dec (Int.zero_ext n i) i && zle 0 n + | I i, Sgn n => Int.eq_dec (Int.sign_ext n i) i && zlt 0 n + | Uns n, Uns m => zle n m + | Uns n, Sgn m => zlt n m + | Sgn n, Sgn m => zle n m + | L i, L j => Int64.eq_dec i j + | F i, F j => Float.eq_dec i j + | F i, Fsingle => Float.is_single_dec i + | Fsingle, Fsingle => true + | Ptr p, Ptr q => pincl p q + | Ptr p, Ifptr q => pincl p q + | Ifptr p, Ifptr q => pincl p q + | _, Ifptr _ => true + | _, _ => false + end. + +Lemma vincl_ge: forall v w, vincl v w = true -> vge w v. +Proof. + unfold vincl; destruct v; destruct w; intros; try discriminate; auto with va. + InvBooleans. subst; auto with va. + InvBooleans. constructor; auto. rewrite is_uns_zero_ext; auto. + InvBooleans. constructor; auto. rewrite is_sgn_sign_ext; auto. + InvBooleans. constructor; auto with va. + InvBooleans. constructor; auto with va. + InvBooleans. constructor; auto with va. + InvBooleans. subst; auto with va. + InvBooleans. subst; auto with va. + InvBooleans. auto with va. + constructor; apply pincl_ge; auto. + constructor; apply pincl_ge; auto. + constructor; apply pincl_ge; auto. +Qed. + +(** Loading constants *) + +Definition genv_match (ge: genv) : Prop := + (forall id b, Genv.find_symbol ge id = Some b <-> bc b = BCglob id) +/\(forall b, Plt b (Genv.genv_next ge) -> bc b <> BCinvalid /\ bc b <> BCstack). + +Definition symbol_address (ge: genv) (id: ident) (ofs: int) : val := + match Genv.find_symbol ge id with Some b => Vptr b ofs | None => Vundef end. + +Lemma symbol_address_sound: + forall ge id ofs, + genv_match ge -> + vmatch (symbol_address ge id ofs) (Ptr (Gl id ofs)). +Proof. + intros. unfold symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:F. + constructor. constructor. apply H; auto. + constructor. +Qed. + +Lemma vmatch_ptr_gl: + forall ge v id ofs, + genv_match ge -> + vmatch v (Ptr (Gl id ofs)) -> + Val.lessdef v (symbol_address ge id ofs). +Proof. + intros. unfold symbol_address. inv H0. +- inv H3. replace (Genv.find_symbol ge id) with (Some b). constructor. + symmetry. apply H; auto. +- constructor. +Qed. + +Lemma vmatch_ptr_stk: + forall v ofs sp, + vmatch v (Ptr(Stk ofs)) -> + bc sp = BCstack -> + Val.lessdef v (Vptr sp ofs). +Proof. + intros. inv H. +- inv H3. replace b with sp by (eapply bc_stack; eauto). constructor. +- constructor. +Qed. + +(** Generic operations that just do constant propagation. *) + +Definition unop_int (sem: int -> int) (x: aval) := + match x with I n => I (sem n) | _ => itop end. + +Lemma unop_int_sound: + forall sem v x, + vmatch v x -> + vmatch (match v with Vint i => Vint(sem i) | _ => Vundef end) (unop_int sem x). +Proof. + intros. unfold unop_int; inv H; auto with va. +Qed. + +Definition binop_int (sem: int -> int -> int) (x y: aval) := + match x, y with I n, I m => I (sem n m) | _, _ => itop end. + +Lemma binop_int_sound: + forall sem v x w y, + vmatch v x -> vmatch w y -> + vmatch (match v, w with Vint i, Vint j => Vint(sem i j) | _, _ => Vundef end) (binop_int sem x y). +Proof. + intros. unfold binop_int; inv H; auto with va; inv H0; auto with va. +Qed. + +Definition unop_float (sem: float -> float) (x: aval) := + match x with F n => F (sem n) | _ => ftop end. + +Lemma unop_float_sound: + forall sem v x, + vmatch v x -> + vmatch (match v with Vfloat i => Vfloat(sem i) | _ => Vundef end) (unop_float sem x). +Proof. + intros. unfold unop_float; inv H; auto with va. +Qed. + +Definition binop_float (sem: float -> float -> float) (x y: aval) := + match x, y with F n, F m => F (sem n m) | _, _ => ftop end. + +Lemma binop_float_sound: + forall sem v x w y, + vmatch v x -> vmatch w y -> + vmatch (match v, w with Vfloat i, Vfloat j => Vfloat(sem i j) | _, _ => Vundef end) (binop_float sem x y). +Proof. + intros. unfold binop_float; inv H; auto with va; inv H0; auto with va. +Qed. + +(** Logical operations *) + +Definition shl (v w: aval) := + match w with + | I amount => + if Int.ltu amount Int.iwordsize then + match v with + | I i => I (Int.shl i amount) + | Uns n => uns (n + Int.unsigned amount) + | Sgn n => sgn (n + Int.unsigned amount) + | _ => itop + end + else itop + | _ => itop + end. + +Lemma shl_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shl v w) (shl x y). +Proof. + intros. + assert (DEFAULT: vmatch (Val.shl v w) itop). + { + destruct v; destruct w; simpl; try constructor. + destruct (Int.ltu i0 Int.iwordsize); constructor. + } + destruct y; auto. simpl. inv H0. unfold Val.shl. + destruct (Int.ltu n Int.iwordsize) eqn:LTU; auto. + exploit Int.ltu_inv; eauto. intros RANGE. + inv H; auto with va. +- apply vmatch_uns'. red; intros. rewrite Int.bits_shl by omega. + destruct (zlt m (Int.unsigned n)). auto. apply H1; xomega. +- apply vmatch_sgn'. red; intros. zify. + rewrite ! Int.bits_shl by omega. + rewrite ! zlt_false by omega. + rewrite H1 by omega. symmetry. rewrite H1 by omega. auto. +- destruct v; constructor. +Qed. + +Definition shru (v w: aval) := + match w with + | I amount => + if Int.ltu amount Int.iwordsize then + match v with + | I i => I (Int.shru i amount) + | Uns n => uns (n - Int.unsigned amount) + | _ => uns (Int.zwordsize - Int.unsigned amount) + end + else itop + | _ => itop + end. + +Lemma shru_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shru v w) (shru x y). +Proof. + intros. + assert (DEFAULT: vmatch (Val.shru v w) itop). + { + destruct v; destruct w; simpl; try constructor. + destruct (Int.ltu i0 Int.iwordsize); constructor. + } + destruct y; auto. inv H0. unfold shru, Val.shru. + destruct (Int.ltu n Int.iwordsize) eqn:LTU; auto. + exploit Int.ltu_inv; eauto. intros RANGE. change (Int.unsigned Int.iwordsize) with Int.zwordsize in RANGE. + assert (DEFAULT2: forall i, vmatch (Vint (Int.shru i n)) (uns (Int.zwordsize - Int.unsigned n))). + { + intros. apply vmatch_uns. red; intros. + rewrite Int.bits_shru by omega. apply zlt_false. omega. + } + inv H; auto with va. +- apply vmatch_uns'. red; intros. zify. + rewrite Int.bits_shru by omega. + destruct (zlt (m + Int.unsigned n) Int.zwordsize); auto. + apply H1; omega. +- destruct v; constructor. +Qed. + +Definition shr (v w: aval) := + match w with + | I amount => + if Int.ltu amount Int.iwordsize then + match v with + | I i => I (Int.shr i amount) + | Uns n => sgn (n + 1 - Int.unsigned amount) + | Sgn n => sgn (n - Int.unsigned amount) + | _ => sgn (Int.zwordsize - Int.unsigned amount) + end + else itop + | _ => itop + end. + +Lemma shr_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shr v w) (shr x y). +Proof. + intros. + assert (DEFAULT: vmatch (Val.shr v w) itop). + { + destruct v; destruct w; simpl; try constructor. + destruct (Int.ltu i0 Int.iwordsize); constructor. + } + destruct y; auto. inv H0. unfold shr, Val.shr. + destruct (Int.ltu n Int.iwordsize) eqn:LTU; auto. + exploit Int.ltu_inv; eauto. intros RANGE. change (Int.unsigned Int.iwordsize) with Int.zwordsize in RANGE. + assert (DEFAULT2: forall i, vmatch (Vint (Int.shr i n)) (sgn (Int.zwordsize - Int.unsigned n))). + { + intros. apply vmatch_sgn. red; intros. + rewrite ! Int.bits_shr by omega. f_equal. + destruct (zlt (m + Int.unsigned n) Int.zwordsize); + destruct (zlt (Int.zwordsize - 1 + Int.unsigned n) Int.zwordsize); + omega. + } + assert (SGN: forall i p, is_sgn p i -> 0 < p -> vmatch (Vint (Int.shr i n)) (sgn (p - Int.unsigned n))). + { + intros. apply vmatch_sgn'. red; intros. zify. + rewrite ! Int.bits_shr by omega. + transitivity (Int.testbit i (Int.zwordsize - 1)). + destruct (zlt (m + Int.unsigned n) Int.zwordsize). + apply H0; omega. + auto. + symmetry. + destruct (zlt (Int.zwordsize - 1 + Int.unsigned n) Int.zwordsize). + apply H0; omega. + auto. + } + inv H; eauto with va. +- destruct v; constructor. +Qed. + +Definition and (v w: aval) := + match v, w with + | I i1, I i2 => I (Int.and i1 i2) + | I i, Uns n | Uns n, I i => uns (Z.min n (usize i)) + | I i, _ | _, I i => uns (usize i) + | Uns n1, Uns n2 => uns (Z.min n1 n2) + | Uns n, _ | _, Uns n => uns n + | Sgn n1, Sgn n2 => sgn (Z.max n1 n2) + | _, _ => itop + end. + +Lemma and_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.and v w) (and x y). +Proof. + assert (UNS_l: forall i j n, is_uns n i -> is_uns n (Int.and i j)). + { + intros; red; intros. rewrite Int.bits_and by auto. rewrite (H m) by auto. + apply andb_false_l. + } + assert (UNS_r: forall i j n, is_uns n i -> is_uns n (Int.and j i)). + { + intros. rewrite Int.and_commut. eauto. + } + assert (UNS: forall i j n m, is_uns n i -> is_uns m j -> is_uns (Z.min n m) (Int.and i j)). + { + intros. apply Z.min_case; auto. + } + assert (SGN: forall i j n m, is_sgn n i -> is_sgn m j -> is_sgn (Z.max n m) (Int.and i j)). + { + intros; red; intros. rewrite ! Int.bits_and by auto with va. + rewrite H by auto with va. rewrite H0 by auto with va. auto. + } + intros. unfold and, Val.and; inv H; eauto with va; inv H0; eauto with va. +Qed. + +Definition or (v w: aval) := + match v, w with + | I i1, I i2 => I (Int.or i1 i2) + | I i, Uns n | Uns n, I i => uns (Z.max n (usize i)) + | Uns n1, Uns n2 => uns (Z.max n1 n2) + | Sgn n1, Sgn n2 => sgn (Z.max n1 n2) + | _, _ => itop + end. + +Lemma or_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.or v w) (or x y). +Proof. + assert (UNS: forall i j n m, is_uns n i -> is_uns m j -> is_uns (Z.max n m) (Int.or i j)). + { + intros; red; intros. rewrite Int.bits_or by auto. + rewrite H by xomega. rewrite H0 by xomega. auto. + } + assert (SGN: forall i j n m, is_sgn n i -> is_sgn m j -> is_sgn (Z.max n m) (Int.or i j)). + { + intros; red; intros. rewrite ! Int.bits_or by xomega. + rewrite H by xomega. rewrite H0 by xomega. auto. + } + intros. unfold or, Val.or; inv H; eauto with va; inv H0; eauto with va. +Qed. + +Definition xor (v w: aval) := + match v, w with + | I i1, I i2 => I (Int.xor i1 i2) + | I i, Uns n | Uns n, I i => uns (Z.max n (usize i)) + | Uns n1, Uns n2 => uns (Z.max n1 n2) + | Sgn n1, Sgn n2 => sgn (Z.max n1 n2) + | _, _ => itop + end. + +Lemma xor_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.xor v w) (xor x y). +Proof. + assert (UNS: forall i j n m, is_uns n i -> is_uns m j -> is_uns (Z.max n m) (Int.xor i j)). + { + intros; red; intros. rewrite Int.bits_xor by auto. + rewrite H by xomega. rewrite H0 by xomega. auto. + } + assert (SGN: forall i j n m, is_sgn n i -> is_sgn m j -> is_sgn (Z.max n m) (Int.xor i j)). + { + intros; red; intros. rewrite ! Int.bits_xor by xomega. + rewrite H by xomega. rewrite H0 by xomega. auto. + } + intros. unfold xor, Val.xor; inv H; eauto with va; inv H0; eauto with va. +Qed. + +Definition notint (v: aval) := + match v with + | I i => I (Int.not i) + | Uns n => sgn (n + 1) + | Sgn n => Sgn n + | _ => itop + end. + +Lemma notint_sound: + forall v x, vmatch v x -> vmatch (Val.notint v) (notint x). +Proof. + assert (SGN: forall n i, is_sgn n i -> is_sgn n (Int.not i)). + { + intros; red; intros. rewrite ! Int.bits_not by omega. + f_equal. apply H; auto. + } + intros. unfold Val.notint, notint; inv H; eauto with va. +Qed. + +Definition ror (x y: aval) := + match y, x with + | I j, I i => if Int.ltu j Int.iwordsize then I(Int.ror i j) else itop + | _, _ => itop + end. + +Lemma ror_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.ror v w) (ror x y). +Proof. + intros. + assert (DEFAULT: vmatch (Val.ror v w) itop). + { + destruct v; destruct w; simpl; try constructor. + destruct (Int.ltu i0 Int.iwordsize); constructor. + } + unfold ror; destruct y; auto. inv H0. unfold Val.ror. + destruct (Int.ltu n Int.iwordsize) eqn:LTU. + inv H; auto with va. + inv H; auto with va. +Qed. + +Definition rolm (x: aval) (amount mask: int) := + match x with + | I i => I (Int.rolm i amount mask) + | _ => uns (usize mask) + end. + +Lemma rolm_sound: + forall v x amount mask, + vmatch v x -> vmatch (Val.rolm v amount mask) (rolm x amount mask). +Proof. + intros. + assert (UNS: forall i, vmatch (Vint (Int.rolm i amount mask)) (uns (usize mask))). + { + intros. + change (vmatch (Val.and (Vint (Int.rol i amount)) (Vint mask)) + (and itop (I mask))). + apply and_sound; auto with va. + } + unfold Val.rolm, rolm. inv H; auto with va. +Qed. + +(** Integer arithmetic operations *) + +Definition neg := unop_int Int.neg. + +Lemma neg_sound: + forall v x, vmatch v x -> vmatch (Val.neg v) (neg x). +Proof (unop_int_sound Int.neg). + +Definition add (x y: aval) := + match x, y with + | I i, I j => I (Int.add i j) + | Ptr p, I i | I i, Ptr p => Ptr (padd p i) + | Ptr p, _ | _, Ptr p => Ptr (poffset p) + | Ifptr p, I i | I i, Ifptr p => Ifptr (padd p i) + | Ifptr p, Ifptr q => Ifptr (plub (poffset p) (poffset q)) + | Ifptr p, _ | _, Ifptr p => Ifptr (poffset p) + | _, _ => Vtop + end. + +Lemma add_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.add v w) (add x y). +Proof. + intros. unfold Val.add, add; inv H; inv H0; constructor; + ((apply padd_sound; assumption) || (eapply poffset_sound; eassumption) || idtac). + apply pmatch_lub_r. eapply poffset_sound; eauto. + apply pmatch_lub_l. eapply poffset_sound; eauto. +Qed. + +Definition sub (v w: aval) := + match v, w with + | I i1, I i2 => I (Int.sub i1 i2) + | Ptr p, I i => Ptr (psub p i) +(* problem with undefs *) +(* + | Ptr p1, Ptr p2 => match psub2 p1 p2 with Some n => I n | _ => itop end +*) + | Ptr p, _ => Ifptr (poffset p) + | Ifptr p, I i => Ifptr (psub p i) + | Ifptr p, (Uns _ | Sgn _) => Ifptr (poffset p) + | Ifptr p1, Ptr p2 => itop + | _, _ => Vtop + end. + +Lemma sub_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.sub v w) (sub x y). +Proof. + intros. inv H; inv H0; simpl; + try (destruct (eq_block b b0)); + try (constructor; (apply psub_sound || eapply poffset_sound); eauto). + change Vtop with (Ifptr (poffset Ptop)). + constructor; eapply poffset_sound. eapply pmatch_top'; eauto. +Qed. + +Definition mul := binop_int Int.mul. + +Lemma mul_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mul v w) (mul x y). +Proof (binop_int_sound Int.mul). + +Definition mulhs := binop_int Int.mulhs. + +Lemma mulhs_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulhs v w) (mulhs x y). +Proof (binop_int_sound Int.mulhs). + +Definition mulhu := binop_int Int.mulhu. + +Lemma mulhu_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulhu v w) (mulhu x y). +Proof (binop_int_sound Int.mulhu). + +Definition divs (v w: aval) := + match w, v with + | I i2, I i1 => + if Int.eq i2 Int.zero + || Int.eq i1 (Int.repr Int.min_signed) && Int.eq i2 Int.mone + then if strict then Vbot else itop + else I (Int.divs i1 i2) + | _, _ => itop + end. + +Lemma divs_sound: + forall v w u x y, vmatch v x -> vmatch w y -> Val.divs v w = Some u -> vmatch u (divs x y). +Proof. + intros. destruct v; destruct w; try discriminate; simpl in H1. + destruct (Int.eq i0 Int.zero + || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone) eqn:E; inv H1. + inv H; inv H0; auto with va. simpl. rewrite E. constructor. +Qed. + +Definition divu (v w: aval) := + match w, v with + | I i2, I i1 => + if Int.eq i2 Int.zero + then if strict then Vbot else itop + else I (Int.divu i1 i2) + | _, _ => itop + end. + +Lemma divu_sound: + forall v w u x y, vmatch v x -> vmatch w y -> Val.divu v w = Some u -> vmatch u (divu x y). +Proof. + intros. destruct v; destruct w; try discriminate; simpl in H1. + destruct (Int.eq i0 Int.zero) eqn:E; inv H1. + inv H; inv H0; auto with va. simpl. rewrite E. constructor. +Qed. + +Definition mods (v w: aval) := + match w, v with + | I i2, I i1 => + if Int.eq i2 Int.zero + || Int.eq i1 (Int.repr Int.min_signed) && Int.eq i2 Int.mone + then if strict then Vbot else itop + else I (Int.mods i1 i2) + | _, _ => itop + end. + +Lemma mods_sound: + forall v w u x y, vmatch v x -> vmatch w y -> Val.mods v w = Some u -> vmatch u (mods x y). +Proof. + intros. destruct v; destruct w; try discriminate; simpl in H1. + destruct (Int.eq i0 Int.zero + || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone) eqn:E; inv H1. + inv H; inv H0; auto with va. simpl. rewrite E. constructor. +Qed. + +Definition modu (v w: aval) := + match w, v with + | I i2, I i1 => + if Int.eq i2 Int.zero + then if strict then Vbot else itop + else I (Int.modu i1 i2) + | I i2, _ => uns (usize i2) + | _, _ => itop + end. + +Lemma modu_sound: + forall v w u x y, vmatch v x -> vmatch w y -> Val.modu v w = Some u -> vmatch u (modu x y). +Proof. + assert (UNS: forall i j, j <> Int.zero -> is_uns (usize j) (Int.modu i j)). + { + intros. apply is_uns_mon with (usize (Int.modu i j)); auto with va. + unfold usize, Int.size. apply Int.Zsize_monotone. + generalize (Int.unsigned_range_2 j); intros RANGE. + assert (Int.unsigned j <> 0). + { red; intros; elim H. rewrite <- (Int.repr_unsigned j). rewrite H0. auto. } + exploit (Z_mod_lt (Int.unsigned i) (Int.unsigned j)). omega. intros MOD. + unfold Int.modu. rewrite Int.unsigned_repr. omega. omega. + } + intros. destruct v; destruct w; try discriminate; simpl in H1. + destruct (Int.eq i0 Int.zero) eqn:Z; inv H1. + assert (i0 <> Int.zero) by (generalize (Int.eq_spec i0 Int.zero); rewrite Z; auto). + unfold modu. inv H; inv H0; auto with va. rewrite Z. constructor. +Qed. + +Definition shrx (v w: aval) := + match v, w with + | I i, I j => if Int.ltu j (Int.repr 31) then I(Int.shrx i j) else itop + | _, _ => itop + end. + +Lemma shrx_sound: + forall v w u x y, vmatch v x -> vmatch w y -> Val.shrx v w = Some u -> vmatch u (shrx x y). +Proof. + intros. + destruct v; destruct w; try discriminate; simpl in H1. + destruct (Int.ltu i0 (Int.repr 31)) eqn:LTU; inv H1. + unfold shrx; inv H; auto with va; inv H0; auto with va. + rewrite LTU; auto with va. +Qed. + +(** Floating-point arithmetic operations *) + +Definition negf := unop_float Float.neg. + +Lemma negf_sound: + forall v x, vmatch v x -> vmatch (Val.negf v) (negf x). +Proof (unop_float_sound Float.neg). + +Definition absf := unop_float Float.abs. + +Lemma absf_sound: + forall v x, vmatch v x -> vmatch (Val.absf v) (absf x). +Proof (unop_float_sound Float.abs). + +Definition addf := binop_float Float.add. + +Lemma addf_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.addf v w) (addf x y). +Proof (binop_float_sound Float.add). + +Definition subf := binop_float Float.sub. + +Lemma subf_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.subf v w) (subf x y). +Proof (binop_float_sound Float.sub). + +Definition mulf := binop_float Float.mul. + +Lemma mulf_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulf v w) (mulf x y). +Proof (binop_float_sound Float.mul). + +Definition divf := binop_float Float.div. + +Lemma divf_sound: + forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.divf v w) (divf x y). +Proof (binop_float_sound Float.div). + +(** Conversions *) + +Definition zero_ext (nbits: Z) (v: aval) := + match v with + | I i => I (Int.zero_ext nbits i) + | Uns n => uns (Z.min n nbits) + | _ => uns nbits + end. + +Lemma zero_ext_sound: + forall nbits v x, vmatch v x -> vmatch (Val.zero_ext nbits v) (zero_ext nbits x). +Proof. + assert (DFL: forall nbits i, is_uns nbits (Int.zero_ext nbits i)). + { + intros; red; intros. rewrite Int.bits_zero_ext by omega. apply zlt_false; auto. + } + intros. inv H; simpl; auto with va. apply vmatch_uns. + red; intros. zify. + rewrite Int.bits_zero_ext by omega. + destruct (zlt m nbits); auto. apply H1; omega. +Qed. + +Definition sign_ext (nbits: Z) (v: aval) := + match v with + | I i => I (Int.sign_ext nbits i) + | Uns n => if zlt n nbits then Uns n else sgn nbits + | Sgn n => sgn (Z.min n nbits) + | _ => sgn nbits + end. + +Lemma sign_ext_sound: + forall nbits v x, 0 < nbits -> vmatch v x -> vmatch (Val.sign_ext nbits v) (sign_ext nbits x). +Proof. + assert (DFL: forall nbits i, 0 < nbits -> vmatch (Vint (Int.sign_ext nbits i)) (sgn nbits)). + { + intros. apply vmatch_sgn. apply is_sign_ext_sgn; auto with va. + } + intros. inv H0; simpl; auto with va. +- destruct (zlt n nbits); eauto with va. + constructor; auto. eapply is_sign_ext_uns; eauto with va. +- destruct (zlt n nbits); auto with va. +- apply vmatch_sgn. apply is_sign_ext_sgn; auto with va. + apply Z.min_case; auto with va. +Qed. + +Definition singleoffloat (v: aval) := + match v with + | F f => F (Float.singleoffloat f) + | _ => Fsingle + end. + +Lemma singleoffloat_sound: + forall v x, vmatch v x -> vmatch (Val.singleoffloat v) (singleoffloat x). +Proof. + intros. + assert (DEFAULT: vmatch (Val.singleoffloat v) Fsingle). + { destruct v; constructor. apply Float.singleoffloat_is_single. } + destruct x; auto. inv H. constructor. +Qed. + +Definition intoffloat (x: aval) := + match x with + | F f => + match Float.intoffloat f with + | Some i => I i + | None => if strict then Vbot else itop + end + | _ => itop + end. + +Lemma intoffloat_sound: + forall v x w, vmatch v x -> Val.intoffloat v = Some w -> vmatch w (intoffloat x). +Proof. + unfold Val.intoffloat; intros. destruct v; try discriminate. + destruct (Float.intoffloat f) as [i|] eqn:E; simpl in H0; inv H0. + inv H; simpl; auto with va. rewrite E; constructor. +Qed. + +Definition intuoffloat (x: aval) := + match x with + | F f => + match Float.intuoffloat f with + | Some i => I i + | None => if strict then Vbot else itop + end + | _ => itop + end. + +Lemma intuoffloat_sound: + forall v x w, vmatch v x -> Val.intuoffloat v = Some w -> vmatch w (intuoffloat x). +Proof. + unfold Val.intuoffloat; intros. destruct v; try discriminate. + destruct (Float.intuoffloat f) as [i|] eqn:E; simpl in H0; inv H0. + inv H; simpl; auto with va. rewrite E; constructor. +Qed. + +Definition floatofint (x: aval) := + match x with + | I i => F(Float.floatofint i) + | _ => ftop + end. + +Lemma floatofint_sound: + forall v x w, vmatch v x -> Val.floatofint v = Some w -> vmatch w (floatofint x). +Proof. + unfold Val.floatofint; intros. destruct v; inv H0. + inv H; simpl; auto with va. +Qed. + +Definition floatofintu (x: aval) := + match x with + | I i => F(Float.floatofintu i) + | _ => ftop + end. + +Lemma floatofintu_sound: + forall v x w, vmatch v x -> Val.floatofintu v = Some w -> vmatch w (floatofintu x). +Proof. + unfold Val.floatofintu; intros. destruct v; inv H0. + inv H; simpl; auto with va. +Qed. + +Definition floatofwords (x y: aval) := + match x, y with + | I i, I j => F(Float.from_words i j) + | _, _ => ftop + end. + +Lemma floatofwords_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.floatofwords v w) (floatofwords x y). +Proof. + intros. unfold floatofwords, ftop; inv H; simpl; auto with va; inv H0; auto with va. +Qed. + +Definition longofwords (x y: aval) := + match x, y with + | I i, I j => L(Int64.ofwords i j) + | _, _ => ltop + end. + +Lemma longofwords_sound: + forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.longofwords v w) (longofwords x y). +Proof. + intros. unfold longofwords, ltop; inv H; simpl; auto with va; inv H0; auto with va. +Qed. + +Definition loword (x: aval) := + match x with + | L i => I(Int64.loword i) + | _ => itop + end. + +Lemma loword_sound: forall v x, vmatch v x -> vmatch (Val.loword v) (loword x). +Proof. + destruct 1; simpl; auto with va. +Qed. + +Definition hiword (x: aval) := + match x with + | L i => I(Int64.hiword i) + | _ => itop + end. + +Lemma hiword_sound: forall v x, vmatch v x -> vmatch (Val.hiword v) (hiword x). +Proof. + destruct 1; simpl; auto with va. +Qed. + +(** Comparisons *) + +Definition cmpu_bool (c: comparison) (v w: aval) : abool := + match v, w with + | I i1, I i2 => Just (Int.cmpu c i1 i2) +(* there are cute things to do with Uns/Sgn compared against an integer *) + | Ptr _, (I _ | Uns _ | Sgn _) => cmp_different_blocks c + | (I _ | Uns _ | Sgn _), Ptr _ => cmp_different_blocks c + | Ptr p1, Ptr p2 => pcmp c p1 p2 + | Ptr p1, Ifptr p2 => club (pcmp c p1 p2) (cmp_different_blocks c) + | Ifptr p1, Ptr p2 => club (pcmp c p1 p2) (cmp_different_blocks c) + | _, _ => Btop + end. + +Lemma cmpu_bool_sound: + forall valid c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmpu_bool valid c v w) (cmpu_bool c x y). +Proof. + intros. + assert (IP: forall i b ofs, + cmatch (Val.cmpu_bool valid c (Vint i) (Vptr b ofs)) (cmp_different_blocks c)). + { + intros. simpl. destruct (Int.eq i Int.zero). apply cmp_different_blocks_sound. apply cmp_different_blocks_none. + } + assert (PI: forall i b ofs, + cmatch (Val.cmpu_bool valid c (Vptr b ofs) (Vint i)) (cmp_different_blocks c)). + { + intros. simpl. destruct (Int.eq i Int.zero). apply cmp_different_blocks_sound. apply cmp_different_blocks_none. + } + unfold cmpu_bool; inv H; inv H0; + auto using cmatch_top, cmp_different_blocks_none, pcmp_none, + cmatch_lub_l, cmatch_lub_r, pcmp_sound. +- constructor. +Qed. + +Definition cmp_bool (c: comparison) (v w: aval) : abool := + match v, w with + | I i1, I i2 => Just (Int.cmp c i1 i2) + | _, _ => Btop + end. + +Lemma cmp_bool_sound: + forall c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmp_bool c v w) (cmp_bool c x y). +Proof. + intros. inv H; try constructor; inv H0; constructor. +Qed. + +Definition cmpf_bool (c: comparison) (v w: aval) : abool := + match v, w with + | F f1, F f2 => Just (Float.cmp c f1 f2) + | _, _ => Btop + end. + +Lemma cmpf_bool_sound: + forall c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmpf_bool c v w) (cmpf_bool c x y). +Proof. + intros. inv H; try constructor; inv H0; constructor. +Qed. + +Definition maskzero (x: aval) (mask: int) : abool := + match x with + | I i => Just (Int.eq (Int.and i mask) Int.zero) + | Uns n => if Int.eq (Int.zero_ext n mask) Int.zero then Maybe true else Btop + | _ => Btop + end. + +Lemma maskzero_sound: + forall mask v x, + vmatch v x -> + cmatch (Val.maskzero_bool v mask) (maskzero x mask). +Proof. + intros. inv H; simpl; auto with va. + predSpec Int.eq Int.eq_spec (Int.zero_ext n mask) Int.zero; auto with va. + replace (Int.and i mask) with Int.zero. + rewrite Int.eq_true. constructor. + rewrite is_uns_zero_ext in H1. rewrite Int.zero_ext_and in * by auto. + rewrite <- H1. rewrite Int.and_assoc. rewrite Int.and_commut in H. rewrite H. + rewrite Int.and_zero; auto. + destruct (Int.eq (Int.zero_ext n mask) Int.zero); constructor. +Qed. + +Definition of_optbool (ab: abool) : aval := + match ab with + | Just b => I (if b then Int.one else Int.zero) + | _ => Uns 1 + end. + +Lemma of_optbool_sound: + forall ob ab, cmatch ob ab -> vmatch (Val.of_optbool ob) (of_optbool ab). +Proof. + intros. + assert (DEFAULT: vmatch (Val.of_optbool ob) (Uns 1)). + { + destruct ob; simpl; auto with va. + destruct b; constructor; try omega. + change 1 with (usize Int.one). apply is_uns_usize. + red; intros. apply Int.bits_zero. + } + inv H; auto. simpl. destruct b; constructor. +Qed. + +Definition resolve_branch (ab: abool) : option bool := + match ab with + | Just b => Some b + | Maybe b => Some b + | _ => None + end. + +Lemma resolve_branch_sound: + forall b ab b', + cmatch (Some b) ab -> resolve_branch ab = Some b' -> b' = b. +Proof. + intros. inv H; simpl in H0; congruence. +Qed. + +(** Normalization at load time *) + +Definition vnormalize (chunk: memory_chunk) (v: aval) := + match chunk, v with + | _, Vbot => Vbot + | Mint8signed, I i => I (Int.sign_ext 8 i) + | Mint8signed, Uns n => if zlt n 8 then Uns n else Sgn 8 + | Mint8signed, Sgn n => Sgn (Z.min n 8) + | Mint8signed, _ => Sgn 8 + | Mint8unsigned, I i => I (Int.zero_ext 8 i) + | Mint8unsigned, Uns n => Uns (Z.min n 8) + | Mint8unsigned, _ => Uns 8 + | Mint16signed, I i => I (Int.sign_ext 16 i) + | Mint16signed, Uns n => if zlt n 16 then Uns n else Sgn 16 + | Mint16signed, Sgn n => Sgn (Z.min n 16) + | Mint16signed, _ => Sgn 16 + | Mint16unsigned, I i => I (Int.zero_ext 16 i) + | Mint16unsigned, Uns n => Uns (Z.min n 16) + | Mint16unsigned, _ => Uns 16 + | Mint32, (I _ | Ptr _ | Ifptr _) => v + | Mint64, L _ => v + | Mfloat32, F f => F (Float.singleoffloat f) + | Mfloat32, _ => Fsingle + | (Mfloat64 | Mfloat64al32), F f => v + | _, _ => Ifptr Pbot + end. + +Lemma vnormalize_sound: + forall chunk v x, vmatch v x -> vmatch (Val.load_result chunk v) (vnormalize chunk x). +Proof. + unfold Val.load_result, vnormalize; induction 1; destruct chunk; auto with va. +- destruct (zlt n 8); constructor; auto with va. + apply is_sign_ext_uns; auto. + apply is_sign_ext_sgn; auto with va. +- constructor. xomega. apply is_zero_ext_uns. apply Z.min_case; auto with va. +- destruct (zlt n 16); constructor; auto with va. + apply is_sign_ext_uns; auto. + apply is_sign_ext_sgn; auto with va. +- constructor. xomega. apply is_zero_ext_uns. apply Z.min_case; auto with va. +- destruct (zlt n 8); auto with va. +- destruct (zlt n 16); auto with va. +- constructor. xomega. apply is_sign_ext_sgn; auto with va. apply Z.min_case; auto with va. +- constructor. omega. apply is_zero_ext_uns; auto with va. +- constructor. xomega. apply is_sign_ext_sgn; auto with va. apply Z.min_case; auto with va. +- constructor. omega. apply is_zero_ext_uns; auto with va. +- constructor. apply Float.singleoffloat_is_single. +- constructor. omega. apply is_sign_ext_sgn; auto with va. +- constructor. omega. apply is_zero_ext_uns; auto with va. +- constructor. omega. apply is_sign_ext_sgn; auto with va. +- constructor. omega. apply is_zero_ext_uns; auto with va. +- constructor. apply Float.singleoffloat_is_single. +Qed. + +Lemma vnormalize_cast: + forall chunk m b ofs v p, + Mem.load chunk m b ofs = Some v -> + vmatch v (Ifptr p) -> + vmatch v (vnormalize chunk (Ifptr p)). +Proof. + intros. exploit Mem.load_cast; eauto. exploit Mem.load_type; eauto. + destruct chunk; simpl; intros. +- (* int8signed *) + rewrite H2. destruct v; simpl; constructor. omega. apply is_sign_ext_sgn; auto with va. +- (* int8unsigned *) + rewrite H2. destruct v; simpl; constructor. omega. apply is_zero_ext_uns; auto with va. +- (* int16signed *) + rewrite H2. destruct v; simpl; constructor. omega. apply is_sign_ext_sgn; auto with va. +- (* int16unsigned *) + rewrite H2. destruct v; simpl; constructor. omega. apply is_zero_ext_uns; auto with va. +- (* int32 *) + auto. +- (* int64 *) + destruct v; try contradiction; constructor. +- (* float32 *) + rewrite H2. destruct v; simpl; constructor. apply Float.singleoffloat_is_single. +- (* float64 *) + destruct v; try contradiction; constructor. +- (* float64 *) + destruct v; try contradiction; constructor. +Qed. + +Lemma vnormalize_monotone: + forall chunk x y, + vge x y -> vge (vnormalize chunk x) (vnormalize chunk y). +Proof. + induction 1; destruct chunk; simpl; auto with va. +- destruct (zlt n 8); constructor; auto with va. + apply is_sign_ext_uns; auto with va. + apply is_sign_ext_sgn; auto with va. +- constructor; auto with va. apply is_zero_ext_uns; auto with va. + apply Z.min_case; auto with va. +- destruct (zlt n 16); constructor; auto with va. + apply is_sign_ext_uns; auto with va. + apply is_sign_ext_sgn; auto with va. +- constructor; auto with va. apply is_zero_ext_uns; auto with va. + apply Z.min_case; auto with va. +- destruct (zlt n1 8). rewrite zlt_true by omega. auto with va. + destruct (zlt n2 8); auto with va. +- destruct (zlt n1 16). rewrite zlt_true by omega. auto with va. + destruct (zlt n2 16); auto with va. +- constructor; auto with va. apply is_sign_ext_sgn; auto with va. + apply Z.min_case; auto with va. +- constructor; auto with va. apply is_zero_ext_uns; auto with va. +- constructor; auto with va. apply is_sign_ext_sgn; auto with va. + apply Z.min_case; auto with va. +- constructor; auto with va. apply is_zero_ext_uns; auto with va. +- destruct (zlt n2 8); constructor; auto with va. +- destruct (zlt n2 16); constructor; auto with va. +- constructor. apply Float.singleoffloat_is_single. +- constructor; auto with va. apply is_sign_ext_sgn; auto with va. +- constructor; auto with va. apply is_zero_ext_uns; auto with va. +- constructor; auto with va. apply is_sign_ext_sgn; auto with va. +- constructor; auto with va. apply is_zero_ext_uns; auto with va. +- constructor. apply Float.singleoffloat_is_single. +- destruct (zlt n 8); constructor; auto with va. +- destruct (zlt n 16); constructor; auto with va. +Qed. + +(** Abstracting memory blocks *) + +Inductive acontent : Type := + | ACany + | ACval (chunk: memory_chunk) (av: aval). + +Definition eq_acontent : forall (c1 c2: acontent), {c1=c2} + {c1<>c2}. +Proof. + intros. generalize chunk_eq eq_aval. decide equality. +Defined. + +Record ablock : Type := ABlock { + ab_contents: ZMap.t acontent; + ab_summary: aptr; + ab_default: fst ab_contents = ACany +}. + +Local Notation "a ## b" := (ZMap.get b a) (at level 1). + +Definition ablock_init (p: aptr) : ablock := + {| ab_contents := ZMap.init ACany; ab_summary := p; ab_default := refl_equal _ |}. + +Definition chunk_compat (chunk chunk': memory_chunk) : bool := + match chunk, chunk' with + | (Mint8signed | Mint8unsigned), (Mint8signed | Mint8unsigned) => true + | (Mint16signed | Mint16unsigned), (Mint16signed | Mint16unsigned) => true + | Mint32, Mint32 => true + | Mfloat32, Mfloat32 => true + | Mint64, Mint64 => true + | (Mfloat64 | Mfloat64al32), Mfloat64 => true + | Mfloat64al32, Mfloat64al32 => true + | _, _ => false + end. + +Definition ablock_load (chunk: memory_chunk) (ab: ablock) (i: Z) : aval := + match ab.(ab_contents)##i with + | ACany => vnormalize chunk (Ifptr ab.(ab_summary)) + | ACval chunk' av => + if chunk_compat chunk chunk' + then vnormalize chunk av + else vnormalize chunk (Ifptr ab.(ab_summary)) + end. + +Definition ablock_load_anywhere (chunk: memory_chunk) (ab: ablock) : aval := + vnormalize chunk (Ifptr ab.(ab_summary)). + +Function inval_after (lo: Z) (hi: Z) (c: ZMap.t acontent) { wf (Zwf lo) hi } : ZMap.t acontent := + if zle lo hi + then inval_after lo (hi - 1) (ZMap.set hi ACany c) + else c. +Proof. + intros; red; omega. + apply Zwf_well_founded. +Qed. + +Definition inval_if (hi: Z) (lo: Z) (c: ZMap.t acontent) := + match c##lo with + | ACany => c + | ACval chunk av => if zle (lo + size_chunk chunk) hi then c else ZMap.set lo ACany c + end. + +Function inval_before (hi: Z) (lo: Z) (c: ZMap.t acontent) { wf (Zwf_up hi) lo } : ZMap.t acontent := + if zlt lo hi + then inval_before hi (lo + 1) (inval_if hi lo c) + else c. +Proof. + intros; red; omega. + apply Zwf_up_well_founded. +Qed. + +Remark fst_inval_after: forall lo hi c, fst (inval_after lo hi c) = fst c. +Proof. + intros. functional induction (inval_after lo hi c); auto. +Qed. + +Remark fst_inval_before: forall hi lo c, fst (inval_before hi lo c) = fst c. +Proof. + intros. functional induction (inval_before hi lo c); auto. + rewrite IHt. unfold inval_if. destruct c##lo; auto. + destruct (zle (lo + size_chunk chunk) hi); auto. +Qed. + +Program Definition ablock_store (chunk: memory_chunk) (ab: ablock) (i: Z) (av: aval) : ablock := + {| ab_contents := + ZMap.set i (ACval chunk av) + (inval_before i (i - 7) + (inval_after (i + 1) (i + size_chunk chunk - 1) ab.(ab_contents))); + ab_summary := + vplub av ab.(ab_summary); + ab_default := _ |}. +Next Obligation. + rewrite fst_inval_before, fst_inval_after. apply ab_default. +Qed. + +Definition ablock_store_anywhere (chunk: memory_chunk) (ab: ablock) (av: aval) : ablock := + ablock_init (vplub av ab.(ab_summary)). + +Definition ablock_loadbytes (ab: ablock) : aptr := ab.(ab_summary). + +Program Definition ablock_storebytes (ab: ablock) (p: aptr) (ofs: Z) (sz: Z) := + {| ab_contents := + inval_before ofs (ofs - 7) + (inval_after ofs (ofs + sz - 1) ab.(ab_contents)); + ab_summary := + plub p ab.(ab_summary); + ab_default := _ |}. +Next Obligation. + rewrite fst_inval_before, fst_inval_after. apply ab_default. +Qed. + +Definition ablock_storebytes_anywhere (ab: ablock) (p: aptr) := + ablock_init (plub p ab.(ab_summary)). + +Definition smatch (m: mem) (b: block) (p: aptr) : Prop := + (forall chunk ofs v, Mem.load chunk m b ofs = Some v -> vmatch v (Ifptr p)) +/\(forall ofs b' ofs' i, Mem.loadbytes m b ofs 1 = Some (Pointer b' ofs' i :: nil) -> pmatch b' ofs' p). + +Remark loadbytes_load_ext: + forall b m m', + (forall ofs n bytes, Mem.loadbytes m' b ofs n = Some bytes -> n >= 0 -> Mem.loadbytes m b ofs n = Some bytes) -> + forall chunk ofs v, Mem.load chunk m' b ofs = Some v -> Mem.load chunk m b ofs = Some v. +Proof. + intros. exploit Mem.load_loadbytes; eauto. intros [bytes [A B]]. + exploit Mem.load_valid_access; eauto. intros [C D]. + subst v. apply Mem.loadbytes_load; auto. apply H; auto. generalize (size_chunk_pos chunk); omega. +Qed. + +Lemma smatch_ext: + forall m b p m', + smatch m b p -> + (forall ofs n bytes, Mem.loadbytes m' b ofs n = Some bytes -> n >= 0 -> Mem.loadbytes m b ofs n = Some bytes) -> + smatch m' b p. +Proof. + intros. destruct H. split; intros. + eapply H; eauto. eapply loadbytes_load_ext; eauto. + eapply H1; eauto. apply H0; eauto. omega. +Qed. + +Lemma smatch_inv: + forall m b p m', + smatch m b p -> + (forall ofs n, n >= 0 -> Mem.loadbytes m' b ofs n = Mem.loadbytes m b ofs n) -> + smatch m' b p. +Proof. + intros. eapply smatch_ext; eauto. + intros. rewrite <- H0; eauto. +Qed. + +Lemma smatch_ge: + forall m b p q, smatch m b p -> pge q p -> smatch m b q. +Proof. + intros. destruct H as [A B]. split; intros. + apply vmatch_ge with (Ifptr p); eauto with va. + apply pmatch_ge with p; eauto with va. +Qed. + +Lemma In_loadbytes: + forall m b byte n ofs bytes, + Mem.loadbytes m b ofs n = Some bytes -> + In byte bytes -> + exists ofs', ofs <= ofs' < ofs + n /\ Mem.loadbytes m b ofs' 1 = Some(byte :: nil). +Proof. + intros until n. pattern n. + apply well_founded_ind with (R := Zwf 0). +- apply Zwf_well_founded. +- intros sz REC ofs bytes LOAD IN. + destruct (zle sz 0). + + rewrite (Mem.loadbytes_empty m b ofs sz) in LOAD by auto. + inv LOAD. contradiction. + + exploit (Mem.loadbytes_split m b ofs 1 (sz - 1) bytes). + replace (1 + (sz - 1)) with sz by omega. auto. + omega. + omega. + intros (bytes1 & bytes2 & LOAD1 & LOAD2 & CONCAT). + subst bytes. + exploit Mem.loadbytes_length. eexact LOAD1. change (nat_of_Z 1) with 1%nat. intros LENGTH1. + rewrite in_app_iff in IN. destruct IN. + * destruct bytes1; try discriminate. destruct bytes1; try discriminate. + simpl in H. destruct H; try contradiction. subst m0. + exists ofs; split. omega. auto. + * exploit (REC (sz - 1)). red; omega. eexact LOAD2. auto. + intros (ofs' & A & B). + exists ofs'; split. omega. auto. +Qed. + +Lemma smatch_loadbytes: + forall m b p b' ofs' i n ofs bytes, + Mem.loadbytes m b ofs n = Some bytes -> + smatch m b p -> + In (Pointer b' ofs' i) bytes -> + pmatch b' ofs' p. +Proof. + intros. exploit In_loadbytes; eauto. intros (ofs1 & A & B). + eapply H0; eauto. +Qed. + +Lemma loadbytes_provenance: + forall m b ofs' byte n ofs bytes, + Mem.loadbytes m b ofs n = Some bytes -> + Mem.loadbytes m b ofs' 1 = Some (byte :: nil) -> + ofs <= ofs' < ofs + n -> + In byte bytes. +Proof. + intros until n. pattern n. + apply well_founded_ind with (R := Zwf 0). +- apply Zwf_well_founded. +- intros sz REC ofs bytes LOAD LOAD1 IN. + exploit (Mem.loadbytes_split m b ofs 1 (sz - 1) bytes). + replace (1 + (sz - 1)) with sz by omega. auto. + omega. + omega. + intros (bytes1 & bytes2 & LOAD3 & LOAD4 & CONCAT). subst bytes. rewrite in_app_iff. + destruct (zeq ofs ofs'). ++ subst ofs'. rewrite LOAD1 in LOAD3; inv LOAD3. left; simpl; auto. ++ right. eapply (REC (sz - 1)). red; omega. eexact LOAD4. auto. omega. +Qed. + +Lemma storebytes_provenance: + forall m b ofs bytes m' b' ofs' b'' ofs'' i, + Mem.storebytes m b ofs bytes = Some m' -> + Mem.loadbytes m' b' ofs' 1 = Some (Pointer b'' ofs'' i :: nil) -> + In (Pointer b'' ofs'' i) bytes + \/ Mem.loadbytes m b' ofs' 1 = Some (Pointer b'' ofs'' i :: nil). +Proof. + intros. + assert (EITHER: + (b' <> b \/ ofs' + 1 <= ofs \/ ofs + Z.of_nat (length bytes) <= ofs') + \/ (b' = b /\ ofs <= ofs' < ofs + Z.of_nat (length bytes))). + { + destruct (eq_block b' b); auto. + destruct (zle (ofs' + 1) ofs); auto. + destruct (zle (ofs + Z.of_nat (length bytes)) ofs'); auto. + right. split. auto. omega. + } + destruct EITHER as [A | (A & B)]. +- right. rewrite <- H0. symmetry. eapply Mem.loadbytes_storebytes_other; eauto. omega. +- subst b'. left. + eapply loadbytes_provenance; eauto. + eapply Mem.loadbytes_storebytes_same; eauto. +Qed. + +Lemma store_provenance: + forall chunk m b ofs v m' b' ofs' b'' ofs'' i, + Mem.store chunk m b ofs v = Some m' -> + Mem.loadbytes m' b' ofs' 1 = Some (Pointer b'' ofs'' i :: nil) -> + v = Vptr b'' ofs'' /\ chunk = Mint32 + \/ Mem.loadbytes m b' ofs' 1 = Some (Pointer b'' ofs'' i :: nil). +Proof. + intros. exploit storebytes_provenance; eauto. eapply Mem.store_storebytes; eauto. + intros [A|A]; auto. left. + assert (IN_ENC_BYTES: forall bl, ~In (Pointer b'' ofs'' i) (inj_bytes bl)). + { + induction bl; simpl. tauto. red; intros; elim IHbl. destruct H1. congruence. auto. + } + assert (IN_REP_UNDEF: forall n, ~In (Pointer b'' ofs'' i) (list_repeat n Undef)). + { + intros; red; intros. exploit in_list_repeat; eauto. congruence. + } + unfold encode_val in A; destruct chunk, v; + try (eelim IN_ENC_BYTES; eassumption); + try (eelim IN_REP_UNDEF; eassumption). + simpl in A. split; auto. intuition congruence. +Qed. + +Lemma smatch_store: + forall chunk m b ofs v m' b' p av, + Mem.store chunk m b ofs v = Some m' -> + smatch m b' p -> + vmatch v av -> + smatch m' b' (vplub av p). +Proof. + intros. destruct H0 as [A B]. split. +- intros chunk' ofs' v' LOAD. destruct v'; auto with va. + exploit Mem.load_pointer_store; eauto. + intros [(P & Q & R & S & T) | DISJ]. ++ subst. apply vmatch_vplub_l. auto. ++ apply vmatch_vplub_r. apply A with (chunk := chunk') (ofs := ofs'). + rewrite <- LOAD. symmetry. eapply Mem.load_store_other; eauto. +- intros. exploit store_provenance; eauto. intros [[P Q] | P]. ++ subst. + assert (V: vmatch (Vptr b'0 ofs') (Ifptr (vplub av p))). + { + apply vmatch_vplub_l. auto. + } + inv V; auto. ++ apply pmatch_vplub. eapply B; eauto. +Qed. + +Lemma smatch_storebytes: + forall m b ofs bytes m' b' p p', + Mem.storebytes m b ofs bytes = Some m' -> + smatch m b' p -> + (forall b' ofs' i, In (Pointer b' ofs' i) bytes -> pmatch b' ofs' p') -> + smatch m' b' (plub p' p). +Proof. + intros. destruct H0 as [A B]. split. +- intros. apply vmatch_ifptr. intros bx ofsx EQ; subst v. + exploit Mem.load_loadbytes; eauto. intros (bytes' & P & Q). + exploit decode_val_pointer_inv; eauto. intros [U V]. + subst chunk bytes'. + exploit In_loadbytes; eauto. + instantiate (1 := Pointer bx ofsx 3%nat). simpl; auto. + intros (ofs' & X & Y). + exploit storebytes_provenance; eauto. intros [Z | Z]. + apply pmatch_lub_l. eauto. + apply pmatch_lub_r. eauto. +- intros. exploit storebytes_provenance; eauto. intros [Z | Z]. + apply pmatch_lub_l. eauto. + apply pmatch_lub_r. eauto. +Qed. + +Definition bmatch (m: mem) (b: block) (ab: ablock) : Prop := + smatch m b ab.(ab_summary) /\ + forall chunk ofs v, Mem.load chunk m b ofs = Some v -> vmatch v (ablock_load chunk ab ofs). + +Lemma bmatch_ext: + forall m b ab m', + bmatch m b ab -> + (forall ofs n bytes, Mem.loadbytes m' b ofs n = Some bytes -> n >= 0 -> Mem.loadbytes m b ofs n = Some bytes) -> + bmatch m' b ab. +Proof. + intros. destruct H as [A B]. split; intros. + apply smatch_ext with m; auto. + eapply B; eauto. eapply loadbytes_load_ext; eauto. +Qed. + +Lemma bmatch_inv: + forall m b ab m', + bmatch m b ab -> + (forall ofs n, n >= 0 -> Mem.loadbytes m' b ofs n = Mem.loadbytes m b ofs n) -> + bmatch m' b ab. +Proof. + intros. eapply bmatch_ext; eauto. + intros. rewrite <- H0; eauto. +Qed. + +Lemma ablock_load_sound: + forall chunk m b ofs v ab, + Mem.load chunk m b ofs = Some v -> + bmatch m b ab -> + vmatch v (ablock_load chunk ab ofs). +Proof. + intros. destruct H0. eauto. +Qed. + +Lemma ablock_load_anywhere_sound: + forall chunk m b ofs v ab, + Mem.load chunk m b ofs = Some v -> + bmatch m b ab -> + vmatch v (ablock_load_anywhere chunk ab). +Proof. + intros. destruct H0. destruct H0. unfold ablock_load_anywhere. + eapply vnormalize_cast; eauto. +Qed. + +Lemma ablock_init_sound: + forall m b p, smatch m b p -> bmatch m b (ablock_init p). +Proof. + intros; split; auto; intros. + unfold ablock_load, ablock_init; simpl. rewrite ZMap.gi. + eapply vnormalize_cast; eauto. eapply H; eauto. +Qed. + +Lemma ablock_store_anywhere_sound: + forall chunk m b ofs v m' b' ab av, + Mem.store chunk m b ofs v = Some m' -> + bmatch m b' ab -> + vmatch v av -> + bmatch m' b' (ablock_store_anywhere chunk ab av). +Proof. + intros. destruct H0 as [A B]. unfold ablock_store_anywhere. + apply ablock_init_sound. eapply smatch_store; eauto. +Qed. + +Remark inval_after_outside: + forall i lo hi c, i < lo \/ i > hi -> (inval_after lo hi c)##i = c##i. +Proof. + intros until c. functional induction (inval_after lo hi c); intros. + rewrite IHt by omega. apply ZMap.gso. unfold ZIndexed.t; omega. + auto. +Qed. + +Remark inval_after_contents: + forall chunk av i lo hi c, + (inval_after lo hi c)##i = ACval chunk av -> + c##i = ACval chunk av /\ (i < lo \/ i > hi). +Proof. + intros until c. functional induction (inval_after lo hi c); intros. + destruct (zeq i hi). + subst i. rewrite inval_after_outside in H by omega. rewrite ZMap.gss in H. discriminate. + exploit IHt; eauto. intros [A B]. rewrite ZMap.gso in A by auto. split. auto. omega. + split. auto. omega. +Qed. + +Remark inval_before_outside: + forall i hi lo c, i < lo \/ i >= hi -> (inval_before hi lo c)##i = c##i. +Proof. + intros until c. functional induction (inval_before hi lo c); intros. + rewrite IHt by omega. unfold inval_if. destruct (c##lo); auto. + destruct (zle (lo + size_chunk chunk) hi); auto. + apply ZMap.gso. unfold ZIndexed.t; omega. + auto. +Qed. + +Remark inval_before_contents_1: + forall i chunk av lo hi c, + lo <= i < hi -> (inval_before hi lo c)##i = ACval chunk av -> + c##i = ACval chunk av /\ i + size_chunk chunk <= hi. +Proof. + intros until c. functional induction (inval_before hi lo c); intros. +- destruct (zeq lo i). ++ subst i. rewrite inval_before_outside in H0 by omega. + unfold inval_if in H0. destruct (c##lo) eqn:C. congruence. + destruct (zle (lo + size_chunk chunk0) hi). + rewrite C in H0; inv H0. auto. + rewrite ZMap.gss in H0. congruence. ++ exploit IHt. omega. auto. intros [A B]; split; auto. + unfold inval_if in A. destruct (c##lo) eqn:C. auto. + destruct (zle (lo + size_chunk chunk0) hi); auto. + rewrite ZMap.gso in A; auto. +- omegaContradiction. +Qed. + +Lemma max_size_chunk: forall chunk, size_chunk chunk <= 8. +Proof. + destruct chunk; simpl; omega. +Qed. + +Remark inval_before_contents: + forall i c chunk' av' j, + (inval_before i (i - 7) c)##j = ACval chunk' av' -> + c##j = ACval chunk' av' /\ (j + size_chunk chunk' <= i \/ i <= j). +Proof. + intros. destruct (zlt j (i - 7)). + rewrite inval_before_outside in H by omega. + split. auto. left. generalize (max_size_chunk chunk'); omega. + destruct (zlt j i). + exploit inval_before_contents_1; eauto. omega. tauto. + rewrite inval_before_outside in H by omega. + split. auto. omega. +Qed. + +Lemma ablock_store_contents: + forall chunk ab i av j chunk' av', + (ablock_store chunk ab i av).(ab_contents)##j = ACval chunk' av' -> + (i = j /\ chunk' = chunk /\ av' = av) + \/ (ab.(ab_contents)##j = ACval chunk' av' + /\ (j + size_chunk chunk' <= i \/ i + size_chunk chunk <= j)). +Proof. + unfold ablock_store; simpl; intros. + destruct (zeq i j). + subst j. rewrite ZMap.gss in H. inv H; auto. + right. rewrite ZMap.gso in H by auto. + exploit inval_before_contents; eauto. intros [A B]. + exploit inval_after_contents; eauto. intros [C D]. + split. auto. omega. +Qed. + +Lemma chunk_compat_true: + forall c c', + chunk_compat c c' = true -> + size_chunk c = size_chunk c' /\ align_chunk c <= align_chunk c' /\ type_of_chunk c = type_of_chunk c'. +Proof. + destruct c, c'; intros; try discriminate; simpl; auto with va. +Qed. + +Lemma ablock_store_sound: + forall chunk m b ofs v m' ab av, + Mem.store chunk m b ofs v = Some m' -> + bmatch m b ab -> + vmatch v av -> + bmatch m' b (ablock_store chunk ab ofs av). +Proof. + intros until av; intros STORE BIN VIN. destruct BIN as [BIN1 BIN2]. split. + eapply smatch_store; eauto. + intros chunk' ofs' v' LOAD. + assert (SUMMARY: vmatch v' (vnormalize chunk' (Ifptr (vplub av ab.(ab_summary))))). + { exploit smatch_store; eauto. intros [A B]. eapply vnormalize_cast; eauto. } + unfold ablock_load. + destruct ((ab_contents (ablock_store chunk ab ofs av)) ## ofs') as [ | chunk1 av1] eqn:C. + apply SUMMARY. + destruct (chunk_compat chunk' chunk1) eqn:COMPAT; auto. + exploit chunk_compat_true; eauto. intros (U & V & W). + exploit ablock_store_contents; eauto. intros [(P & Q & R) | (P & Q)]. +- (* same offset and compatible chunks *) + subst. + assert (v' = Val.load_result chunk' v). + { exploit Mem.load_store_similar_2; eauto. congruence. } + subst v'. apply vnormalize_sound; auto. +- (* disjoint load/store *) + assert (Mem.load chunk' m b ofs' = Some v'). + { rewrite <- LOAD. symmetry. eapply Mem.load_store_other; eauto. + rewrite U. auto. } + exploit BIN2; eauto. unfold ablock_load. rewrite P. rewrite COMPAT. auto. +Qed. + +Lemma ablock_loadbytes_sound: + forall m b ab b' ofs' i n ofs bytes, + Mem.loadbytes m b ofs n = Some bytes -> + bmatch m b ab -> + In (Pointer b' ofs' i) bytes -> + pmatch b' ofs' (ablock_loadbytes ab). +Proof. + intros. destruct H0. eapply smatch_loadbytes; eauto. +Qed. + +Lemma ablock_storebytes_anywhere_sound: + forall m b ofs bytes p m' b' ab, + Mem.storebytes m b ofs bytes = Some m' -> + (forall b' ofs' i, In (Pointer b' ofs' i) bytes -> pmatch b' ofs' p) -> + bmatch m b' ab -> + bmatch m' b' (ablock_storebytes_anywhere ab p). +Proof. + intros. destruct H1 as [A B]. apply ablock_init_sound. + eapply smatch_storebytes; eauto. +Qed. + +Lemma ablock_storebytes_contents: + forall ab p i sz j chunk' av', + (ablock_storebytes ab p i sz).(ab_contents)##j = ACval chunk' av' -> + ab.(ab_contents)##j = ACval chunk' av' + /\ (j + size_chunk chunk' <= i \/ i + Zmax sz 0 <= j). +Proof. + unfold ablock_storebytes; simpl; intros. + exploit inval_before_contents; eauto. clear H. intros [A B]. + exploit inval_after_contents; eauto. clear A. intros [C D]. + split. auto. xomega. +Qed. + +Lemma ablock_storebytes_sound: + forall m b ofs bytes m' p ab sz, + Mem.storebytes m b ofs bytes = Some m' -> + length bytes = nat_of_Z sz -> + (forall b' ofs' i, In (Pointer b' ofs' i) bytes -> pmatch b' ofs' p) -> + bmatch m b ab -> + bmatch m' b (ablock_storebytes ab p ofs sz). +Proof. + intros until sz; intros STORE LENGTH CONTENTS BM. destruct BM as [BM1 BM2]. split. + eapply smatch_storebytes; eauto. + intros chunk' ofs' v' LOAD'. + assert (SUMMARY: vmatch v' (vnormalize chunk' (Ifptr (plub p ab.(ab_summary))))). + { exploit smatch_storebytes; eauto. intros [A B]. eapply vnormalize_cast; eauto. } + unfold ablock_load. + destruct (ab_contents (ablock_storebytes ab p ofs sz))##ofs' eqn:C. + exact SUMMARY. + destruct (chunk_compat chunk' chunk) eqn:COMPAT; auto. + exploit chunk_compat_true; eauto. intros (U & V & W). + exploit ablock_storebytes_contents; eauto. intros [A B]. + assert (Mem.load chunk' m b ofs' = Some v'). + { rewrite <- LOAD'; symmetry. eapply Mem.load_storebytes_other; eauto. + rewrite U. rewrite LENGTH. rewrite nat_of_Z_max. right; omega. } + exploit BM2; eauto. unfold ablock_load. rewrite A. rewrite COMPAT. auto. +Qed. + +(** Boolean equality *) + +Definition bbeq (ab1 ab2: ablock) : bool := + eq_aptr ab1.(ab_summary) ab2.(ab_summary) && + PTree.beq (fun c1 c2 => proj_sumbool (eq_acontent c1 c2)) + (snd ab1.(ab_contents)) (snd ab2.(ab_contents)). + +Lemma bbeq_load: + forall ab1 ab2, + bbeq ab1 ab2 = true -> + ab1.(ab_summary) = ab2.(ab_summary) + /\ (forall chunk i, ablock_load chunk ab1 i = ablock_load chunk ab2 i). +Proof. + unfold bbeq; intros. InvBooleans. split. +- unfold ablock_load_anywhere; intros; congruence. +- rewrite PTree.beq_correct in H1. + assert (A: forall i, ZMap.get i (ab_contents ab1) = ZMap.get i (ab_contents ab2)). + { + intros. unfold ZMap.get, PMap.get. set (j := ZIndexed.index i). + specialize (H1 j). + destruct (snd (ab_contents ab1))!j; destruct (snd (ab_contents ab2))!j; try contradiction. + InvBooleans; auto. + rewrite ! ab_default. auto. + } + intros. unfold ablock_load. rewrite A, H. + destruct (ab_contents ab2)##i; auto. +Qed. + +Lemma bbeq_sound: + forall ab1 ab2, + bbeq ab1 ab2 = true -> + forall m b, bmatch m b ab1 <-> bmatch m b ab2. +Proof. + intros. exploit bbeq_load; eauto. intros [A B]. + unfold bmatch. rewrite A. intuition. rewrite <- B; eauto. rewrite B; eauto. +Qed. + +(** Least upper bound *) + +Definition combine_acontents_opt (c1 c2: option acontent) : option acontent := + match c1, c2 with + | Some (ACval chunk1 v1), Some (ACval chunk2 v2) => + if chunk_eq chunk1 chunk2 then Some(ACval chunk1 (vlub v1 v2)) else None + | _, _ => + None + end. + +Definition combine_contentmaps (m1 m2: ZMap.t acontent) : ZMap.t acontent := + (ACany, PTree.combine combine_acontents_opt (snd m1) (snd m2)). + +Definition blub (ab1 ab2: ablock) : ablock := + {| ab_contents := combine_contentmaps ab1.(ab_contents) ab2.(ab_contents); + ab_summary := plub ab1.(ab_summary) ab2.(ab_summary); + ab_default := refl_equal _ |}. + +Definition combine_acontents (c1 c2: acontent) : acontent := + match c1, c2 with + | ACval chunk1 v1, ACval chunk2 v2 => + if chunk_eq chunk1 chunk2 then ACval chunk1 (vlub v1 v2) else ACany + | _, _ => ACany + end. + +Lemma get_combine_contentmaps: + forall m1 m2 i, + fst m1 = ACany -> fst m2 = ACany -> + ZMap.get i (combine_contentmaps m1 m2) = combine_acontents (ZMap.get i m1) (ZMap.get i m2). +Proof. + intros. destruct m1 as [dfl1 pt1]. destruct m2 as [dfl2 pt2]; simpl in *. + subst dfl1 dfl2. unfold combine_contentmaps, ZMap.get, PMap.get, fst, snd. + set (j := ZIndexed.index i). + rewrite PTree.gcombine by auto. + destruct (pt1!j) as [[]|]; destruct (pt2!j) as [[]|]; simpl; auto. + destruct (chunk_eq chunk chunk0); auto. +Qed. + +Lemma smatch_lub_l: + forall m b p q, smatch m b p -> smatch m b (plub p q). +Proof. + intros. destruct H as [A B]. split; intros. + change (vmatch v (vlub (Ifptr p) (Ifptr q))). apply vmatch_lub_l. eapply A; eauto. + apply pmatch_lub_l. eapply B; eauto. +Qed. + +Lemma smatch_lub_r: + forall m b p q, smatch m b q -> smatch m b (plub p q). +Proof. + intros. destruct H as [A B]. split; intros. + change (vmatch v (vlub (Ifptr p) (Ifptr q))). apply vmatch_lub_r. eapply A; eauto. + apply pmatch_lub_r. eapply B; eauto. +Qed. + +Lemma bmatch_lub_l: + forall m b x y, bmatch m b x -> bmatch m b (blub x y). +Proof. + intros. destruct H as [BM1 BM2]. split; unfold blub; simpl. +- apply smatch_lub_l; auto. +- intros. + assert (SUMMARY: vmatch v (vnormalize chunk (Ifptr (plub (ab_summary x) (ab_summary y)))) +). + { exploit smatch_lub_l; eauto. instantiate (1 := ab_summary y). + intros [SUMM _]. eapply vnormalize_cast; eauto. } + exploit BM2; eauto. + unfold ablock_load; simpl. rewrite get_combine_contentmaps by (apply ab_default). + unfold combine_acontents; destruct (ab_contents x)##ofs, (ab_contents y)##ofs; auto. + destruct (chunk_eq chunk0 chunk1); auto. subst chunk0. + destruct (chunk_compat chunk chunk1); auto. + intros. eapply vmatch_ge; eauto. apply vnormalize_monotone. apply vge_lub_l. +Qed. + +Lemma bmatch_lub_r: + forall m b x y, bmatch m b y -> bmatch m b (blub x y). +Proof. + intros. destruct H as [BM1 BM2]. split; unfold blub; simpl. +- apply smatch_lub_r; auto. +- intros. + assert (SUMMARY: vmatch v (vnormalize chunk (Ifptr (plub (ab_summary x) (ab_summary y)))) +). + { exploit smatch_lub_r; eauto. instantiate (1 := ab_summary x). + intros [SUMM _]. eapply vnormalize_cast; eauto. } + exploit BM2; eauto. + unfold ablock_load; simpl. rewrite get_combine_contentmaps by (apply ab_default). + unfold combine_acontents; destruct (ab_contents x)##ofs, (ab_contents y)##ofs; auto. + destruct (chunk_eq chunk0 chunk1); auto. subst chunk0. + destruct (chunk_compat chunk chunk1); auto. + intros. eapply vmatch_ge; eauto. apply vnormalize_monotone. apply vge_lub_r. +Qed. + +(** * Abstracting read-only global variables *) + +Definition romem := PTree.t ablock. + +Definition romatch (m: mem) (rm: romem) : Prop := + forall b id ab, + bc b = BCglob id -> + rm!id = Some ab -> + pge Glob ab.(ab_summary) + /\ bmatch m b ab + /\ forall ofs, ~Mem.perm m b ofs Max Writable. + +Lemma romatch_store: + forall chunk m b ofs v m' rm, + Mem.store chunk m b ofs v = Some m' -> + romatch m rm -> + romatch m' rm. +Proof. + intros; red; intros. exploit H0; eauto. intros (A & B & C). split; auto. split. +- exploit Mem.store_valid_access_3; eauto. intros [P _]. + apply bmatch_inv with m; auto. ++ intros. eapply Mem.loadbytes_store_other; eauto. + left. red; intros; subst b0. elim (C ofs). apply Mem.perm_cur_max. + apply P. generalize (size_chunk_pos chunk); omega. +- intros; red; intros; elim (C ofs0). eauto with mem. +Qed. + +Lemma romatch_storebytes: + forall m b ofs bytes m' rm, + Mem.storebytes m b ofs bytes = Some m' -> + bytes <> nil -> + romatch m rm -> + romatch m' rm. +Proof. + intros; red; intros. exploit H1; eauto. intros (A & B & C). split; auto. split. +- apply bmatch_inv with m; auto. + intros. eapply Mem.loadbytes_storebytes_other; eauto. + left. red; intros; subst b0. elim (C ofs). apply Mem.perm_cur_max. + eapply Mem.storebytes_range_perm; eauto. + destruct bytes. congruence. simpl length. rewrite inj_S. omega. +- intros; red; intros; elim (C ofs0). eauto with mem. +Qed. + +Lemma romatch_ext: + forall m rm m', + romatch m rm -> + (forall b id ofs n bytes, bc b = BCglob id -> Mem.loadbytes m' b ofs n = Some bytes -> Mem.loadbytes m b ofs n = Some bytes) -> + (forall b id ofs p, bc b = BCglob id -> Mem.perm m' b ofs Max p -> Mem.perm m b ofs Max p) -> + romatch m' rm. +Proof. + intros; red; intros. exploit H; eauto. intros (A & B & C). + split. auto. + split. apply bmatch_ext with m; auto. intros. eapply H0; eauto. + intros; red; intros. elim (C ofs). eapply H1; eauto. +Qed. + +Lemma romatch_free: + forall m b lo hi m' rm, + Mem.free m b lo hi = Some m' -> + romatch m rm -> + romatch m' rm. +Proof. + intros. apply romatch_ext with m; auto. + intros. eapply Mem.loadbytes_free_2; eauto. + intros. eauto with mem. +Qed. + +Lemma romatch_alloc: + forall m b lo hi m' rm, + Mem.alloc m lo hi = (m', b) -> + bc_below bc (Mem.nextblock m) -> + romatch m rm -> + romatch m' rm. +Proof. + intros. apply romatch_ext with m; auto. + intros. rewrite <- H3; symmetry. eapply Mem.loadbytes_alloc_unchanged; eauto. + apply H0. congruence. + intros. eapply Mem.perm_alloc_4; eauto. apply Mem.valid_not_valid_diff with m; eauto with mem. + apply H0. congruence. +Qed. + +(** * Abstracting memory states *) + +Record amem : Type := AMem { + am_stack: ablock; + am_glob: PTree.t ablock; + am_nonstack: aptr; + am_top: aptr +}. + +Record mmatch (m: mem) (am: amem) : Prop := mk_mem_match { + mmatch_stack: forall b, + bc b = BCstack -> + bmatch m b am.(am_stack); + mmatch_glob: forall id ab b, + bc b = BCglob id -> + am.(am_glob)!id = Some ab -> + bmatch m b ab; + mmatch_nonstack: forall b, + bc b <> BCstack -> bc b <> BCinvalid -> + smatch m b am.(am_nonstack); + mmatch_top: forall b, + bc b <> BCinvalid -> + smatch m b am.(am_top); + mmatch_below: + bc_below bc (Mem.nextblock m) +}. + +Definition minit (p: aptr) := + {| am_stack := ablock_init p; + am_glob := PTree.empty _; + am_nonstack := p; + am_top := p |}. + +Definition mbot := minit Pbot. +Definition mtop := minit Ptop. + +Definition load (chunk: memory_chunk) (rm: romem) (m: amem) (p: aptr) : aval := + match p with + | Pbot => if strict then Vbot else Vtop + | Gl id ofs => + match rm!id with + | Some ab => ablock_load chunk ab (Int.unsigned ofs) + | None => + match m.(am_glob)!id with + | Some ab => ablock_load chunk ab (Int.unsigned ofs) + | None => vnormalize chunk (Ifptr m.(am_nonstack)) + end + end + | Glo id => + match rm!id with + | Some ab => ablock_load_anywhere chunk ab + | None => + match m.(am_glob)!id with + | Some ab => ablock_load_anywhere chunk ab + | None => vnormalize chunk (Ifptr m.(am_nonstack)) + end + end + | Stk ofs => ablock_load chunk m.(am_stack) (Int.unsigned ofs) + | Stack => ablock_load_anywhere chunk m.(am_stack) + | Glob | Nonstack => vnormalize chunk (Ifptr m.(am_nonstack)) + | Ptop => vnormalize chunk (Ifptr m.(am_top)) + end. + +Definition loadv (chunk: memory_chunk) (rm: romem) (m: amem) (addr: aval) : aval := + load chunk rm m (aptr_of_aval addr). + +Definition store (chunk: memory_chunk) (m: amem) (p: aptr) (av: aval) : amem := + {| am_stack := + match p with + | Stk ofs => ablock_store chunk m.(am_stack) (Int.unsigned ofs) av + | Stack | Ptop => ablock_store_anywhere chunk m.(am_stack) av + | _ => m.(am_stack) + end; + am_glob := + match p with + | Gl id ofs => + let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + PTree.set id (ablock_store chunk ab (Int.unsigned ofs) av) m.(am_glob) + | Glo id => + let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + PTree.set id (ablock_store_anywhere chunk ab av) m.(am_glob) + | Glob | Nonstack | Ptop => PTree.empty _ + | _ => m.(am_glob) + end; + am_nonstack := + match p with + | Gl _ _ | Glo _ | Glob | Nonstack | Ptop => vplub av m.(am_nonstack) + | _ => m.(am_nonstack) + end; + am_top := vplub av m.(am_top) + |}. + +Definition storev (chunk: memory_chunk) (m: amem) (addr: aval) (v: aval): amem := + store chunk m (aptr_of_aval addr) v. + +Definition loadbytes (m: amem) (rm: romem) (p: aptr) : aptr := + match p with + | Pbot => if strict then Pbot else Ptop + | Gl id _ | Glo id => + match rm!id with + | Some ab => ablock_loadbytes ab + | None => + match m.(am_glob)!id with + | Some ab => ablock_loadbytes ab + | None => m.(am_nonstack) + end + end + | Stk _ | Stack => ablock_loadbytes m.(am_stack) + | Glob | Nonstack => m.(am_nonstack) + | Ptop => m.(am_top) + end. + +Definition storebytes (m: amem) (dst: aptr) (sz: Z) (p: aptr) : amem := + {| am_stack := + match dst with + | Stk ofs => ablock_storebytes m.(am_stack) p (Int.unsigned ofs) sz + | Stack | Ptop => ablock_storebytes_anywhere m.(am_stack) p + | _ => m.(am_stack) + end; + am_glob := + match dst with + | Gl id ofs => + let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + PTree.set id (ablock_storebytes ab p (Int.unsigned ofs) sz) m.(am_glob) + | Glo id => + let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + PTree.set id (ablock_storebytes_anywhere ab p) m.(am_glob) + | Glob | Nonstack | Ptop => PTree.empty _ + | _ => m.(am_glob) + end; + am_nonstack := + match dst with + | Gl _ _ | Glo _ | Glob | Nonstack | Ptop => plub p m.(am_nonstack) + | _ => m.(am_nonstack) + end; + am_top := plub p m.(am_top) + |}. + +Theorem load_sound: + forall chunk m b ofs v rm am p, + Mem.load chunk m b (Int.unsigned ofs) = Some v -> + romatch m rm -> + mmatch m am -> + pmatch b ofs p -> + vmatch v (load chunk rm am p). +Proof. + intros. unfold load. inv H2. +- (* Gl id ofs *) + destruct (rm!id) as [ab|] eqn:RM. + eapply ablock_load_sound; eauto. eapply H0; eauto. + destruct (am_glob am)!id as [ab|] eqn:AM. + eapply ablock_load_sound; eauto. eapply mmatch_glob; eauto. + eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto; congruence. +- (* Glo id *) + destruct (rm!id) as [ab|] eqn:RM. + eapply ablock_load_anywhere_sound; eauto. eapply H0; eauto. + destruct (am_glob am)!id as [ab|] eqn:AM. + eapply ablock_load_anywhere_sound; eauto. eapply mmatch_glob; eauto. + eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto; congruence. +- (* Glob *) + eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto. congruence. congruence. +- (* Stk ofs *) + eapply ablock_load_sound; eauto. eapply mmatch_stack; eauto. +- (* Stack *) + eapply ablock_load_anywhere_sound; eauto. eapply mmatch_stack; eauto. +- (* Nonstack *) + eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto. +- (* Top *) + eapply vnormalize_cast; eauto. eapply mmatch_top; eauto. +Qed. + +Theorem loadv_sound: + forall chunk m addr v rm am aaddr, + Mem.loadv chunk m addr = Some v -> + romatch m rm -> + mmatch m am -> + vmatch addr aaddr -> + vmatch v (loadv chunk rm am aaddr). +Proof. + intros. destruct addr; simpl in H; try discriminate. + eapply load_sound; eauto. apply match_aptr_of_aval; auto. +Qed. + +Theorem store_sound: + forall chunk m b ofs v m' am p av, + Mem.store chunk m b (Int.unsigned ofs) v = Some m' -> + mmatch m am -> + pmatch b ofs p -> + vmatch v av -> + mmatch m' (store chunk am p av). +Proof. + intros until av; intros STORE MM PM VM. + unfold store; constructor; simpl; intros. +- (* Stack *) + assert (DFL: bc b <> BCstack -> bmatch m' b0 (am_stack am)). + { intros. apply bmatch_inv with m. eapply mmatch_stack; eauto. + intros. eapply Mem.loadbytes_store_other; eauto. left; congruence. } + inv PM; try (apply DFL; congruence). + + assert (b0 = b) by (eapply bc_stack; eauto). subst b0. + eapply ablock_store_sound; eauto. eapply mmatch_stack; eauto. + + assert (b0 = b) by (eapply bc_stack; eauto). subst b0. + eapply ablock_store_anywhere_sound; eauto. eapply mmatch_stack; eauto. + + eapply ablock_store_anywhere_sound; eauto. eapply mmatch_stack; eauto. + +- (* Globals *) + rename b0 into b'. + assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab -> + bmatch m' b' ab). + { intros. apply bmatch_inv with m. eapply mmatch_glob; eauto. + intros. eapply Mem.loadbytes_store_other; eauto. left; congruence. } + inv PM. + + rewrite PTree.gsspec in H0. destruct (peq id id0). + subst id0; inv H0. + assert (b' = b) by (eapply bc_glob; eauto). subst b'. + eapply ablock_store_sound; eauto. + destruct (am_glob am)!id as [ab0|] eqn:GL. + eapply mmatch_glob; eauto. + apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. + eapply DFL; eauto. congruence. + + rewrite PTree.gsspec in H0. destruct (peq id id0). + subst id0; inv H0. + assert (b' = b) by (eapply bc_glob; eauto). subst b'. + eapply ablock_store_anywhere_sound; eauto. + destruct (am_glob am)!id as [ab0|] eqn:GL. + eapply mmatch_glob; eauto. + apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. + eapply DFL; eauto. congruence. + + rewrite PTree.gempty in H0; congruence. + + eapply DFL; eauto. congruence. + + eapply DFL; eauto. congruence. + + rewrite PTree.gempty in H0; congruence. + + rewrite PTree.gempty in H0; congruence. + +- (* Nonstack *) + assert (DFL: smatch m' b0 (vplub av (am_nonstack am))). + { eapply smatch_store; eauto. eapply mmatch_nonstack; eauto. } + assert (STK: bc b = BCstack -> smatch m' b0 (am_nonstack am)). + { intros. apply smatch_inv with m. eapply mmatch_nonstack; eauto; congruence. + intros. eapply Mem.loadbytes_store_other; eauto. left. congruence. + } + inv PM; (apply DFL || apply STK; congruence). + +- (* Top *) + eapply smatch_store; eauto. eapply mmatch_top; eauto. + +- (* Below *) + erewrite Mem.nextblock_store by eauto. eapply mmatch_below; eauto. +Qed. + +Theorem storev_sound: + forall chunk m addr v m' am aaddr av, + Mem.storev chunk m addr v = Some m' -> + mmatch m am -> + vmatch addr aaddr -> + vmatch v av -> + mmatch m' (storev chunk am aaddr av). +Proof. + intros. destruct addr; simpl in H; try discriminate. + eapply store_sound; eauto. apply match_aptr_of_aval; auto. +Qed. + +Theorem loadbytes_sound: + forall m b ofs sz bytes am rm p, + Mem.loadbytes m b (Int.unsigned ofs) sz = Some bytes -> + romatch m rm -> + mmatch m am -> + pmatch b ofs p -> + forall b' ofs' i, In (Pointer b' ofs' i) bytes -> pmatch b' ofs' (loadbytes am rm p). +Proof. + intros. unfold loadbytes; inv H2. +- (* Gl id ofs *) + destruct (rm!id) as [ab|] eqn:RM. + exploit H0; eauto. intros (A & B & C). eapply ablock_loadbytes_sound; eauto. + destruct (am_glob am)!id as [ab|] eqn:GL. + eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto. + eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va. +- (* Glo id *) + destruct (rm!id) as [ab|] eqn:RM. + exploit H0; eauto. intros (A & B & C). eapply ablock_loadbytes_sound; eauto. + destruct (am_glob am)!id as [ab|] eqn:GL. + eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto. + eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va. +- (* Glob *) + eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va. +- (* Stk ofs *) + eapply ablock_loadbytes_sound; eauto. eapply mmatch_stack; eauto. +- (* Stack *) + eapply ablock_loadbytes_sound; eauto. eapply mmatch_stack; eauto. +- (* Nonstack *) + eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va. +- (* Top *) + eapply smatch_loadbytes; eauto. eapply mmatch_top; eauto with va. +Qed. + +Theorem storebytes_sound: + forall m b ofs bytes m' am p sz q, + Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' -> + mmatch m am -> + pmatch b ofs p -> + length bytes = nat_of_Z sz -> + (forall b' ofs' i, In (Pointer b' ofs' i) bytes -> pmatch b' ofs' q) -> + mmatch m' (storebytes am p sz q). +Proof. + intros until q; intros STORE MM PM LENGTH BYTES. + unfold storebytes; constructor; simpl; intros. +- (* Stack *) + assert (DFL: bc b <> BCstack -> bmatch m' b0 (am_stack am)). + { intros. apply bmatch_inv with m. eapply mmatch_stack; eauto. + intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence. } + inv PM; try (apply DFL; congruence). + + assert (b0 = b) by (eapply bc_stack; eauto). subst b0. + eapply ablock_storebytes_sound; eauto. eapply mmatch_stack; eauto. + + assert (b0 = b) by (eapply bc_stack; eauto). subst b0. + eapply ablock_storebytes_anywhere_sound; eauto. eapply mmatch_stack; eauto. + + eapply ablock_storebytes_anywhere_sound; eauto. eapply mmatch_stack; eauto. + +- (* Globals *) + rename b0 into b'. + assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab -> + bmatch m' b' ab). + { intros. apply bmatch_inv with m. eapply mmatch_glob; eauto. + intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence. } + inv PM. + + rewrite PTree.gsspec in H0. destruct (peq id id0). + subst id0; inv H0. + assert (b' = b) by (eapply bc_glob; eauto). subst b'. + eapply ablock_storebytes_sound; eauto. + destruct (am_glob am)!id as [ab0|] eqn:GL. + eapply mmatch_glob; eauto. + apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. + eapply DFL; eauto. congruence. + + rewrite PTree.gsspec in H0. destruct (peq id id0). + subst id0; inv H0. + assert (b' = b) by (eapply bc_glob; eauto). subst b'. + eapply ablock_storebytes_anywhere_sound; eauto. + destruct (am_glob am)!id as [ab0|] eqn:GL. + eapply mmatch_glob; eauto. + apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. + eapply DFL; eauto. congruence. + + rewrite PTree.gempty in H0; congruence. + + eapply DFL; eauto. congruence. + + eapply DFL; eauto. congruence. + + rewrite PTree.gempty in H0; congruence. + + rewrite PTree.gempty in H0; congruence. + +- (* Nonstack *) + assert (DFL: smatch m' b0 (plub q (am_nonstack am))). + { eapply smatch_storebytes; eauto. eapply mmatch_nonstack; eauto. } + assert (STK: bc b = BCstack -> smatch m' b0 (am_nonstack am)). + { intros. apply smatch_inv with m. eapply mmatch_nonstack; eauto; congruence. + intros. eapply Mem.loadbytes_storebytes_other; eauto. left. congruence. + } + inv PM; (apply DFL || apply STK; congruence). + +- (* Top *) + eapply smatch_storebytes; eauto. eapply mmatch_top; eauto. + +- (* Below *) + erewrite Mem.nextblock_storebytes by eauto. eapply mmatch_below; eauto. +Qed. + +Lemma mmatch_ext: + forall m am m', + mmatch m am -> + (forall b ofs n bytes, bc b <> BCinvalid -> n >= 0 -> Mem.loadbytes m' b ofs n = Some bytes -> Mem.loadbytes m b ofs n = Some bytes) -> + Ple (Mem.nextblock m) (Mem.nextblock m') -> + mmatch m' am. +Proof. + intros. inv H. constructor; intros. +- apply bmatch_ext with m; auto with va. +- apply bmatch_ext with m; eauto with va. +- apply smatch_ext with m; auto with va. +- apply smatch_ext with m; auto with va. +- red; intros. exploit mmatch_below0; eauto. xomega. +Qed. + +Lemma mmatch_free: + forall m b lo hi m' am, + Mem.free m b lo hi = Some m' -> + mmatch m am -> + mmatch m' am. +Proof. + intros. apply mmatch_ext with m; auto. + intros. eapply Mem.loadbytes_free_2; eauto. + erewrite <- Mem.nextblock_free by eauto. xomega. +Qed. + +Lemma mmatch_top': + forall m am, mmatch m am -> mmatch m mtop. +Proof. + intros. constructor; simpl; intros. +- apply ablock_init_sound. apply smatch_ge with (ab_summary (am_stack am)). + eapply mmatch_stack; eauto. constructor. +- rewrite PTree.gempty in H1; discriminate. +- eapply smatch_ge. eapply mmatch_nonstack; eauto. constructor. +- eapply smatch_ge. eapply mmatch_top; eauto. constructor. +- eapply mmatch_below; eauto. +Qed. + +(** Boolean equality *) + +Definition mbeq (m1 m2: amem) : bool := + eq_aptr m1.(am_top) m2.(am_top) + && eq_aptr m1.(am_nonstack) m2.(am_nonstack) + && bbeq m1.(am_stack) m2.(am_stack) + && PTree.beq bbeq m1.(am_glob) m2.(am_glob). + +Lemma mbeq_sound: + forall m1 m2, mbeq m1 m2 = true -> forall m, mmatch m m1 <-> mmatch m m2. +Proof. + unfold mbeq; intros. InvBooleans. rewrite PTree.beq_correct in H1. + split; intros M; inv M; constructor; intros. +- erewrite <- bbeq_sound; eauto. +- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m1)!id eqn:G; try contradiction. + erewrite <- bbeq_sound; eauto. +- rewrite <- H; eauto. +- rewrite <- H0; eauto. +- auto. +- erewrite bbeq_sound; eauto. +- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m2)!id eqn:G; try contradiction. + erewrite bbeq_sound; eauto. +- rewrite H; eauto. +- rewrite H0; eauto. +- auto. +Qed. + +(** Least upper bound *) + +Definition combine_ablock (ob1 ob2: option ablock) : option ablock := + match ob1, ob2 with + | Some b1, Some b2 => Some (blub b1 b2) + | _, _ => None + end. + +Definition mlub (m1 m2: amem) : amem := +{| am_stack := blub m1.(am_stack) m2.(am_stack); + am_glob := PTree.combine combine_ablock m1.(am_glob) m2.(am_glob); + am_nonstack := plub m1.(am_nonstack) m2.(am_nonstack); + am_top := plub m1.(am_top) m2.(am_top) |}. + +Lemma mmatch_lub_l: + forall m x y, mmatch m x -> mmatch m (mlub x y). +Proof. + intros. inv H. constructor; simpl; intros. +- apply bmatch_lub_l; auto. +- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0. + destruct (am_glob x)!id as [b1|] eqn:G1; + destruct (am_glob y)!id as [b2|] eqn:G2; + inv H0. + apply bmatch_lub_l; eauto. +- apply smatch_lub_l; auto. +- apply smatch_lub_l; auto. +- auto. +Qed. + +Lemma mmatch_lub_r: + forall m x y, mmatch m y -> mmatch m (mlub x y). +Proof. + intros. inv H. constructor; simpl; intros. +- apply bmatch_lub_r; auto. +- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0. + destruct (am_glob x)!id as [b1|] eqn:G1; + destruct (am_glob y)!id as [b2|] eqn:G2; + inv H0. + apply bmatch_lub_r; eauto. +- apply smatch_lub_r; auto. +- apply smatch_lub_r; auto. +- auto. +Qed. + +End MATCH. + +(** * Monotonicity properties when the block classification changes. *) + +Lemma genv_match_exten: + forall ge (bc1 bc2: block_classification), + genv_match bc1 ge -> + (forall b id, bc1 b = BCglob id <-> bc2 b = BCglob id) -> + (forall b, bc1 b = BCother -> bc2 b = BCother) -> + genv_match bc2 ge. +Proof. + intros. destruct H as [A B]. split; intros. +- rewrite <- H0. eauto. +- exploit B; eauto. destruct (bc1 b) eqn:BC1. + + intuition congruence. + + rewrite H0 in BC1. intuition congruence. + + intuition congruence. + + erewrite H1 by eauto. intuition congruence. +Qed. + +Lemma romatch_exten: + forall (bc1 bc2: block_classification) m rm, + romatch bc1 m rm -> + (forall b id, bc2 b = BCglob id <-> bc1 b = BCglob id) -> + romatch bc2 m rm. +Proof. + intros; red; intros. rewrite H0 in H1. exploit H; eauto. intros (A & B & C). + split; auto. split; auto. + assert (PM: forall b ofs p, pmatch bc1 b ofs p -> pmatch bc1 b ofs (ab_summary ab) -> pmatch bc2 b ofs p). + { + intros. + assert (pmatch bc1 b0 ofs Glob) by (eapply pmatch_ge; eauto). + inv H5. + assert (bc2 b0 = BCglob id0) by (rewrite H0; auto). + inv H3; econstructor; eauto with va. + } + assert (VM: forall v x, vmatch bc1 v x -> vmatch bc1 v (Ifptr (ab_summary ab)) -> vmatch bc2 v x). + { + intros. inv H3; constructor; auto; inv H4; eapply PM; eauto. + } + destruct B as [[B1 B2] B3]. split. split. +- intros. apply VM; eauto. +- intros. apply PM; eauto. +- intros. apply VM; eauto. +Qed. + +Definition bc_incr (bc1 bc2: block_classification) : Prop := + forall b, bc1 b <> BCinvalid -> bc2 b = bc1 b. + +Section MATCH_INCR. + +Variables bc1 bc2: block_classification. +Hypothesis INCR: bc_incr bc1 bc2. + +Lemma pmatch_incr: forall b ofs p, pmatch bc1 b ofs p -> pmatch bc2 b ofs p. +Proof. + induction 1; + assert (bc2 b = bc1 b) by (apply INCR; congruence); + econstructor; eauto with va. rewrite H0; eauto. +Qed. + +Lemma vmatch_incr: forall v x, vmatch bc1 v x -> vmatch bc2 v x. +Proof. + induction 1; constructor; auto; apply pmatch_incr; auto. +Qed. + +Lemma smatch_incr: forall m b p, smatch bc1 m b p -> smatch bc2 m b p. +Proof. + intros. destruct H as [A B]. split; intros. + apply vmatch_incr; eauto. + apply pmatch_incr; eauto. +Qed. + +Lemma bmatch_incr: forall m b ab, bmatch bc1 m b ab -> bmatch bc2 m b ab. +Proof. + intros. destruct H as [B1 B2]. split. + apply smatch_incr; auto. + intros. apply vmatch_incr; eauto. +Qed. + +End MATCH_INCR. + +(** * Matching and memory injections. *) + +Definition inj_of_bc (bc: block_classification) : meminj := + fun b => match bc b with BCinvalid => None | _ => Some(b, 0) end. + +Lemma inj_of_bc_valid: + forall (bc: block_classification) b, bc b <> BCinvalid -> inj_of_bc bc b = Some(b, 0). +Proof. + intros. unfold inj_of_bc. destruct (bc b); congruence. +Qed. + +Lemma inj_of_bc_inv: + forall (bc: block_classification) b b' delta, + inj_of_bc bc b = Some(b', delta) -> bc b <> BCinvalid /\ b' = b /\ delta = 0. +Proof. + unfold inj_of_bc; intros. destruct (bc b); intuition congruence. +Qed. + +Lemma pmatch_inj: + forall bc b ofs p, pmatch bc b ofs p -> inj_of_bc bc b = Some(b, 0). +Proof. + intros. apply inj_of_bc_valid. inv H; congruence. +Qed. + +Lemma vmatch_inj: + forall bc v x, vmatch bc v x -> val_inject (inj_of_bc bc) v v. +Proof. + induction 1; econstructor. + eapply pmatch_inj; eauto. rewrite Int.add_zero; auto. + eapply pmatch_inj; eauto. rewrite Int.add_zero; auto. +Qed. + +Lemma vmatch_list_inj: + forall bc vl xl, list_forall2 (vmatch bc) vl xl -> val_list_inject (inj_of_bc bc) vl vl. +Proof. + induction 1; constructor. eapply vmatch_inj; eauto. auto. +Qed. + +Lemma mmatch_inj: + forall bc m am, mmatch bc m am -> bc_below bc (Mem.nextblock m) -> Mem.inject (inj_of_bc bc) m m. +Proof. + intros. constructor. constructor. +- (* perms *) + intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst. + rewrite Zplus_0_r. auto. +- (* alignment *) + intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst. + apply Zdivide_0. +- (* contents *) + intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst. + rewrite Zplus_0_r. + set (mv := ZMap.get ofs (Mem.mem_contents m)#b1). + assert (Mem.loadbytes m b1 ofs 1 = Some (mv :: nil)). + { + Local Transparent Mem.loadbytes. + unfold Mem.loadbytes. rewrite pred_dec_true. reflexivity. + red; intros. replace ofs0 with ofs by omega. auto. + } + destruct mv; econstructor. + apply inj_of_bc_valid. + assert (PM: pmatch bc b i Ptop). + { exploit mmatch_top; eauto. intros [P Q]. + eapply pmatch_top'. eapply Q; eauto. } + inv PM; auto. + rewrite Int.add_zero; auto. +- (* free blocks *) + intros. unfold inj_of_bc. erewrite bc_below_invalid; eauto. +- (* mapped blocks *) + intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst. + apply H0; auto. +- (* overlap *) + red; intros. + exploit inj_of_bc_inv. eexact H2. intros (A1 & B & C); subst. + exploit inj_of_bc_inv. eexact H3. intros (A2 & B & C); subst. + auto. +- (* overflow *) + intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst. + rewrite Zplus_0_r. split. omega. apply Int.unsigned_range_2. +Qed. + +Lemma inj_of_bc_preserves_globals: + forall bc ge, genv_match bc ge -> meminj_preserves_globals ge (inj_of_bc bc). +Proof. + intros. destruct H as [A B]. + split. intros. apply inj_of_bc_valid. rewrite A in H. congruence. + split. intros. apply inj_of_bc_valid. apply B. eapply Genv.genv_vars_range; eauto. + intros. exploit inj_of_bc_inv; eauto. intros (P & Q & R). auto. +Qed. + +Lemma pmatch_inj_top: + forall bc b b' delta ofs, inj_of_bc bc b = Some(b', delta) -> pmatch bc b ofs Ptop. +Proof. + intros. exploit inj_of_bc_inv; eauto. intros (A & B & C). constructor; auto. +Qed. + +Lemma vmatch_inj_top: + forall bc v v', val_inject (inj_of_bc bc) v v' -> vmatch bc v Vtop. +Proof. + intros. inv H; constructor. eapply pmatch_inj_top; eauto. +Qed. + +Lemma mmatch_inj_top: + forall bc m m', Mem.inject (inj_of_bc bc) m m' -> mmatch bc m mtop. +Proof. + intros. + assert (SM: forall b, bc b <> BCinvalid -> smatch bc m b Ptop). + { + intros; split; intros. + - exploit Mem.load_inject. eauto. eauto. apply inj_of_bc_valid; auto. + intros (v' & A & B). eapply vmatch_inj_top; eauto. + - exploit Mem.loadbytes_inject. eauto. eauto. apply inj_of_bc_valid; auto. + intros (bytes' & A & B). inv B. inv H4. eapply pmatch_inj_top; eauto. + } + constructor; simpl; intros. + - apply ablock_init_sound. apply SM. congruence. + - rewrite PTree.gempty in H1; discriminate. + - apply SM; auto. + - apply SM; auto. + - red; intros. eapply Mem.valid_block_inject_1. eapply inj_of_bc_valid; eauto. eauto. +Qed. + +(** * Abstracting RTL register environments *) + +Module AVal <: SEMILATTICE_WITH_TOP. + + Definition t := aval. + Definition eq (x y: t) := (x = y). + Definition eq_refl: forall x, eq x x := (@refl_equal t). + Definition eq_sym: forall x y, eq x y -> eq y x := (@sym_equal t). + Definition eq_trans: forall x y z, eq x y -> eq y z -> eq x z := (@trans_equal t). + Definition beq (x y: t) : bool := proj_sumbool (eq_aval x y). + Lemma beq_correct: forall x y, beq x y = true -> eq x y. + Proof. unfold beq; intros. InvBooleans. auto. Qed. + Definition ge := vge. + Lemma ge_refl: forall x y, eq x y -> ge x y. + Proof. unfold eq, ge; intros. subst y. apply vge_refl. Qed. + Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z. + Proof. unfold ge; intros. eapply vge_trans; eauto. Qed. + Definition bot : t := Vbot. + Lemma ge_bot: forall x, ge x bot. + Proof. intros. constructor. Qed. + Definition top : t := Vtop. + Lemma ge_top: forall x, ge top x. + Proof. intros. apply vge_top. Qed. + Definition lub := vlub. + Lemma ge_lub_left: forall x y, ge (lub x y) x. + Proof vge_lub_l. + Lemma ge_lub_right: forall x y, ge (lub x y) y. + Proof vge_lub_r. +End AVal. + +Module AE := LPMap(AVal). + +Definition aenv := AE.t. + +Section MATCHENV. + +Variable bc: block_classification. + +Definition ematch (e: regset) (ae: aenv) : Prop := + forall r, vmatch bc e#r (AE.get r ae). + +Lemma ematch_ge: + forall e ae1 ae2, + ematch e ae1 -> AE.ge ae2 ae1 -> ematch e ae2. +Proof. + intros; red; intros. apply vmatch_ge with (AE.get r ae1); auto. apply H0. +Qed. + +Lemma ematch_update: + forall e ae v av r, + ematch e ae -> vmatch bc v av -> ematch (e#r <- v) (AE.set r av ae). +Proof. + intros; red; intros. rewrite AE.gsspec. rewrite PMap.gsspec. + destruct (peq r0 r); auto. + red; intros. specialize (H xH). subst ae. simpl in H. inv H. + unfold AVal.eq; red; intros. subst av. inv H0. +Qed. + +Fixpoint einit_regs (rl: list reg) : aenv := + match rl with + | r1 :: rs => AE.set r1 (Ifptr Nonstack) (einit_regs rs) + | nil => AE.top + end. + +Lemma ematch_init: + forall rl vl, + (forall v, In v vl -> vmatch bc v (Ifptr Nonstack)) -> + ematch (init_regs vl rl) (einit_regs rl). +Proof. + induction rl; simpl; intros. +- red; intros. rewrite Regmap.gi. simpl AE.get. rewrite PTree.gempty. + constructor. +- destruct vl as [ | v1 vs ]. + + assert (ematch (init_regs nil rl) (einit_regs rl)). + { apply IHrl. simpl; tauto. } + replace (init_regs nil rl) with (Regmap.init Vundef) in H0 by (destruct rl; auto). + red; intros. rewrite AE.gsspec. destruct (peq r a). + rewrite Regmap.gi. constructor. + apply H0. + red; intros EQ; rewrite EQ in H0. specialize (H0 xH). simpl in H0. inv H0. + unfold AVal.eq, AVal.bot. congruence. + + assert (ematch (init_regs vs rl) (einit_regs rl)). + { apply IHrl. eauto with coqlib. } + red; intros. rewrite Regmap.gsspec. rewrite AE.gsspec. destruct (peq r a). + auto with coqlib. + apply H0. + red; intros EQ; rewrite EQ in H0. specialize (H0 xH). simpl in H0. inv H0. + unfold AVal.eq, AVal.bot. congruence. +Qed. + +Fixpoint eforget (rl: list reg) (ae: aenv) {struct rl} : aenv := + match rl with + | nil => ae + | r1 :: rs => eforget rs (AE.set r1 Vtop ae) + end. + +Lemma eforget_ge: + forall rl ae, AE.ge (eforget rl ae) ae. +Proof. + unfold AE.ge; intros. revert rl ae; induction rl; intros; simpl. + apply AVal.ge_refl. apply AVal.eq_refl. + destruct ae. unfold AE.get at 2. apply AVal.ge_bot. + eapply AVal.ge_trans. apply IHrl. rewrite AE.gsspec. + destruct (peq p a). apply AVal.ge_top. apply AVal.ge_refl. apply AVal.eq_refl. + congruence. + unfold AVal.eq, Vtop, AVal.bot. congruence. +Qed. + +Lemma ematch_forget: + forall e rl ae, ematch e ae -> ematch e (eforget rl ae). +Proof. + intros. eapply ematch_ge; eauto. apply eforget_ge. +Qed. + +End MATCHENV. + +Lemma ematch_incr: + forall bc bc' e ae, ematch bc e ae -> bc_incr bc bc' -> ematch bc' e ae. +Proof. + intros; red; intros. apply vmatch_incr with bc; auto. +Qed. + +(** * Lattice for dataflow analysis *) + +Module VA <: SEMILATTICE. + + Inductive t' := Bot | State (ae: aenv) (am: amem). + Definition t := t'. + + Definition eq (x y: t) := + match x, y with + | Bot, Bot => True + | State ae1 am1, State ae2 am2 => + AE.eq ae1 ae2 /\ forall bc m, mmatch bc m am1 <-> mmatch bc m am2 + | _, _ => False + end. + + Lemma eq_refl: forall x, eq x x. + Proof. + destruct x; simpl. auto. split. apply AE.eq_refl. tauto. + Qed. + Lemma eq_sym: forall x y, eq x y -> eq y x. + Proof. + destruct x, y; simpl; auto. intros [A B]. + split. apply AE.eq_sym; auto. intros. rewrite B. tauto. + Qed. + Lemma eq_trans: forall x y z, eq x y -> eq y z -> eq x z. + Proof. + destruct x, y, z; simpl; try tauto. intros [A B] [C D]; split. + eapply AE.eq_trans; eauto. + intros. rewrite B; auto. + Qed. + + Definition beq (x y: t) : bool := + match x, y with + | Bot, Bot => true + | State ae1 am1, State ae2 am2 => AE.beq ae1 ae2 && mbeq am1 am2 + | _, _ => false + end. + + Lemma beq_correct: forall x y, beq x y = true -> eq x y. + Proof. + destruct x, y; simpl; intros. + auto. + congruence. + congruence. + InvBooleans; split. + apply AE.beq_correct; auto. + intros. apply mbeq_sound; auto. + Qed. + + Definition ge (x y: t) : Prop := + match x, y with + | _, Bot => True + | Bot, _ => False + | State ae1 am1, State ae2 am2 => AE.ge ae1 ae2 /\ forall bc m, mmatch bc m am2 -> mmatch bc m am1 + end. + + Lemma ge_refl: forall x y, eq x y -> ge x y. + Proof. + destruct x, y; simpl; try tauto. intros [A B]; split. + apply AE.ge_refl; auto. + intros. rewrite B; auto. + Qed. + Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z. + Proof. + destruct x, y, z; simpl; try tauto. intros [A B] [C D]; split. + eapply AE.ge_trans; eauto. + eauto. + Qed. + + Definition bot : t := Bot. + Lemma ge_bot: forall x, ge x bot. + Proof. + destruct x; simpl; auto. + Qed. + + Definition lub (x y: t) : t := + match x, y with + | Bot, _ => y + | _, Bot => x + | State ae1 am1, State ae2 am2 => State (AE.lub ae1 ae2) (mlub am1 am2) + end. + + Lemma ge_lub_left: forall x y, ge (lub x y) x. + Proof. + destruct x, y. + apply ge_refl; apply eq_refl. + simpl. auto. + apply ge_refl; apply eq_refl. + simpl. split. apply AE.ge_lub_left. intros; apply mmatch_lub_l; auto. + Qed. + Lemma ge_lub_right: forall x y, ge (lub x y) y. + Proof. + destruct x, y. + apply ge_refl; apply eq_refl. + apply ge_refl; apply eq_refl. + simpl. auto. + simpl. split. apply AE.ge_lub_right. intros; apply mmatch_lub_r; auto. + Qed. + +End VA. + +Hint Constructors cmatch : va. +Hint Constructors pmatch: va. +Hint Constructors vmatch : va. +Hint Resolve cnot_sound symbol_address_sound + shl_sound shru_sound shr_sound + and_sound or_sound xor_sound notint_sound + ror_sound rolm_sound + neg_sound add_sound sub_sound + mul_sound mulhs_sound mulhu_sound + divs_sound divu_sound mods_sound modu_sound shrx_sound + negf_sound absf_sound + addf_sound subf_sound mulf_sound divf_sound + zero_ext_sound sign_ext_sound singleoffloat_sound + intoffloat_sound intuoffloat_sound floatofint_sound floatofintu_sound + longofwords_sound loword_sound hiword_sound + cmpu_bool_sound cmp_bool_sound cmpf_bool_sound maskzero_sound : va. -- cgit v1.2.3