summaryrefslogtreecommitdiff
path: root/backend/Alloctyping.v
blob: c0abf0d6d3319f04926ada588f450f6954c766b1 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
(** 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.
Proof.
  apply type_function_correct; auto.
Qed.

Lemma alloc_type: forall r, Loc.type (alloc r) = env r.
Proof.
  intro. eapply regalloc_preserves_types; eauto.
Qed.

Lemma alloc_types:
  forall rl, List.map Loc.type (List.map alloc rl) = List.map env rl.
Proof.
  intros. rewrite list_map_compose. apply list_map_exten.
  intros. symmetry. apply alloc_type. 
Qed.

Lemma alloc_acceptable: 
  forall r, loc_acceptable (alloc r).
Proof.
  intros. eapply regalloc_acceptable; eauto.
Qed.

Lemma allocs_acceptable:
  forall rl, locs_acceptable (List.map alloc rl).
Proof.
  intros. eapply regsalloc_acceptable; eauto.
Qed.

Remark transf_unroll:
  tf = transf_fun f live alloc.
Proof.
  generalize TRANSL. unfold transf_function.
  rewrite TYPE_RTL. rewrite LIVE. rewrite ALLOC. congruence.
Qed.

Lemma valid_successor_transf:
  forall s,
  RTLtyping.valid_successor f s ->
  LTLtyping.valid_successor tf s.
Proof.
  unfold RTLtyping.valid_successor, LTLtyping.valid_successor.
  intros s [i AT].
  rewrite transf_unroll; simpl. rewrite PTree.gmap. 
  rewrite AT. exists (transf_instr f live alloc s i). auto.
Qed.  

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 ->
  wt_instr tf (transf_instr f live alloc pc instr).
Proof.
  intros. inv H; simpl.
  (* nop *)
  WT.
  (* move *)
  destruct (Regset.mem r live!!pc).
  destruct (is_redundant_move Omove (r1 :: nil) r alloc); WT.
  WT.
  (* other ops *)
  destruct (Regset.mem res live!!pc).
  destruct (is_redundant_move op args res alloc); WT.
  WT.
  (* load *)
  destruct (Regset.mem dst live!!pc); WT.
  (* store *)
  WT.
  (* call *)
  WT.
  destruct ros; simpl. autorewrite with allocty; auto. auto.
  destruct ros; simpl; auto with allocty.
  (* tailcall *)
  WT.
  destruct ros; simpl. autorewrite with allocty; auto. auto.
  destruct ros; simpl; auto with allocty.
  rewrite transf_unroll; auto.
  (* alloc *)
  WT.
  (* cond *)
  WT.
  (* return *)
  WT.
  rewrite transf_unroll; simpl. 
  destruct optres; simpl. autorewrite with allocty. auto. auto.
  destruct optres; simpl; auto with allocty.
Qed.

End TYPING_FUNCTION.

Lemma wt_transf_function:
  forall f tf,
  transf_function f = OK tf -> wt_function tf.
Proof.
  intros. generalize H; unfold transf_function.
  caseEq (type_function f). intros env TYP. 
  caseEq (analyze f). intros live ANL.
  change (transfer f (RTL.fn_entrypoint f)
                     live!!(RTL.fn_entrypoint f))
    with (live0 f live).
  caseEq (regalloc f live (live0 f live) env).
  intros alloc ALLOC.
  intro EQ; injection EQ; intro.
  assert (RTLtyping.wt_function f env). apply type_function_correct; auto.
  inversion H1.
  constructor; rewrite <- H0; simpl.
  rewrite (alloc_types _ _ _ _ ALLOC). auto.
  eapply regsalloc_acceptable; eauto.
  eapply regalloc_norepet_norepet; eauto.
  eapply regalloc_correct_2; eauto.
  intros until instr. rewrite PTree.gmap. 
  caseEq (RTL.fn_code f)!pc; simpl; intros.
  inversion H3. eapply wt_transf_instr; eauto. congruence. discriminate.
  eapply valid_successor_transf; eauto. congruence. 
  congruence. congruence. congruence.
Qed.

Lemma wt_transf_fundef:
  forall f tf,
  transf_fundef f = OK tf -> wt_fundef tf.
Proof.
  intros until tf; destruct f; simpl. 
  caseEq (transf_function f); simpl. intros g TF EQ. inversion EQ.
  constructor. eapply wt_transf_function; eauto.
  congruence.
  intros. inversion H. constructor. 
Qed.

Lemma program_typing_preserved:
  forall (p: RTL.program) (tp: LTL.program),
  transf_program p = OK tp ->
  LTLtyping.wt_program tp.
Proof.
  intros; red; intros.
  generalize (transform_partial_program_function transf_fundef p i f H H0).
  intros [f0 [IN TRANSF]].
  apply wt_transf_fundef with f0; auto.
Qed.