Library RTLgenspec

Abstract specification of RTL generation

In this module, we define inductive predicates that specify the translations from Cminor to RTL performed by the functions in module RTLgen. We then show that these functions satisfy these relational specifications. The relational specifications will then be used instead of the actual functions to show semantic equivalence between the source Cminor code and the the generated RTL code (see module RTLgenproof).

Require Import Coqlib.
Require Errors.
Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Events.
Require Import Mem.
Require Import Globalenvs.
Require Import Switch.
Require Import Op.
Require Import Registers.
Require Import CminorSel.
Require Import RTL.
Require Import RTLgen.

Reasoning about monadic computations


The tactics below simplify hypotheses of the form f ... = OK x s i where f is a monadic computation. For instance, the hypothesis (do x <- a; b) s = OK y s' i will generate the additional witnesses x, s1, i1, i' and the additional hypotheses a s = OK x s1 i1 and b x s1 = OK y s' i', reflecting the fact that both monadic computations a and b succeeded.

Remark bind_inversion:
  forall (A B: Type) (f: mon A) (g: A -> mon B)
         (y: B) (s1 s3: state) (i: state_incr s1 s3),
  bind f g s1 = OK y s3 i ->
  exists x, exists s2, exists i1, exists i2,
  f s1 = OK x s2 i1 /\ g x s2 = OK y s3 i2.


Remark bind2_inversion:
  forall (A B C: Type) (f: mon (A*B)) (g: A -> B -> mon C)
         (z: C) (s1 s3: state) (i: state_incr s1 s3),
  bind2 f g s1 = OK z s3 i ->
  exists x, exists y, exists s2, exists i1, exists i2,
  f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2.


