Library Alloctyping

Preservation of typing during register allocation.

Require Import Coqlib.
Require Import Maps.
Require Import Errors.
Require Import AST.
Require Import Op.
Require Import Registers.
Require Import RTL.
Require Import Locations.
Require Import LTL.
Require Import Coloring.
Require Import Coloringproof.
Require Import Allocation.
Require Import Allocproof.
Require Import RTLtyping.
Require Import LTLtyping.
Require Import Conventions.

This file proves that register allocation (the translation from RTL to LTL defined in file Allocation) preserves typing: given a well-typed RTL input, it produces LTL code that is well-typed.

Section TYPING_FUNCTION.

Variable f: RTL.function.
Variable env: regenv.
Variable live: PMap.t Regset.t.
Variable alloc: reg -> loc.
Variable tf: LTL.function.

Hypothesis TYPE_RTL: type_function f = OK env.
Hypothesis LIVE: analyze f = Some live.
Hypothesis ALLOC: regalloc f live (live0 f live) env = Some alloc.
Hypothesis TRANSL: transf_function f = OK tf.

Lemma wt_rtl_function: RTLtyping.wt_function f env.


Lemma alloc_type: forall r, Loc.type (alloc r) = env r.


Lemma alloc_types:
  forall rl, List.map Loc.type (List.map alloc rl) = List.map env rl.


Lemma alloc_acceptable:
  forall r, loc_acceptable (alloc r).


Lemma allocs_acceptable:
  forall rl, locs_acceptable (List.map alloc rl).


Remark transf_unroll:
  tf = transf_fun f live alloc.


Lemma valid_successor_transf:
  forall s,
  RTLtyping.valid_successor f s ->
  LTLtyping.valid_successor tf s.


Hint Resolve alloc_acceptable allocs_acceptable: allocty.
Hint Rewrite alloc_type alloc_types: allocty.
Hint Resolve valid_successor_transf: allocty.

Type preservation during translation from RTL to LTL


Ltac WT :=
  constructor; auto with allocty; autorewrite with allocty; auto.

Lemma wt_transf_instr:
  forall pc instr,
  RTLtyping.wt_instr env f instr ->
  f.(RTL.fn_code)!pc = Some instr ->
  wt_instr tf (transf_instr f live alloc pc instr).


End TYPING_FUNCTION.

Lemma wt_transf_function:
  forall f tf,
  transf_function f = OK tf -> wt_function tf.


Lemma wt_transf_fundef:
  forall f tf,
  transf_fundef f = OK tf -> wt_fundef tf.


Lemma program_typing_preserved:
  forall (p: RTL.program) (tp: LTL.program),
  transf_program p = OK tp ->
  LTLtyping.wt_program tp.