Library CSEproof

Correctness proof for common subexpression elimination.

Require Import Coqlib.
Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Op.
Require Import Registers.
Require Import RTL.
Require Import Kildall.
Require Import CSE.

Semantic properties of value numberings


Well-formedness of numberings


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.

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 wf_equation (next: valnum) (vr: valnum) (rh: rhs) : Prop :=
  Plt vr next /\ wf_rhs next rh.

Definition wf_numbering (n: numbering) : Prop :=
  (forall v rh,
    In (v, rh) n.(num_eqs) -> wf_equation n.(num_next) v rh)
/\
  (forall r v,
    PTree.get r n.(num_reg) = Some v -> Plt v n.(num_next)).

Lemma wf_empty_numbering:
  wf_numbering empty_numbering.


Lemma wf_rhs_increasing:
  forall next1 next2 rh,
  Ple next1 next2 ->
  wf_rhs next1 rh -> wf_rhs next2 rh.


Lemma wf_equation_increasing:
  forall next1 next2 vr rh,
  Ple next1 next2 ->
  wf_equation next1 vr rh -> wf_equation next2 vr rh.


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 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).


Lemma find_valnum_rhs_correct:
  forall rh vn eqs,
  find_valnum_rhs rh eqs = Some vn -> In (vn, rh) eqs.


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 wf_add_op:
  forall n rd op rs,
  wf_numbering n ->
  wf_numbering (add_op n rd op rs).


Lemma wf_add_load:
  forall n rd chunk addr rs,
  wf_numbering n ->
  wf_numbering (add_load n rd chunk addr rs).


Lemma wf_add_unknown:
  forall n rd,
  wf_numbering n ->
  wf_numbering (add_unknown n rd).


Lemma kill_load_eqs_incl:
  forall eqs, List.incl (kill_load_eqs eqs) eqs.


Lemma wf_kill_loads:
  forall n, wf_numbering n -> wf_numbering (kill_loads n).


Lemma wf_transfer:
  forall f pc n, wf_numbering n -> wf_numbering (transfer f pc n).


As a consequence, the numberings computed by the static analysis are well formed.

Theorem wf_analyze:
  forall f pc, wf_numbering (analyze f)!!pc.


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.
Variable m: mem.

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 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 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.


The numbering_holds predicate (defined in file CSE) is extensional with respect to valu_agree.

Lemma numbering_holds_exten:
  forall valu1 valu2 n rs,
  valu_agree valu1 valu2 n.(num_next) ->
  wf_numbering n ->
  numbering_holds valu1 ge sp rs m n ->
  numbering_holds valu2 ge sp rs m n.


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,
  wf_numbering n ->
  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).


Lemma valnum_regs_holds:
  forall rs rl valu1 n n' vl,
  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).


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) (v: val) : Prop :=
  match rh with
  | Op op vl =>
      eval_operation ge sp op (List.map valu vl) = Some v
  | Load chunk addr vl =>
      exists a,
      eval_addressing ge sp addr (List.map valu vl) = Some a /\
      loadv chunk m a = Some v
  end.

Lemma equation_evals_to_holds_1:
  forall valu rh v vres,
  rhs_evals_to valu rh v ->
  equation_holds valu ge sp m vres rh ->
  valu vres = v.