Ltac monadInv1 H :=
  match type of H with
  | (OK _ _ _ = OK _ _ _) =>
      inversion H; clear H; try subst
  | (Error _ _ = OK _ _ _) =>
      discriminate
  | (ret _ _ = OK _ _ _) =>
      inversion H; clear H; try subst
  | (error _ _ = OK _ _ _) =>
      discriminate
  | (bind ?F ?G ?S = OK ?X ?S' ?I) =>
      let x := fresh "x" in (
      let s := fresh "s" in (
      let i1 := fresh "INCR" in (
      let i2 := fresh "INCR" in (
      let EQ1 := fresh "EQ" in (
      let EQ2 := fresh "EQ" in (
      destruct (bind_inversion _ _ F G X S S' I H) as [x [s [i1 [i2 [EQ1 EQ2]]]]];
      clear H;
      try (monadInv1 EQ2)))))))
  | (bind2 ?F ?G ?S = OK ?X ?S' ?I) =>
      let x1 := fresh "x" in (
      let x2 := fresh "x" in (
      let s := fresh "s" in (
      let i1 := fresh "INCR" in (
      let i2 := fresh "INCR" in (
      let EQ1 := fresh "EQ" in (
      let EQ2 := fresh "EQ" in (
      destruct (bind2_inversion _ _ _ F G X S S' I H) as [x1 [x2 [s [i1 [i2 [EQ1 EQ2]]]]]];
      clear H;
      try (monadInv1 EQ2))))))))
  end.

Ltac monadInv H :=
  match type of H with
  | (ret _ _ = OK _ _ _) => monadInv1 H
  | (error _ _ = OK _ _ _) => monadInv1 H
  | (bind ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H
  | (bind2 ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H
  | (?F _ _ _ _ _ _ _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ _ _ _ _ _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ _ _ _ _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ _ _ _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ _ _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  | (?F _ = OK _ _ _) =>
      ((progress simpl in H) || unfold F in H); monadInv1 H
  end.

Monotonicity properties of the state


Lemma instr_at_incr:
  forall s1 s2 n i,
  state_incr s1 s2 -> s1.(st_code)!n = Some i -> s2.(st_code)!n = Some i.

Hint Resolve instr_at_incr: rtlg.

Hint Resolve state_incr_refl state_incr_trans: rtlg.

Validity and freshness of registers


An RTL pseudo-register is valid in a given state if it was created earlier, that is, it is less than the next fresh register of the state. Otherwise, the pseudo-register is said to be fresh.

Definition reg_valid (r: reg) (s: state) : Prop :=
  Plt r s.(st_nextreg).

Definition reg_fresh (r: reg) (s: state) : Prop :=
  ~(Plt r s.(st_nextreg)).

Lemma valid_fresh_absurd:
  forall r s, reg_valid r s -> reg_fresh r s -> False.

Hint Resolve valid_fresh_absurd: rtlg.

Lemma valid_fresh_different:
  forall r1 r2 s, reg_valid r1 s -> reg_fresh r2 s -> r1 <> r2.

Hint Resolve valid_fresh_different: rtlg.

Lemma reg_valid_incr:
  forall r s1 s2, state_incr s1 s2 -> reg_valid r s1 -> reg_valid r s2.

Hint Resolve reg_valid_incr: rtlg.

Lemma reg_fresh_decr:
  forall r s1 s2, state_incr s1 s2 -> reg_fresh r s2 -> reg_fresh r s1.

Hint Resolve reg_fresh_decr: rtlg.

Validity of a list of registers.

Definition regs_valid (rl: list reg) (s: state) : Prop :=
  forall r, In r rl -> reg_valid r s.

Lemma regs_valid_nil:
  forall s, regs_valid nil s.

Hint Resolve regs_valid_nil: rtlg.

Lemma regs_valid_cons:
  forall r1 rl s,
  reg_valid r1 s -> regs_valid rl s -> regs_valid (r1 :: rl) s.


Lemma regs_valid_incr:
  forall s1 s2 rl, state_incr s1 s2 -> regs_valid rl s1 -> regs_valid rl s2.

Hint Resolve regs_valid_incr: rtlg.

A register is ``in'' a mapping if it is associated with a Cminor local or let-bound variable.

Definition reg_in_map (m: mapping) (r: reg) : Prop :=
  (exists id, m.(map_vars)!id = Some r) \/ In r m.(map_letvars).

A compilation environment (mapping) is valid in a given state if the registers associated with Cminor local variables and let-bound variables are valid in the state.

Definition map_valid (m: mapping) (s: state) : Prop :=
  forall r, reg_in_map m r -> reg_valid r s.

Lemma map_valid_incr:
  forall s1 s2 m,
  state_incr s1 s2 -> map_valid m s1 -> map_valid m s2.

Hint Resolve map_valid_incr: rtlg.

Properties of basic operations over the state


Properties of add_instr.

Lemma add_instr_at:
  forall s1 s2 incr i n,
  add_instr i s1 = OK n s2 incr -> s2.(st_code)!n = Some i.

Hint Resolve add_instr_at: rtlg.

Properties of update_instr.

Lemma update_instr_at:
  forall n i s1 s2 incr u,
  update_instr n i s1 = OK u s2 incr -> s2.(st_code)!n = Some i.

Hint Resolve update_instr_at: rtlg.

Properties of new_reg.

Lemma new_reg_valid:
  forall s1 s2 r i,
  new_reg s1 = OK r s2 i -> reg_valid r s2.

Hint Resolve new_reg_valid: rtlg.

Lemma new_reg_fresh:
  forall s1 s2 r i,
  new_reg s1 = OK r s2 i -> reg_fresh r s1.

Hint Resolve new_reg_fresh: rtlg.

Lemma new_reg_not_in_map:
  forall s1 s2 m r i,
  new_reg s1 = OK r s2 i -> map_valid m s1 -> ~(reg_in_map m r).

Hint Resolve new_reg_not_in_map: rtlg.

Properties of operations over compilation environments


Lemma init_mapping_valid:
  forall s, map_valid init_mapping s.


Properties of find_var.

Lemma find_var_in_map:
  forall s1 s2 map name r i,
  find_var map name s1 = OK r s2 i -> reg_in_map map r.

Hint Resolve find_var_in_map: rtlg.

Lemma find_var_valid:
  forall s1 s2 map name r i,
  find_var map name s1 = OK r s2 i -> map_valid map s1 -> reg_valid r s1.

Hint Resolve find_var_valid: rtlg.

Properties of find_letvar.

Lemma find_letvar_in_map:
  forall s1 s2 map idx r i,
  find_letvar map idx s1 = OK r s2 i -> reg_in_map map r.

Hint Resolve find_letvar_in_map: rtlg.

Lemma find_letvar_valid:
  forall s1 s2 map idx r i,
  find_letvar map idx s1 = OK r s2 i -> map_valid map s1 -> reg_valid r s1.

Hint Resolve find_letvar_valid: rtlg.

Properties of add_var.

Lemma add_var_valid:
  forall s1 s2 map1 map2 name r i,
  add_var map1 name s1 = OK (r, map2) s2 i ->
  map_valid map1 s1 ->
  reg_valid r s2 /\ map_valid map2 s2.


Lemma add_var_find:
  forall s1 s2 map name r map' i,
  add_var map name s1 = OK (r,map') s2 i -> map'.(map_vars)!name = Some r.


Lemma add_vars_valid:
  forall namel s1 s2 map1 map2 rl i,
  add_vars map1 namel s1 = OK (rl, map2) s2 i ->
  map_valid map1 s1 ->
  regs_valid rl s2 /\ map_valid map2 s2.


Lemma add_var_letenv:
  forall map1 id s1 r map2 s2 i,
  add_var map1 id s1 = OK (r, map2) s2 i ->
  map2.(map_letvars) = map1.(map_letvars).


Lemma add_vars_letenv:
  forall il map1 s1 rl map2 s2 i,
  add_vars map1 il s1 = OK (rl, map2) s2 i ->
  map2.(map_letvars) = map1.(map_letvars).


Properties of add_letvar.

Lemma add_letvar_valid:
  forall map s r,
  map_valid map s ->
  reg_valid r s ->
  map_valid (add_letvar map r) s.


Properties of alloc_reg and alloc_regs


Lemma alloc_reg_valid:
  forall a s1 s2 map r i,
  map_valid map s1 ->
  alloc_reg map a s1 = OK r s2 i -> reg_valid r s2.

Hint Resolve alloc_reg_valid: rtlg.

Lemma alloc_reg_fresh_or_in_map:
  forall map a s r s' i,
  map_valid map s ->
  alloc_reg map a s = OK r s' i ->
  reg_in_map map r \/ reg_fresh r s.


Lemma alloc_regs_valid:
  forall al s1 s2 map rl i,
  map_valid map s1 ->
  alloc_regs map al s1 = OK rl s2 i ->
  regs_valid rl s2.

Hint Resolve alloc_regs_valid: rtlg.

Lemma alloc_regs_fresh_or_in_map:
  forall map al s rl s' i,
  map_valid map s ->
  alloc_regs map al s = OK rl s' i ->
  forall r, In r rl -> reg_in_map map r \/ reg_fresh r s.


A register is an adequate target for holding the value of an expression if
  • either the register is associated with a Cminor let-bound variable or a Cminor local variable;
  • or the register is not associated with any Cminor variable and does not belong to the preserved set pr.

Inductive target_reg_ok (map: mapping) (pr: list reg): expr -> reg -> Prop :=
  | target_reg_var:
      forall id r,
      map.(map_vars)!id = Some r ->
      target_reg_ok map pr (Evar id) r
  | target_reg_letvar:
      forall idx r,
      nth_error map.(map_letvars) idx = Some r ->
      target_reg_ok map pr (Eletvar idx) r
  | target_reg_other:
      forall a r,
      ~(reg_in_map map r) -> ~In r pr ->
      target_reg_ok map pr a r.

Inductive target_regs_ok (map: mapping) (pr: list reg): exprlist -> list reg -> Prop :=
  | target_regs_nil:
      target_regs_ok map pr Enil nil
  | target_regs_cons: forall a1 al r1 rl,
      target_reg_ok map pr a1 r1 ->
      target_regs_ok map (r1 :: pr) al rl ->
      target_regs_ok map pr (Econs a1 al) (r1 :: rl).

Lemma target_reg_ok_append:
  forall map pr a r,
  target_reg_ok map pr a r ->
  forall pr',
  (forall r', In r' pr' -> reg_in_map map r' \/ r' <> r) ->
  target_reg_ok map (pr' ++ pr) a r.


Lemma target_reg_ok_cons:
  forall map pr a r,
  target_reg_ok map pr a r ->
  forall r',
  reg_in_map map r' \/ r' <> r ->
  target_reg_ok map (r' :: pr) a r.


Lemma new_reg_target_ok:
  forall map pr s1 a r s2 i,
  map_valid map s1 ->
  regs_valid pr s1 ->
  new_reg s1 = OK r s2 i ->
  target_reg_ok map pr a r.


Lemma alloc_reg_target_ok:
  forall map pr s1 a r s2 i,
  map_valid map s1 ->
  regs_valid pr s1 ->
  alloc_reg map a s1 = OK r s2 i ->
  target_reg_ok map pr a r.


Lemma alloc_regs_target_ok:
  forall map al pr s1 rl s2 i,
  map_valid map s1 ->
  regs_valid pr s1 ->
  alloc_regs map al s1 = OK rl s2 i ->
  target_regs_ok map pr al rl.


Hint Resolve new_reg_target_ok alloc_reg_target_ok alloc_regs_target_ok: rtlg.

The following predicate is a variant of target_reg_ok used to characterize registers that are adequate for holding the return value of a function.

Inductive return_reg_ok: state -> mapping -> option reg -> Prop :=
  | return_reg_ok_none:
      forall s map,
      return_reg_ok s map None
  | return_reg_ok_some:
      forall s map r,
      ~(reg_in_map map r) -> reg_valid r s ->
      return_reg_ok s map (Some r).

Lemma return_reg_ok_incr:
  forall s map rret, return_reg_ok s map rret ->
  forall s', state_incr s s' -> return_reg_ok s' map rret.

Hint Resolve return_reg_ok_incr: rtlg.

Lemma new_reg_return_ok:
  forall s1 r s2 map sig i,
  new_reg s1 = OK r s2 i ->
  map_valid map s1 ->
  return_reg_ok s2 map (ret_reg sig r).


Relational specification of the translation


We now define inductive predicates that characterize the fact that the control-flow graph produced by compilation of a Cminor function contains sub-graphs that correspond to the translation of each Cminor expression or statement in the original code.

tr_move c ns rs nd rd holds if the graph c, between nodes ns and nd, contains instructions that move the value of register rs to register rd.

Inductive tr_move (c: code):
       node -> reg -> node -> reg -> Prop :=
  | tr_move_0: forall n r,
      tr_move c n r n r
  | tr_move_1: forall ns rs nd rd,
      c!ns = Some (Iop Omove (rs :: nil) rd nd) ->
      tr_move c ns rs nd rd.

tr_expr c map pr expr ns nd rd holds if the graph c, between nodes ns and nd, contains instructions that compute the value of the Cminor expression expr and deposit this value in register rd. map is a mapping from Cminor variables to the corresponding RTL registers. pr is a list of RTL registers whose values must be preserved during this computation. All registers mentioned in map must also be preserved during this computation. To ensure this, we demand that the result registers of the instructions appearing on the path from ns to nd are neither in pr nor in map.

Inductive tr_expr (c: code):
       mapping -> list reg -> expr -> node -> node -> reg -> Prop :=
  | tr_Evar: forall map pr id ns nd r rd,
      map.(map_vars)!id = Some r ->
      (rd = r \/ ~reg_in_map map rd /\ ~In rd pr) ->
      tr_move c ns r nd rd ->
      tr_expr c map pr (Evar id) ns nd rd
  | tr_Eop: forall map pr op al ns nd rd n1 rl,
      tr_exprlist c map pr al ns n1 rl ->
      c!n1 = Some (Iop op rl rd nd) ->
      ~reg_in_map map rd -> ~In rd pr ->
      tr_expr c map pr (Eop op al) ns nd rd
  | tr_Eload: forall map pr chunk addr al ns nd rd n1 rl,
      tr_exprlist c map pr al ns n1 rl ->
      c!n1 = Some (Iload chunk addr rl rd nd) ->
      ~reg_in_map map rd -> ~In rd pr ->
      tr_expr c map pr (Eload chunk addr al) ns nd rd
  | tr_Econdition: forall map pr b ifso ifnot ns nd rd ntrue nfalse,
      tr_condition c map pr b ns ntrue nfalse ->
      tr_expr c map pr ifso ntrue nd rd ->
      tr_expr c map pr ifnot nfalse nd rd ->
      tr_expr c map pr (Econdition b ifso ifnot) ns nd rd
  | tr_Elet: forall map pr b1 b2 ns nd rd n1 r,
      ~reg_in_map map r ->
      tr_expr c map pr b1 ns n1 r ->
      tr_expr c (add_letvar map r) pr b2 n1 nd rd ->
      tr_expr c map pr (Elet b1 b2) ns nd rd
  | tr_Eletvar: forall map pr n ns nd rd r,
      List.nth_error map.(map_letvars) n = Some r ->
      (rd = r \/ ~reg_in_map map rd /\ ~In rd pr) ->
      tr_move c ns r nd rd ->
      tr_expr c map pr (Eletvar n) ns nd rd

tr_condition c map pr cond ns ntrue nfalse rd holds if the graph c, starting at node ns, contains instructions that compute the truth value of the Cminor conditional expression cond and terminate on node ntrue if the condition holds and on node nfalse otherwise.

with tr_condition (c: code):
       mapping -> list reg -> condexpr -> node -> node -> node -> Prop :=
  | tr_CEtrue: forall map pr ntrue nfalse,
      tr_condition c map pr CEtrue ntrue ntrue nfalse
  | tr_CEfalse: forall map pr ntrue nfalse,
      tr_condition c map pr CEfalse nfalse ntrue nfalse
  | tr_CEcond: forall map pr cond bl ns ntrue nfalse n1 rl,
      tr_exprlist c map pr bl ns n1 rl ->
      c!n1 = Some (Icond cond rl ntrue nfalse) ->
      tr_condition c map pr (CEcond cond bl) ns ntrue nfalse
  | tr_CEcondition: forall map pr b ifso ifnot ns ntrue nfalse ntrue' nfalse',
      tr_condition c map pr b ns ntrue' nfalse' ->
      tr_condition c map pr ifso ntrue' ntrue nfalse ->
      tr_condition c map pr ifnot nfalse' ntrue nfalse ->
      tr_condition c map pr (CEcondition b ifso ifnot) ns ntrue nfalse

tr_exprlist c map pr exprs ns nd rds holds if the graph c, between nodes ns and nd, contains instructions that compute the values of the list of Cminor expression exprlist and deposit these values in registers rds.

with tr_exprlist (c: code):
       mapping -> list reg -> exprlist -> node -> node -> list reg -> Prop :=
  | tr_Enil: forall map pr n,
      tr_exprlist c map pr Enil n n nil
  | tr_Econs: forall map pr a1 al ns nd r1 rl n1,
      tr_expr c map pr a1 ns n1 r1 ->
      tr_exprlist c map (r1 :: pr) al n1 nd rl ->
      tr_exprlist c map pr (Econs a1 al) ns nd (r1 :: rl).

Auxiliary for the compilation of variable assignments.

Definition tr_store_var (c: code) (map: mapping)
                        (rs: reg) (id: ident) (ns nd: node): Prop :=
  exists rv, map.(map_vars)!id = Some rv /\ tr_move c ns rs nd rv.

Definition tr_store_optvar (c: code) (map: mapping)
                 (rs: reg) (optid: option ident) (ns nd: node): Prop :=
  match optid with
  | None => ns = nd
  | Some id => tr_store_var c map rs id ns nd
  end.

Auxiliary for the compilation of switch statements.

Definition tr_jumptable (nexits: list node) (tbl: list nat) (ttbl: list node) : Prop :=
  forall v act,
  list_nth_z tbl v = Some act ->
  exists n, list_nth_z ttbl v = Some n /\ nth_error nexits act = Some n.

Inductive tr_switch
     (c: code) (map: mapping) (r: reg) (nexits: list node):
     comptree -> node -> Prop :=
  | tr_switch_action: forall act n,
      nth_error nexits act = Some n ->
      tr_switch c map r nexits (CTaction act) n
  | tr_switch_ifeq: forall key act t' n ncont nfound,
      tr_switch c map r nexits t' ncont ->
      nth_error nexits act = Some nfound ->
      c!n = Some(Icond (Ccompimm Ceq key) (r :: nil) nfound ncont) ->
      tr_switch c map r nexits (CTifeq key act t') n
  | tr_switch_iflt: forall key t1 t2 n n1 n2,
      tr_switch c map r nexits t1 n1 ->
      tr_switch c map r nexits t2 n2 ->
      c!n = Some(Icond (Ccompuimm Clt key) (r :: nil) n1 n2) ->
      tr_switch c map r nexits (CTiflt key t1 t2) n
  | tr_switch_jumptable: forall ofs sz tbl t n n1 n2 n3 rt ttbl,
      ~reg_in_map map rt -> rt <> r ->
      c!n = Some(Iop (if Int.eq ofs Int.zero then Omove else Oaddimm (Int.neg ofs))
                     (r ::nil) rt n1) ->
      c!n1 = Some(Icond (Ccompuimm Clt sz) (rt :: nil) n2 n3) ->
      c!n2 = Some(Ijumptable rt ttbl) ->
      tr_switch c map r nexits t n3 ->
      tr_jumptable nexits tbl ttbl ->
      tr_switch c map r nexits (CTjumptable ofs sz tbl t) n.

tr_stmt c map stmt ns ncont nexits nret rret holds if the graph c, starting at node ns, contains instructions that perform the Cminor statement stmt. These instructions branch to node ncont if the statement terminates normally, to the n-th node in nexits if the statement terminates prematurely on a exit n statement, and to nret if the statement terminates prematurely on a return statement. Moreover, if the return statement has an argument, its value is deposited in register rret.

Inductive tr_stmt (c: code) (map: mapping):
     stmt -> node -> node -> list node -> labelmap -> node -> option reg -> Prop :=
  | tr_Sskip: forall ns nexits ngoto nret rret,
     tr_stmt c map Sskip ns ns nexits ngoto nret rret
  | tr_Sassign: forall id a ns nd nexits ngoto nret rret rt n,
     tr_expr c map nil a ns n rt ->
     tr_store_var c map rt id n nd ->
     tr_stmt c map (Sassign id a) ns nd nexits ngoto nret rret
  | tr_Sstore: forall chunk addr al b ns nd nexits ngoto nret rret rd n1 rl n2,
     tr_exprlist c map nil al ns n1 rl ->
     tr_expr c map rl b n1 n2 rd ->
     c!n2 = Some (Istore chunk addr rl rd nd) ->
     tr_stmt c map (Sstore chunk addr al b) ns nd nexits ngoto nret rret
  | tr_Scall: forall optid sig b cl ns nd nexits ngoto nret rret rd n1 rf n2 n3 rargs,
     tr_expr c map nil b ns n1 rf ->
     tr_exprlist c map (rf :: nil) cl n1 n2 rargs ->
     c!n2 = Some (Icall sig (inl _ rf) rargs rd n3) ->
     tr_store_optvar c map rd optid n3 nd ->
     ~reg_in_map map rd ->
     tr_stmt c map (Scall optid sig b cl) ns nd nexits ngoto nret rret
  | tr_Stailcall: forall sig b cl ns nd nexits ngoto nret rret n1 rf n2 rargs,
     tr_expr c map nil b ns n1 rf ->
     tr_exprlist c map (rf :: nil) cl n1 n2 rargs ->
     c!n2 = Some (Itailcall sig (inl _ rf) rargs) ->
     tr_stmt c map (Stailcall sig b cl) ns nd nexits ngoto nret rret
  | tr_Sseq: forall s1 s2 ns nd nexits ngoto nret rret n,
     tr_stmt c map s2 n nd nexits ngoto nret rret ->
     tr_stmt c map s1 ns n nexits ngoto nret rret ->
     tr_stmt c map (Sseq s1 s2) ns nd nexits ngoto nret rret
  | tr_Sifthenelse: forall a strue sfalse ns nd nexits ngoto nret rret ntrue nfalse,
     tr_stmt c map strue ntrue nd nexits ngoto nret rret ->
     tr_stmt c map sfalse nfalse nd nexits ngoto nret rret ->
     tr_condition c map nil a ns ntrue nfalse ->
     tr_stmt c map (Sifthenelse a strue sfalse) ns nd nexits ngoto nret rret
  | tr_Sloop: forall sbody ns nd nexits ngoto nret rret nloop,
     tr_stmt c map sbody nloop ns nexits ngoto nret rret ->
     c!ns = Some(Inop nloop) ->
     tr_stmt c map (Sloop sbody) ns nd nexits ngoto nret rret
  | tr_Sblock: forall sbody ns nd nexits ngoto nret rret,
     tr_stmt c map sbody ns nd (nd :: nexits) ngoto nret rret ->
     tr_stmt c map (Sblock sbody) ns nd nexits ngoto nret rret
  | tr_Sexit: forall n ns nd nexits ngoto nret rret,
     nth_error nexits n = Some ns ->
     tr_stmt c map (Sexit n) ns nd nexits ngoto nret rret
  | tr_Sswitch: forall a cases default ns nd nexits ngoto nret rret n r t,
     validate_switch default cases t = true ->
     tr_expr c map nil a ns n r ->
     tr_switch c map r nexits t n ->
     tr_stmt c map (Sswitch a cases default) ns nd nexits ngoto nret rret
  | tr_Sreturn_none: forall nret nd nexits ngoto,
     tr_stmt c map (Sreturn None) nret nd nexits ngoto nret None
  | tr_Sreturn_some: forall a ns nd nexits ngoto nret rret,
     tr_expr c map nil a ns nret rret ->
     tr_stmt c map (Sreturn (Some a)) ns nd nexits ngoto nret (Some rret)
  | tr_Slabel: forall lbl s ns nd nexits ngoto nret rret n,
     ngoto!lbl = Some n ->
     c!n = Some (Inop ns) ->
     tr_stmt c map s ns nd nexits ngoto nret rret ->
     tr_stmt c map (Slabel lbl s) ns nd nexits ngoto nret rret
  | tr_Sgoto: forall lbl ns nd nexits ngoto nret rret,
     ngoto!lbl = Some ns ->
     tr_stmt c map (Sgoto lbl) ns nd nexits ngoto nret rret.

tr_function f tf specifies the RTL function tf that RTLgen.transl_function returns.

Inductive tr_function: CminorSel.function -> RTL.function -> Prop :=
  | tr_function_intro:
      forall f code rparams map1 s0 s1 i1 rvars map2 s2 i2 nentry ngoto nret rret orret,
      add_vars init_mapping f.(CminorSel.fn_params) s0 = OK (rparams, map1) s1 i1 ->
      add_vars map1 f.(CminorSel.fn_vars) s1 = OK (rvars, map2) s2 i2 ->
      orret = ret_reg f.(CminorSel.fn_sig) rret ->
      tr_stmt code map2 f.(CminorSel.fn_body) nentry nret nil ngoto nret orret ->
      code!nret = Some(Ireturn orret) ->
      tr_function f (RTL.mkfunction
                       f.(CminorSel.fn_sig)
                       rparams
                       f.(CminorSel.fn_stackspace)
                       code
                       nentry).

Correctness proof of the translation functions


We now show that the translation functions in module RTLgen satisfy the specifications given above as inductive predicates.

Lemma tr_move_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall ns rs nd rd,
  tr_move s1.(st_code) ns rs nd rd ->
  tr_move s2.(st_code) ns rs nd rd.


Lemma tr_expr_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map pr a ns nd rd,
  tr_expr s1.(st_code) map pr a ns nd rd ->
  tr_expr s2.(st_code) map pr a ns nd rd
with tr_condition_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map pr a ns ntrue nfalse,
  tr_condition s1.(st_code) map pr a ns ntrue nfalse ->
  tr_condition s2.(st_code) map pr a ns ntrue nfalse
with tr_exprlist_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map pr al ns nd rl,
  tr_exprlist s1.(st_code) map pr al ns nd rl ->
  tr_exprlist s2.(st_code) map pr al ns nd rl.


Lemma add_move_charact:
  forall s ns rs nd rd s' i,
  add_move rs rd nd s = OK ns s' i ->
  tr_move s'.(st_code) ns rs nd rd.


Lemma transl_expr_charact:
  forall a map rd nd s ns s' pr INCR
     (TR: transl_expr map a rd nd s = OK ns s' INCR)
     (WF: map_valid map s)
     (OK: target_reg_ok map pr a rd)
     (VALID: regs_valid pr s)
     (VALID2: reg_valid rd s),
   tr_expr s'.(st_code) map pr a ns nd rd

with transl_condexpr_charact:
  forall a map ntrue nfalse s ns s' pr INCR
     (TR: transl_condition map a ntrue nfalse s = OK ns s' INCR)
     (VALID: regs_valid pr s)
     (WF: map_valid map s),
   tr_condition s'.(st_code) map pr a ns ntrue nfalse

with transl_exprlist_charact:
  forall al map rl nd s ns s' pr INCR
     (TR: transl_exprlist map al rl nd s = OK ns s' INCR)
     (WF: map_valid map s)
     (OK: target_regs_ok map pr al rl)
     (VALID1: regs_valid pr s)
     (VALID2: regs_valid rl s),
   tr_exprlist s'.(st_code) map pr al ns nd rl.



Lemma tr_store_var_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map rs id ns nd,
  tr_store_var s1.(st_code) map rs id ns nd ->
  tr_store_var s2.(st_code) map rs id ns nd.


Lemma tr_store_optvar_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map rs optid ns nd,
  tr_store_optvar s1.(st_code) map rs optid ns nd ->
  tr_store_optvar s2.(st_code) map rs optid ns nd.


Lemma tr_switch_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map r nexits t ns,
  tr_switch s1.(st_code) map r nexits t ns ->
  tr_switch s2.(st_code) map r nexits t ns.


Lemma tr_stmt_incr:
  forall s1 s2, state_incr s1 s2 ->
  forall map s ns nd nexits ngoto nret rret,
  tr_stmt s1.(st_code) map s ns nd nexits ngoto nret rret ->
  tr_stmt s2.(st_code) map s ns nd nexits ngoto nret rret.


Lemma store_var_charact:
  forall map rs id nd s ns s' incr,
  store_var map rs id nd s = OK ns s' incr ->
  tr_store_var s'.(st_code) map rs id ns nd.


Lemma store_optvar_charact:
  forall map rs optid nd s ns s' incr,
  store_optvar map rs optid nd s = OK ns s' incr ->
  tr_store_optvar s'.(st_code) map rs optid ns nd.


Lemma transl_exit_charact:
  forall nexits n s ne s' incr,
  transl_exit nexits n s = OK ne s' incr ->
  nth_error nexits n = Some ne /\ s' = s.


Lemma transl_jumptable_charact:
  forall nexits tbl s nl s' incr,
  transl_jumptable nexits tbl s = OK nl s' incr ->
  tr_jumptable nexits tbl nl /\ s' = s.


Lemma transl_switch_charact:
  forall map r nexits t s ns s' incr,
  map_valid map s -> reg_valid r s ->
  transl_switch r nexits t s = OK ns s' incr ->
  tr_switch s'.(st_code) map r nexits t ns.


Lemma transl_stmt_charact:
  forall map stmt nd nexits ngoto nret rret s ns s' INCR
    (TR: transl_stmt map stmt nd nexits ngoto nret rret s = OK ns s' INCR)
    (WF: map_valid map s)
    (OK: return_reg_ok s map rret),
  tr_stmt s'.(st_code) map stmt ns nd nexits ngoto nret rret.


Lemma transl_function_charact:
  forall f tf,
  transl_function f = Errors.OK tf ->
  tr_function f tf.