Lemma equation_evals_to_holds_2:
  forall valu rh v vres,
  wf_rhs vres rh ->
  rhs_evals_to valu rh v ->
  equation_holds (VMap.set vres v valu) ge sp m vres rh.


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:
  forall n rh valu1 rs rd v,
  wf_numbering n ->
  wf_rhs n.(num_next) rh ->
  numbering_holds valu1 ge sp rs m n ->
  rhs_evals_to valu1 rh v ->
  numbering_satisfiable ge sp (rs#rd <- v) m (add_rhs n rd rh).


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,
  wf_numbering n ->
  numbering_satisfiable ge sp rs m n ->
  eval_operation ge sp op rs##args = Some v ->
  numbering_satisfiable ge sp (rs#dst <- v) m (add_op n dst op args).


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,
  wf_numbering n ->
  numbering_satisfiable ge sp rs m n ->
  eval_addressing ge sp addr rs##args = Some a ->
  loadv chunk m a = Some v ->
  numbering_satisfiable ge sp
     (rs#dst <- v)
     m (add_load n dst chunk addr args).


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,
  wf_numbering n ->
  numbering_satisfiable ge sp rs m n ->
  numbering_satisfiable ge sp
     (rs#dst <- v) m (add_unknown n dst).


Allocation of a fresh memory block preserves satisfiability.

Lemma alloc_satisfiable:
  forall lo hi b m' rs n,
  Mem.alloc m lo hi = (m', b) ->
  numbering_satisfiable ge sp rs m n ->
  numbering_satisfiable ge sp rs m' n.


kill_load preserves satisfiability. Moreover, the resulting numbering is satisfiable in any concrete memory state.

Lemma kill_load_eqs_ops:
  forall v rhs eqs,
  In (v, rhs) (kill_load_eqs eqs) ->
  match rhs with Op _ _ => True | Load _ _ _ => False end.


Lemma kill_load_satisfiable:
  forall n rs chunk addr v m',
  Mem.storev chunk m addr v = Some m' ->
  numbering_satisfiable ge sp rs m n ->
  numbering_satisfiable ge sp rs m' (kill_loads n).


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, reg_valnum n v = Some r -> n.(num_reg)!r = Some v.


Correctness of find_op and find_load: 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 n rh r,
  numbering_holds valu ge sp rs m n ->
  find_rhs n rh = Some r ->
  rhs_evals_to valu rh rs#r.


Lemma find_op_correct:
  forall rs n op args r,
  wf_numbering n ->
  numbering_satisfiable ge sp rs m n ->
  find_op n op args = Some r ->
  eval_operation ge sp op rs##args = Some rs#r.


Lemma find_load_correct:
  forall rs n chunk addr args r,
  wf_numbering n ->
  numbering_satisfiable ge sp rs m n ->
  find_load n chunk addr args = Some r ->
  exists a,
  eval_addressing ge sp addr rs##args = Some a /\
  loadv chunk m a = Some rs#r.


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 from pc to pc', satisfiability at pc implies satisfiability at pc'.

Theorem analysis_correct_1:
  forall ge sp rs m f pc pc' i,
  f.(fn_code)!pc = Some i -> In pc' (successors_instr i) ->
  numbering_satisfiable ge sp rs m (transfer f pc (analyze f)!!pc) ->
  numbering_satisfiable ge sp rs m (analyze f)!!pc'.


Theorem analysis_correct_entry:
  forall ge sp rs m f,
  numbering_satisfiable ge sp rs m (analyze f)!!(f.(fn_entrypoint)).


Semantic preservation


Section PRESERVATION.

Variable prog: program.
Let tprog := transf_program prog.
Let ge := Genv.globalenv prog.
Let tge := Genv.globalenv tprog.

Lemma symbols_preserved:
  forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof (Genv.find_symbol_transf transf_fundef prog).

Lemma functions_translated:
  forall (v: val) (f: RTL.fundef),
  Genv.find_funct ge v = Some f ->
  Genv.find_funct tge v = Some (transf_fundef f).
Proof (@Genv.find_funct_transf _ _ _ transf_fundef prog).

Lemma funct_ptr_translated:
  forall (b: block) (f: RTL.fundef),
  Genv.find_funct_ptr ge b = Some f ->
  Genv.find_funct_ptr tge b = Some (transf_fundef f).
Proof (@Genv.find_funct_ptr_transf _ _ _ transf_fundef prog).

Lemma sig_preserved:
  forall f, funsig (transf_fundef f) = funsig f.


Lemma find_function_translated:
  forall ros rs f,
  find_function ge ros rs = Some f ->
  find_function tge ros rs = Some (transf_fundef f).


The proof of semantic preservation is a simulation argument using diagrams of the following form:
           st1 --------------- st2
            |                   |
           t|                   |t
            |                   |
            v                   v
           st1'--------------- st2'


Left: RTL execution in the original program. Right: RTL execution in the optimized program. Precondition (top) and postcondition (bottom): agreement between the states, including the fact that the numbering at pc (returned by the static analysis) is satisfiable.

Inductive match_stackframes: stackframe -> stackframe -> Prop :=
  match_stackframes_intro:
    forall res c sp pc rs f,
    c = f.(RTL.fn_code) ->
    (forall v m, numbering_satisfiable ge sp (rs#res <- v) m (analyze f)!!pc) ->
    match_stackframes
      (Stackframe res c sp pc rs)
      (Stackframe res (transf_code (analyze f) c) sp pc rs).

Inductive match_states: state -> state -> Prop :=
  | match_states_intro:
      forall s c sp pc rs m s' f
             (CF: c = f.(RTL.fn_code))
             (SAT: numbering_satisfiable ge sp rs m (analyze f)!!pc)
             (STACKS: list_forall2 match_stackframes s s'),
      match_states (State s c sp pc rs m)
                   (State s' (transf_code (analyze f) c) sp pc rs m)
  | match_states_call:
      forall s f args m s',
      list_forall2 match_stackframes s s' ->
      match_states (Callstate s f args m)
                   (Callstate s' (transf_fundef f) args m)
  | match_states_return:
      forall s s' v m,
      list_forall2 match_stackframes s s' ->
      match_states (Returnstate s v m)
                   (Returnstate s' v m).

Ltac TransfInstr :=
  match goal with
  | H1: (PTree.get ?pc ?c = Some ?instr), f: function |- _ =>
      cut ((transf_code (analyze f) c)!pc = Some(transf_instr (analyze f)!!pc instr));
      [ simpl | unfold transf_code; rewrite PTree.gmap; unfold option_map; rewrite H1; reflexivity ]
  end.

The proof of simulation is a case analysis over the transition in the source code.

Lemma transf_step_correct:
  forall s1 t s2, step ge s1 t s2 ->
  forall s1' (MS: match_states s1 s1'),
  exists s2', step tge s1' t s2' /\ match_states s2 s2'.


Lemma transf_initial_states:
  forall st1, initial_state prog st1 ->
  exists st2, initial_state tprog st2 /\ match_states st1 st2.


Lemma transf_final_states:
  forall st1 st2 r,
  match_states st1 st2 -> final_state st1 r -> final_state st2 r.


Theorem transf_program_correct:
  forall (beh: program_behavior), not_wrong beh ->
  exec_program prog beh -> exec_program tprog beh.


End PRESERVATION.