summaryrefslogtreecommitdiff
path: root/backend
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-04-20 07:54:52 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-04-20 07:54:52 +0000
commit255cee09b71255051c2b40eae0c88bffce1f6f32 (patch)
tree7951b1b13e8fd5e525b9223e8be0580e83550f55 /backend
parent6e5041958df01c56762e90770abd704b95a36e5d (diff)
Big merge of the newregalloc-int64 branch. Lots of changes in two directions:
1- new register allocator (+ live range splitting, spilling&reloading, etc) based on a posteriori validation using the Rideau-Leroy algorithm 2- support for 64-bit integer arithmetic (type "long long"). git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2200 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend')
-rw-r--r--backend/Allocation.v1214
-rw-r--r--backend/Allocproof.v2502
-rw-r--r--backend/Alloctyping.v205
-rw-r--r--backend/Asmgenproof0.v134
-rw-r--r--backend/Bounds.v160
-rw-r--r--backend/CMtypecheck.ml84
-rw-r--r--backend/CSEproof.v10
-rw-r--r--backend/CleanupLabels.v9
-rw-r--r--backend/CleanupLabelsproof.v52
-rw-r--r--backend/CleanupLabelstyping.v59
-rw-r--r--backend/Cminor.v66
-rw-r--r--backend/CminorSel.v15
-rw-r--r--backend/Constprop.v5
-rw-r--r--backend/Constpropproof.v5
-rw-r--r--backend/Conventions.v208
-rw-r--r--backend/IRC.ml894
-rw-r--r--backend/IRC.mli44
-rw-r--r--backend/Inlining.v18
-rw-r--r--backend/Inliningspec.v33
-rw-r--r--backend/LTL.v349
-rw-r--r--backend/LTLin.v268
-rw-r--r--backend/LTLintyping.v122
-rw-r--r--backend/LTLtyping.v143
-rw-r--r--backend/Linear.v193
-rw-r--r--backend/Linearize.v75
-rw-r--r--backend/Linearizeaux.ml28
-rw-r--r--backend/Linearizeproof.v462
-rw-r--r--backend/Linearizetyping.v112
-rw-r--r--backend/Lineartyping.v224
-rw-r--r--backend/Locations.v401
-rw-r--r--backend/Mach.v98
-rw-r--r--backend/Machtyping.v108
-rw-r--r--backend/PrintCminor.ml42
-rw-r--r--backend/PrintLTL.ml159
-rw-r--r--backend/PrintMach.ml2
-rw-r--r--backend/PrintXTL.ml147
-rw-r--r--backend/RRE.v173
-rw-r--r--backend/RREproof.v658
-rw-r--r--backend/RREtyping.v110
-rw-r--r--backend/RTLgen.v41
-rw-r--r--backend/RTLgenaux.ml2
-rw-r--r--backend/RTLgenproof.v210
-rw-r--r--backend/RTLgenspec.v92
-rw-r--r--backend/RTLtyping.v45
-rw-r--r--backend/Regalloc.ml986
-rw-r--r--backend/Reload.v274
-rw-r--r--backend/Reloadproof.v1487
-rw-r--r--backend/Reloadtyping.v353
-rw-r--r--backend/SelectLong.vp368
-rw-r--r--backend/SelectLongproof.v1063
-rw-r--r--backend/Selection.v53
-rw-r--r--backend/Selectionproof.v252
-rw-r--r--backend/Splitting.ml184
-rw-r--r--backend/Stacking.v32
-rw-r--r--backend/Stackingproof.v819
-rw-r--r--backend/Tunneling.v40
-rw-r--r--backend/Tunnelingproof.v194
-rw-r--r--backend/Tunnelingtyping.v103
-rw-r--r--backend/XTL.ml213
-rw-r--r--backend/XTL.mli100
60 files changed, 9462 insertions, 7040 deletions
diff --git a/backend/Allocation.v b/backend/Allocation.v
index caaf09d..8674335 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -10,107 +10,1169 @@
(* *)
(* *********************************************************************)
-(** Register allocation. *)
+(** Register allocation by external oracle and a posteriori validation. *)
+Require Import FSets.
+Require FSetAVLplus.
Require Import Coqlib.
+Require Import Ordered.
Require Import Errors.
Require Import Maps.
Require Import Lattice.
Require Import AST.
+Require Import Integers.
+Require Import Memdata.
Require Import Op.
Require Import Registers.
Require Import RTL.
-Require Import RTLtyping.
-Require Import Liveness.
+Require Import Kildall.
Require Import Locations.
+Require Import Conventions.
+Require Import RTLtyping.
Require Import LTL.
-Require Import Coloring.
-
-(** * Translation from RTL to LTL *)
-
-(** Each [RTL] instruction translates to an [LTL] instruction.
- The register assignment [assign] returned by register allocation
- is applied to the arguments and results of the RTL
- instruction. Moreover, dead instructions and redundant moves
- are eliminated (turned into a [Lnop] instruction).
- Dead instructions are instructions without side-effects ([Iop] and
- [Iload]) whose result register is dead, i.e. whose result value
- is never used. Redundant moves are moves whose source and destination
- are assigned the same location. *)
-
-Definition is_redundant_move
- (op: operation) (args: list reg) (res: reg) (assign: reg -> loc) : bool :=
- match is_move_operation op args with
- | None => false
- | Some src => if Loc.eq (assign src) (assign res) then true else false
- end.
-
-Definition transf_instr
- (f: RTL.function) (live: PMap.t Regset.t) (assign: reg -> loc)
- (pc: node) (instr: RTL.instruction) : LTL.instruction :=
- match instr with
+
+(** The validation algorithm used here is described in
+ "Validating register allocation and spilling",
+ by Silvain Rideau and Xavier Leroy,
+ in Compiler Construction (CC 2010), LNCS 6011, Springer, 2010. *)
+
+(** * Structural checks *)
+
+(** As a first pass, we check the LTL code returned by the external oracle
+ against the original RTL code for structural conformance.
+ Each RTL instruction was transformed into a LTL basic block whose
+ shape must agree with the RTL instruction. For example, if the RTL
+ instruction is [Istore(Mint32, addr, args, src, s)], the LTL basic block
+ must be of the following shape:
+- zero, one or several "move" instructions
+- a store instruction [Lstore(Mint32, addr, args', src')]
+- a [Lbranch s] instruction.
+
+ The [block_shape] type below describes all possible cases of structural
+ maching between an RTL instruction and an LTL basic block.
+*)
+
+Definition moves := list (loc * loc)%type.
+
+Inductive block_shape: Type :=
+ | BSnop (mv: moves) (s: node)
+ | BSmove (src: reg) (dst: reg) (mv: moves) (s: node)
+ | BSmakelong (src1 src2: reg) (dst: reg) (mv: moves) (s: node)
+ | BSlowlong (src: reg) (dst: reg) (mv: moves) (s: node)
+ | BShighlong (src: reg) (dst: reg) (mv: moves) (s: node)
+ | BSop (op: operation) (args: list reg) (res: reg)
+ (mv1: moves) (args': list mreg) (res': mreg)
+ (mv2: moves) (s: node)
+ | BSopdead (op: operation) (args: list reg) (res: reg)
+ (mv: moves) (s: node)
+ | BSload (chunk: memory_chunk) (addr: addressing) (args: list reg) (dst: reg)
+ (mv1: moves) (args': list mreg) (dst': mreg)
+ (mv2: moves) (s: node)
+ | BSloaddead (chunk: memory_chunk) (addr: addressing) (args: list reg) (dst: reg)
+ (mv: moves) (s: node)
+ | BSload2 (addr1 addr2: addressing) (args: list reg) (dst: reg)
+ (mv1: moves) (args1': list mreg) (dst1': mreg)
+ (mv2: moves) (args2': list mreg) (dst2': mreg)
+ (mv3: moves) (s: node)
+ | BSload2_1 (addr: addressing) (args: list reg) (dst: reg)
+ (mv1: moves) (args': list mreg) (dst': mreg)
+ (mv2: moves) (s: node)
+ | BSload2_2 (addr addr': addressing) (args: list reg) (dst: reg)
+ (mv1: moves) (args': list mreg) (dst': mreg)
+ (mv2: moves) (s: node)
+ | BSstore (chunk: memory_chunk) (addr: addressing) (args: list reg) (src: reg)
+ (mv1: moves) (args': list mreg) (src': mreg)
+ (s: node)
+ | BSstore2 (addr1 addr2: addressing) (args: list reg) (src: reg)
+ (mv1: moves) (args1': list mreg) (src1': mreg)
+ (mv2: moves) (args2': list mreg) (src2': mreg)
+ (s: node)
+ | BScall (sg: signature) (ros: reg + ident) (args: list reg) (res: reg)
+ (mv1: moves) (ros': mreg + ident) (mv2: moves) (s: node)
+ | BStailcall (sg: signature) (ros: reg + ident) (args: list reg)
+ (mv1: moves) (ros': mreg + ident)
+ | BSbuiltin (ef: external_function) (args: list reg) (res: reg)
+ (mv1: moves) (args': list mreg) (res': list mreg)
+ (mv2: moves) (s: node)
+ | BSannot (text: ident) (targs: list annot_arg) (args: list reg) (res: reg)
+ (mv: moves) (args': list loc) (s: node)
+ | BScond (cond: condition) (args: list reg)
+ (mv: moves) (args': list mreg) (s1 s2: node)
+ | BSjumptable (arg: reg)
+ (mv: moves) (arg': mreg) (tbl: list node)
+ | BSreturn (arg: option reg)
+ (mv: moves).
+
+(** Extract the move instructions at the beginning of block [b].
+ Return the list of moves and the suffix of [b] after the moves. *)
+
+Fixpoint extract_moves (accu: moves) (b: bblock) {struct b} : moves * bblock :=
+ match b with
+ | Lgetstack sl ofs ty dst :: b' =>
+ extract_moves ((S sl ofs ty, R dst) :: accu) b'
+ | Lsetstack src sl ofs ty :: b' =>
+ extract_moves ((R src, S sl ofs ty) :: accu) b'
+ | Lop op args res :: b' =>
+ match is_move_operation op args with
+ | Some arg => extract_moves ((R arg, R res) :: accu) b'
+ | None => (List.rev accu, b)
+ end
+ | _ =>
+ (List.rev accu, b)
+ end.
+
+Definition check_succ (s: node) (b: LTL.bblock) : bool :=
+ match b with
+ | Lbranch s' :: _ => peq s s'
+ | _ => false
+ end.
+
+Parameter eq_operation: forall (x y: operation), {x=y} + {x<>y}.
+Parameter eq_addressing: forall (x y: addressing), {x=y} + {x<>y}.
+Parameter eq_opt_addressing: forall (x y: option addressing), {x=y} + {x<>y}.
+Parameter eq_condition: forall (x y: condition), {x=y} + {x<>y}.
+Parameter eq_chunk: forall (x y: memory_chunk), {x=y} + {x<>y}.
+Parameter eq_external_function: forall (x y: external_function), {x=y} + {x<>y}.
+Parameter eq_signature: forall (x y: signature), {x=y} + {x<>y}.
+
+Notation "'do' X <- A ; B" := (match A with Some X => B | None => None end)
+ (at level 200, X ident, A at level 100, B at level 200)
+ : option_monad_scope.
+
+Notation "'assertion' A ; B" := (if A then B else None)
+ (at level 200, A at level 100, B at level 200)
+ : option_monad_scope.
+
+Local Open Scope option_monad_scope.
+
+(** Classify operations into moves, 64-bit integer operations, and other
+ arithmetic/logical operations. *)
+
+Inductive operation_kind: operation -> list reg -> Type :=
+ | operation_Omove: forall arg, operation_kind Omove (arg :: nil)
+ | operation_Omakelong: forall arg1 arg2, operation_kind Omakelong (arg1 :: arg2 :: nil)
+ | operation_Olowlong: forall arg, operation_kind Olowlong (arg :: nil)
+ | operation_Ohighlong: forall arg, operation_kind Ohighlong (arg :: nil)
+ | operation_other: forall op args, operation_kind op args.
+
+Definition classify_operation (op: operation) (args: list reg) : operation_kind op args :=
+ match op, args with
+ | Omove, arg::nil => operation_Omove arg
+ | Omakelong, arg1::arg2::nil => operation_Omakelong arg1 arg2
+ | Olowlong, arg::nil => operation_Olowlong arg
+ | Ohighlong, arg::nil => operation_Ohighlong arg
+ | op, args => operation_other op args
+ end.
+
+(** Check RTL instruction [i] against LTL basic block [b].
+ On success, return [Some] with a [block_shape] describing the correspondence.
+ On error, return [None]. *)
+
+Definition pair_instr_block
+ (i: RTL.instruction) (b: LTL.bblock) : option block_shape :=
+ match i with
| Inop s =>
- Lnop s
+ let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BSnop mv s)
| Iop op args res s =>
- if Regset.mem res live!!pc then
- if is_redundant_move op args res assign then
- Lnop s
- else
- Lop op (List.map assign args) (assign res) s
- else
- Lnop s
+ match classify_operation op args with
+ | operation_Omove arg =>
+ let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BSmove arg res mv s)
+ | operation_Omakelong arg1 arg2 =>
+ let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BSmakelong arg1 arg2 res mv s)
+ | operation_Olowlong arg =>
+ let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BSlowlong arg res mv s)
+ | operation_Ohighlong arg =>
+ let (mv, b1) := extract_moves nil b in
+ assertion (check_succ s b1); Some(BShighlong arg res mv s)
+ | operation_other _ _ =>
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lop op' args' res' :: b2 =>
+ let (mv2, b3) := extract_moves nil b2 in
+ assertion (eq_operation op op');
+ assertion (check_succ s b3);
+ Some(BSop op args res mv1 args' res' mv2 s)
+ | _ =>
+ assertion (check_succ s b1);
+ Some(BSopdead op args res mv1 s)
+ end
+ end
| Iload chunk addr args dst s =>
- if Regset.mem dst live!!pc then
- Lload chunk addr (List.map assign args) (assign dst) s
- else
- Lnop s
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lload chunk' addr' args' dst' :: b2 =>
+ if eq_chunk chunk Mint64 then
+ assertion (eq_chunk chunk' Mint32);
+ let (mv2, b3) := extract_moves nil b2 in
+ match b3 with
+ | Lload chunk'' addr'' args'' dst'' :: b4 =>
+ let (mv3, b5) := extract_moves nil b4 in
+ assertion (eq_chunk chunk'' Mint32);
+ assertion (eq_addressing addr addr');
+ assertion (eq_opt_addressing (offset_addressing addr (Int.repr 4)) (Some addr''));
+ assertion (check_succ s b5);
+ Some(BSload2 addr addr'' args dst mv1 args' dst' mv2 args'' dst'' mv3 s)
+ | _ =>
+ assertion (check_succ s b3);
+ if (eq_addressing addr addr') then
+ Some(BSload2_1 addr args dst mv1 args' dst' mv2 s)
+ else
+ (assertion (eq_opt_addressing (offset_addressing addr (Int.repr 4)) (Some addr'));
+ Some(BSload2_2 addr addr' args dst mv1 args' dst' mv2 s))
+ end
+ else (
+ let (mv2, b3) := extract_moves nil b2 in
+ assertion (eq_chunk chunk chunk');
+ assertion (eq_addressing addr addr');
+ assertion (check_succ s b3);
+ Some(BSload chunk addr args dst mv1 args' dst' mv2 s))
+ | _ =>
+ assertion (check_succ s b1);
+ Some(BSloaddead chunk addr args dst mv1 s)
+ end
| Istore chunk addr args src s =>
- Lstore chunk addr (List.map assign args) (assign src) s
- | Icall sig ros args res s =>
- Lcall sig (sum_left_map assign ros) (List.map assign args)
- (assign res) s
- | Itailcall sig ros args =>
- Ltailcall sig (sum_left_map assign ros) (List.map assign args)
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lstore chunk' addr' args' src' :: b2 =>
+ if eq_chunk chunk Mint64 then
+ let (mv2, b3) := extract_moves nil b2 in
+ match b3 with
+ | Lstore chunk'' addr'' args'' src'' :: b4 =>
+ assertion (eq_chunk chunk' Mint32);
+ assertion (eq_chunk chunk'' Mint32);
+ assertion (eq_addressing addr addr');
+ assertion (eq_opt_addressing (offset_addressing addr (Int.repr 4)) (Some addr''));
+ assertion (check_succ s b4);
+ Some(BSstore2 addr addr'' args src mv1 args' src' mv2 args'' src'' s)
+ | _ => None
+ end
+ else (
+ assertion (eq_chunk chunk chunk');
+ assertion (eq_addressing addr addr');
+ assertion (check_succ s b2);
+ Some(BSstore chunk addr args src mv1 args' src' s))
+ | _ => None
+ end
+ | Icall sg ros args res s =>
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lcall sg' ros' :: b2 =>
+ let (mv2, b3) := extract_moves nil b2 in
+ assertion (eq_signature sg sg');
+ assertion (check_succ s b3);
+ Some(BScall sg ros args res mv1 ros' mv2 s)
+ | _ => None
+ end
+ | Itailcall sg ros args =>
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Ltailcall sg' ros' :: b2 =>
+ assertion (eq_signature sg sg');
+ Some(BStailcall sg ros args mv1 ros')
+ | _ => None
+ end
| Ibuiltin ef args res s =>
- Lbuiltin ef (List.map assign args) (assign res) s
- | Icond cond args ifso ifnot =>
- Lcond cond (List.map assign args) ifso ifnot
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lbuiltin ef' args' res' :: b2 =>
+ let (mv2, b3) := extract_moves nil b2 in
+ assertion (eq_external_function ef ef');
+ assertion (check_succ s b3);
+ Some(BSbuiltin ef args res mv1 args' res' mv2 s)
+ | Lannot ef' args' :: b2 =>
+ assertion (eq_external_function ef ef');
+ assertion (check_succ s b2);
+ match ef with
+ | EF_annot txt typ => Some(BSannot txt typ args res mv1 args' s)
+ | _ => None
+ end
+ | _ => None
+ end
+ | Icond cond args s1 s2 =>
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lcond cond' args' s1' s2' :: b2 =>
+ assertion (eq_condition cond cond');
+ assertion (peq s1 s1');
+ assertion (peq s2 s2');
+ Some(BScond cond args mv1 args' s1 s2)
+ | _ => None
+ end
| Ijumptable arg tbl =>
- Ljumptable (assign arg) tbl
- | Ireturn optarg =>
- Lreturn (option_map assign optarg)
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Ljumptable arg' tbl' :: b2 =>
+ assertion (list_eq_dec peq tbl tbl');
+ Some(BSjumptable arg mv1 arg' tbl)
+ | _ => None
+ end
+ | Ireturn arg =>
+ let (mv1, b1) := extract_moves nil b in
+ match b1 with
+ | Lreturn :: b2 => Some(BSreturn arg mv1)
+ | _ => None
+ end
+ end.
+
+(** Check all instructions of the RTL function [f1] against the corresponding
+ basic blocks of LTL function [f2]. Return a map from CFG nodes to
+ [block_shape] info. *)
+
+Definition pair_codes (f1: RTL.function) (f2: LTL.function) : PTree.t block_shape :=
+ PTree.combine
+ (fun opti optb => do i <- opti; do b <- optb; pair_instr_block i b)
+ (RTL.fn_code f1) (LTL.fn_code f2).
+
+(** Check the entry point code of the LTL function [f2]. It must be
+ a sequence of moves that branches to the same node as the entry point
+ of RTL function [f1]. *)
+
+Definition pair_entrypoints (f1: RTL.function) (f2: LTL.function) : option moves :=
+ do b <- (LTL.fn_code f2)!(LTL.fn_entrypoint f2);
+ let (mv, b1) := extract_moves nil b in
+ assertion (check_succ (RTL.fn_entrypoint f1) b1);
+ Some mv.
+
+(** * Representing sets of equations between RTL registers and LTL locations. *)
+
+(** The Rideau-Leroy validation algorithm manipulates sets of equations of
+ the form [pseudoreg = location [kind]], meaning:
+- if [kind = Full], the value of [location] in the generated LTL code is
+ the same as (or more defined than) the value of [pseudoreg] in the original
+ RTL code;
+- if [kind = Low], the value of [location] in the generated LTL code is
+ the same as (or more defined than) the low 32 bits of the 64-bit
+ integer value of [pseudoreg] in the original RTL code;
+- if [kind = High], the value of [location] in the generated LTL code is
+ the same as (or more defined than) the high 32 bits of the 64-bit
+ integer value of [pseudoreg] in the original RTL code.
+*)
+
+Inductive equation_kind : Type := Full | Low | High.
+
+Record equation := Eq {
+ ekind: equation_kind;
+ ereg: reg;
+ eloc: loc
+}.
+
+(** We use AVL finite sets to represent sets of equations. Therefore, we need
+ total orders over equations and their components. *)
+
+Module IndexedEqKind <: INDEXED_TYPE.
+ Definition t := equation_kind.
+ Definition index (x: t) :=
+ match x with Full => 1%positive | Low => 2%positive | High => 3%positive end.
+ Lemma index_inj: forall x y, index x = index y -> x = y.
+ Proof. destruct x; destruct y; simpl; congruence. Qed.
+ Definition eq (x y: t) : {x=y} + {x<>y}.
+ Proof. decide equality. Defined.
+End IndexedEqKind.
+
+Module OrderedEqKind := OrderedIndexed(IndexedEqKind).
+
+(** This is an order over equations that is lexicgraphic on [ereg], then
+ [eloc], then [ekind]. *)
+
+Module OrderedEquation <: OrderedType.
+ Definition t := equation.
+ Definition eq (x y: t) := x = y.
+ Definition lt (x y: t) :=
+ Plt (ereg x) (ereg y) \/ (ereg x = ereg y /\
+ (OrderedLoc.lt (eloc x) (eloc y) \/ (eloc x = eloc y /\
+ OrderedEqKind.lt (ekind x) (ekind y)))).
+ Lemma eq_refl : forall x : t, eq x x.
+ Proof (@refl_equal t).
+ Lemma eq_sym : forall x y : t, eq x y -> eq y x.
+ Proof (@sym_equal t).
+ Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
+ Proof (@trans_equal t).
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof.
+ unfold lt; intros.
+ destruct H.
+ destruct H0. left; eapply Plt_trans; eauto.
+ destruct H0. rewrite <- H0. auto.
+ destruct H. rewrite H.
+ destruct H0. auto.
+ destruct H0. right; split; auto.
+ intuition.
+ left; eapply OrderedLoc.lt_trans; eauto.
+ left; congruence.
+ left; congruence.
+ right; split. congruence. eapply OrderedEqKind.lt_trans; eauto.
+ Qed.
+ Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Proof.
+ unfold lt, eq; intros; red; intros. subst y. intuition.
+ eelim Plt_strict; eauto.
+ eelim OrderedLoc.lt_not_eq; eauto. red; auto.
+ eelim OrderedEqKind.lt_not_eq; eauto. red; auto.
+ Qed.
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ intros.
+ destruct (OrderedPositive.compare (ereg x) (ereg y)).
+ - apply LT. red; auto.
+ - destruct (OrderedLoc.compare (eloc x) (eloc y)).
+ + apply LT. red; auto.
+ + destruct (OrderedEqKind.compare (ekind x) (ekind y)).
+ * apply LT. red; auto.
+ * apply EQ. red in e; red in e0; red in e1; red.
+ destruct x; destruct y; simpl in *; congruence.
+ * apply GT. red; auto.
+ + apply GT. red; auto.
+ - apply GT. red; auto.
+ Defined.
+ Definition eq_dec (x y: t) : {x = y} + {x <> y}.
+ Proof.
+ intros. decide equality.
+ apply Loc.eq.
+ apply peq.
+ apply IndexedEqKind.eq.
+ Defined.
+End OrderedEquation.
+
+(** This is an alternate order over equations that is lexicgraphic on
+ [eloc], then [ereg], then [ekind]. *)
+
+Module OrderedEquation' <: OrderedType.
+ Definition t := equation.
+ Definition eq (x y: t) := x = y.
+ Definition lt (x y: t) :=
+ OrderedLoc.lt (eloc x) (eloc y) \/ (eloc x = eloc y /\
+ (Plt (ereg x) (ereg y) \/ (ereg x = ereg y /\
+ OrderedEqKind.lt (ekind x) (ekind y)))).
+ Lemma eq_refl : forall x : t, eq x x.
+ Proof (@refl_equal t).
+ Lemma eq_sym : forall x y : t, eq x y -> eq y x.
+ Proof (@sym_equal t).
+ Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
+ Proof (@trans_equal t).
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof.
+ unfold lt; intros.
+ destruct H.
+ destruct H0. left; eapply OrderedLoc.lt_trans; eauto.
+ destruct H0. rewrite <- H0. auto.
+ destruct H. rewrite H.
+ destruct H0. auto.
+ destruct H0. right; split; auto.
+ intuition.
+ left; eapply Plt_trans; eauto.
+ left; congruence.
+ left; congruence.
+ right; split. congruence. eapply OrderedEqKind.lt_trans; eauto.
+ Qed.
+ Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Proof.
+ unfold lt, eq; intros; red; intros. subst y. intuition.
+ eelim OrderedLoc.lt_not_eq; eauto. red; auto.
+ eelim Plt_strict; eauto.
+ eelim OrderedEqKind.lt_not_eq; eauto. red; auto.
+ Qed.
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ intros.
+ destruct (OrderedLoc.compare (eloc x) (eloc y)).
+ - apply LT. red; auto.
+ - destruct (OrderedPositive.compare (ereg x) (ereg y)).
+ + apply LT. red; auto.
+ + destruct (OrderedEqKind.compare (ekind x) (ekind y)).
+ * apply LT. red; auto.
+ * apply EQ. red in e; red in e0; red in e1; red.
+ destruct x; destruct y; simpl in *; congruence.
+ * apply GT. red; auto.
+ + apply GT. red; auto.
+ - apply GT. red; auto.
+ Defined.
+ Definition eq_dec: forall (x y: t), {x = y} + {x <> y} := OrderedEquation.eq_dec.
+End OrderedEquation'.
+
+Module EqSet := FSetAVLplus.Make(OrderedEquation).
+Module EqSet2 := FSetAVLplus.Make(OrderedEquation').
+
+(** We use a redundant representation for sets of equations, comprising
+ two AVL finite sets, containing the same elements, but ordered along
+ the two orders defined above. Playing on properties of lexicographic
+ orders, this redundant representation enables us to quickly find
+ all equations involving a given RTL pseudoregister, or all equations
+ involving a given LTL location or overlapping location. *)
+
+Record eqs := mkeqs {
+ eqs1 :> EqSet.t;
+ eqs2 : EqSet2.t;
+ eqs_same: forall q, EqSet2.In q eqs2 <-> EqSet.In q eqs1
+}.
+
+(** * Operations on sets of equations *)
+
+(** The empty set of equations. *)
+
+Program Definition empty_eqs := mkeqs EqSet.empty EqSet2.empty _.
+Next Obligation.
+ split; intros. eelim EqSet2.empty_1; eauto. eelim EqSet.empty_1; eauto.
+Qed.
+
+(** Adding or removing an equation from a set. *)
+
+Program Definition add_equation (q: equation) (e: eqs) :=
+ mkeqs (EqSet.add q (eqs1 e)) (EqSet2.add q (eqs2 e)) _.
+Next Obligation.
+ split; intros.
+ destruct (OrderedEquation'.eq_dec q q0).
+ apply EqSet.add_1; auto.
+ apply EqSet.add_2. apply (eqs_same e). apply EqSet2.add_3 with q; auto.
+ destruct (OrderedEquation.eq_dec q q0).
+ apply EqSet2.add_1; auto.
+ apply EqSet2.add_2. apply (eqs_same e). apply EqSet.add_3 with q; auto.
+Qed.
+
+Program Definition remove_equation (q: equation) (e: eqs) :=
+ mkeqs (EqSet.remove q (eqs1 e)) (EqSet2.remove q (eqs2 e)) _.
+Next Obligation.
+ split; intros.
+ destruct (OrderedEquation'.eq_dec q q0).
+ eelim EqSet2.remove_1; eauto.
+ apply EqSet.remove_2; auto. apply (eqs_same e). apply EqSet2.remove_3 with q; auto.
+ destruct (OrderedEquation.eq_dec q q0).
+ eelim EqSet.remove_1; eauto.
+ apply EqSet2.remove_2; auto. apply (eqs_same e). apply EqSet.remove_3 with q; auto.
+Qed.
+
+(** [reg_unconstrained r e] is true if [e] contains no equations involving
+ the RTL pseudoregister [r]. In other words, all equations [r' = l [kind]]
+ in [e] are such that [r' <> r]. *)
+
+Definition select_reg_l (r: reg) (q: equation) := Pos.leb r (ereg q).
+Definition select_reg_h (r: reg) (q: equation) := Pos.leb (ereg q) r.
+
+Definition reg_unconstrained (r: reg) (e: eqs) : bool :=
+ negb (EqSet.mem_between (select_reg_l r) (select_reg_h r) (eqs1 e)).
+
+(** [loc_unconstrained l e] is true if [e] contains no equations involving
+ the LTL location [l] or a location that partially overlaps with [l].
+ In other words, all equations [r = l' [kind]] in [e] are such that
+ [Loc.diff l' l]. *)
+
+Definition select_loc_l (l: loc) :=
+ let lb := OrderedLoc.diff_low_bound l in
+ fun (q: equation) => match OrderedLoc.compare (eloc q) lb with LT _ => false | _ => true end.
+Definition select_loc_h (l: loc) :=
+ let lh := OrderedLoc.diff_high_bound l in
+ fun (q: equation) => match OrderedLoc.compare (eloc q) lh with GT _ => false | _ => true end.
+
+Definition loc_unconstrained (l: loc) (e: eqs) : bool :=
+ negb (EqSet2.mem_between (select_loc_l l) (select_loc_h l) (eqs2 e)).
+
+Definition reg_loc_unconstrained (r: reg) (l: loc) (e: eqs) : bool :=
+ reg_unconstrained r e && loc_unconstrained l e.
+
+(** [subst_reg r1 r2 e] simulates the effect of assigning [r2] to [r1] on [e].
+ All equations of the form [r1 = l [kind]] are replaced by [r2 = l [kind]].
+*)
+
+Definition subst_reg (r1 r2: reg) (e: eqs) : eqs :=
+ EqSet.fold
+ (fun q e => add_equation (Eq (ekind q) r2 (eloc q)) (remove_equation q e))
+ (EqSet.elements_between (select_reg_l r1) (select_reg_h r1) (eqs1 e))
+ e.
+
+(** [subst_reg_kind r1 k1 r2 k2 e] simulates the effect of assigning
+ the [k2] part of [r2] to the [k1] part of [r1] on [e].
+ All equations of the form [r1 = l [k1]] are replaced by [r2 = l [k2]].
+*)
+
+Definition subst_reg_kind (r1: reg) (k1: equation_kind) (r2: reg) (k2: equation_kind) (e: eqs) : eqs :=
+ EqSet.fold
+ (fun q e =>
+ if IndexedEqKind.eq (ekind q) k1
+ then add_equation (Eq k2 r2 (eloc q)) (remove_equation q e)
+ else e)
+ (EqSet.elements_between (select_reg_l r1) (select_reg_h r1) (eqs1 e))
+ e.
+
+(** [subst_loc l1 l2 e] simulates the effect of assigning [l2] to [l1] on [e].
+ All equations of the form [r = l1 [kind]] are replaced by [r = l2 [kind]].
+ Return [None] if [e] contains an equation of the form [r = l] with [l]
+ partially overlapping [l1].
+*)
+
+Definition subst_loc (l1 l2: loc) (e: eqs) : option eqs :=
+ EqSet2.fold
+ (fun q opte =>
+ match opte with
+ | None => None
+ | Some e =>
+ if Loc.eq l1 (eloc q) then
+ Some (add_equation (Eq (ekind q) (ereg q) l2) (remove_equation q e))
+ else
+ None
+ end)
+ (EqSet2.elements_between (select_loc_l l1) (select_loc_h l1) (eqs2 e))
+ (Some e).
+
+(** [add_equations [r1...rN] [m1...mN] e] adds to [e] the [N] equations
+ [ri = R mi [Full]]. Return [None] if the two lists have different lengths.
+*)
+
+Fixpoint add_equations (rl: list reg) (ml: list mreg) (e: eqs) : option eqs :=
+ match rl, ml with
+ | nil, nil => Some e
+ | r1 :: rl, m1 :: ml => add_equations rl ml (add_equation (Eq Full r1 (R m1)) e)
+ | _, _ => None
+ end.
+
+(** [add_equations_args] is similar but additionally handles the splitting
+ of pseudoregisters of type [Tlong] in two locations containing the
+ two 32-bit halves of the 64-bit integer. *)
+
+Function add_equations_args (rl: list reg) (tyl: list typ) (ll: list loc) (e: eqs) : option eqs :=
+ match rl, tyl, ll with
+ | nil, nil, nil => Some e
+ | r1 :: rl, Tlong :: tyl, l1 :: l2 :: ll =>
+ add_equations_args rl tyl ll (add_equation (Eq Low r1 l2) (add_equation (Eq High r1 l1) e))
+ | r1 :: rl, (Tint|Tfloat) :: tyl, l1 :: ll =>
+ add_equations_args rl tyl ll (add_equation (Eq Full r1 l1) e)
+ | _, _, _ => None
+ end.
+
+(** [add_equations_res] is similar but is specialized to the case where
+ there is only one pseudo-register. *)
+
+Function add_equations_res (r: reg) (oty: option typ) (ll: list loc) (e: eqs) : option eqs :=
+ match oty with
+ | Some Tlong =>
+ match ll with
+ | l1 :: l2 :: nil => Some (add_equation (Eq Low r l2) (add_equation (Eq High r l1) e))
+ | _ => None
+ end
+ | _ =>
+ match ll with
+ | l1 :: nil => Some (add_equation (Eq Full r l1) e)
+ | _ => None
+ end
end.
-Definition transf_fun (f: RTL.function) (live: PMap.t Regset.t)
- (assign: reg -> loc) : LTL.function :=
- LTL.mkfunction
- (RTL.fn_sig f)
- (List.map assign (RTL.fn_params f))
- (RTL.fn_stacksize f)
- (PTree.map (transf_instr f live assign) (RTL.fn_code f))
- (RTL.fn_entrypoint f).
+(** [remove_equations_res] is similar to [add_equations_res] but removes
+ equations instead of adding them. *)
+
+Function remove_equations_res (r: reg) (oty: option typ) (ll: list loc) (e: eqs) : option eqs :=
+ match oty with
+ | Some Tlong =>
+ match ll with
+ | l1 :: l2 :: nil =>
+ if Loc.diff_dec l2 l1
+ then Some (remove_equation (Eq Low r l2) (remove_equation (Eq High r l1) e))
+ else None
+ | _ => None
+ end
+ | _ =>
+ match ll with
+ | l1 :: nil => Some (remove_equation (Eq Full r l1) e)
+ | _ => None
+ end
+ end.
+
+(** [add_equations_ros] adds an equation, if needed, between an optional
+ pseudoregister and an optional machine register. It is used for the
+ function argument of the [Icall] and [Itailcall] instructions. *)
+
+Definition add_equation_ros (ros: reg + ident) (ros': mreg + ident) (e: eqs) : option eqs :=
+ match ros, ros' with
+ | inl r, inl mr => Some(add_equation (Eq Full r (R mr)) e)
+ | inr id, inr id' => assertion (ident_eq id id'); Some e
+ | _, _ => None
+ end.
+
+(** [can_undef ml] returns true if all machine registers in [ml] are
+ unconstrained and can harmlessly be undefined. *)
+
+Fixpoint can_undef (ml: list mreg) (e: eqs) : bool :=
+ match ml with
+ | nil => true
+ | m1 :: ml => loc_unconstrained (R m1) e && can_undef ml e
+ end.
+
+Fixpoint can_undef_except (l: loc) (ml: list mreg) (e: eqs) : bool :=
+ match ml with
+ | nil => true
+ | m1 :: ml =>
+ (Loc.eq l (R m1) || loc_unconstrained (R m1) e) && can_undef_except l ml e
+ end.
+
+(** [no_caller_saves e] returns [e] if all caller-save locations are
+ unconstrained in [e]. In other words, [e] contains no equations
+ involving a caller-save register or [Outgoing] stack slot. *)
+
+Definition no_caller_saves (e: eqs) : bool :=
+ EqSet.for_all
+ (fun eq =>
+ match eloc eq with
+ | R r =>
+ zle 0 (index_int_callee_save r) || zle 0 (index_float_callee_save r)
+ | S Outgoing _ _ => false
+ | S _ _ _ => true
+ end)
+ e.
+
+(** [compat_left r l e] returns true if all equations in [e] that involve
+ [r] are of the form [r = l [Full]]. *)
+
+Definition compat_left (r: reg) (l: loc) (e: eqs) : bool :=
+ EqSet.for_all_between
+ (fun q =>
+ match ekind q with
+ | Full => Loc.eq l (eloc q)
+ | _ => false
+ end)
+ (select_reg_l r) (select_reg_h r)
+ (eqs1 e).
+
+(** [compat_left2 r l1 l2 e] returns true if all equations in [e] that involve
+ [r] are of the form [r = l1 [High]] or [r = l2 [Low]]. *)
+
+Definition compat_left2 (r: reg) (l1 l2: loc) (e: eqs) : bool :=
+ EqSet.for_all_between
+ (fun q =>
+ match ekind q with
+ | High => Loc.eq l1 (eloc q)
+ | Low => Loc.eq l2 (eloc q)
+ | _ => false
+ end)
+ (select_reg_l r) (select_reg_h r)
+ (eqs1 e).
+
+(** [ros_compatible_tailcall ros] returns true if [ros] is a function
+ name or a caller-save register. This is used to check [Itailcall]
+ instructions. *)
+
+Definition ros_compatible_tailcall (ros: mreg + ident) : bool :=
+ match ros with
+ | inl r => In_dec mreg_eq r destroyed_at_call
+ | inr id => true
+ end.
+
+(** * The validator *)
+
+Definition destroyed_by_move (src dst: loc) :=
+ match src with
+ | S sl ofs ty => destroyed_by_getstack sl
+ | _ => destroyed_by_op Omove
+ end.
+
+(** Simulate the effect of a sequence of moves [mv] on a set of
+ equations [e]. The set [e] is the equations that must hold
+ after the sequence of moves. Return the set of equations that
+ must hold before the sequence of moves. Return [None] if the
+ set of equations [e] cannot hold after the sequence of moves. *)
+
+Fixpoint track_moves (mv: moves) (e: eqs) : option eqs :=
+ match mv with
+ | nil => Some e
+ | (src, dst) :: mv =>
+ do e1 <- track_moves mv e;
+ do e2 <- subst_loc dst src e1;
+ assertion (can_undef_except dst (destroyed_by_move src dst)) e1;
+ Some e2
+ end.
+
+(** [transfer_use_def args res args' res' undefs e] returns the set
+ of equations that must hold "before" in order for the equations [e]
+ to hold "after" the execution of RTL and LTL code of the following form:
+<<
+ RTL LTL
+ use pseudoregs args use machine registers args'
+ define pseudoreg res undefine machine registers undef
+ define machine register res'
+>>
+ As usual, [None] is returned if the equations [e] cannot hold after
+ this execution.
+*)
+
+Definition transfer_use_def (args: list reg) (res: reg) (args': list mreg) (res': mreg)
+ (undefs: list mreg) (e: eqs) : option eqs :=
+ let e1 := remove_equation (Eq Full res (R res')) e in
+ assertion (reg_loc_unconstrained res (R res') e1);
+ assertion (can_undef undefs e1);
+ add_equations args args' e1.
+
+Definition kind_first_word := if big_endian then High else Low.
+Definition kind_second_word := if big_endian then Low else High.
+
+(** The core transfer function. It takes a set [e] of equations that must
+ hold "after" and a block shape [shape] representing a matching pair
+ of an RTL instruction and an LTL basic block. It returns the set of
+ equations that must hold "before" these instructions, or [None] if
+ impossible. *)
+
+Definition transfer_aux (f: RTL.function) (shape: block_shape) (e: eqs) : option eqs :=
+ match shape with
+ | BSnop mv s =>
+ track_moves mv e
+ | BSmove src dst mv s =>
+ track_moves mv (subst_reg dst src e)
+ | BSmakelong src1 src2 dst mv s =>
+ let e1 := subst_reg_kind dst High src1 Full e in
+ let e2 := subst_reg_kind dst Low src2 Full e1 in
+ assertion (reg_unconstrained dst e2);
+ track_moves mv e2
+ | BSlowlong src dst mv s =>
+ let e1 := subst_reg_kind dst Full src Low e in
+ assertion (reg_unconstrained dst e1);
+ track_moves mv e1
+ | BShighlong src dst mv s =>
+ let e1 := subst_reg_kind dst Full src High e in
+ assertion (reg_unconstrained dst e1);
+ track_moves mv e1
+ | BSop op args res mv1 args' res' mv2 s =>
+ do e1 <- track_moves mv2 e;
+ do e2 <- transfer_use_def args res args' res' (destroyed_by_op op) e1;
+ track_moves mv1 e2
+ | BSopdead op args res mv s =>
+ assertion (reg_unconstrained res e);
+ track_moves mv e
+ | BSload chunk addr args dst mv1 args' dst' mv2 s =>
+ do e1 <- track_moves mv2 e;
+ do e2 <- transfer_use_def args dst args' dst' (destroyed_by_load chunk addr) e1;
+ track_moves mv1 e2
+ | BSload2 addr addr' args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s =>
+ do e1 <- track_moves mv3 e;
+ let e2 := remove_equation (Eq kind_second_word dst (R dst2')) e1 in
+ assertion (loc_unconstrained (R dst2') e2);
+ assertion (can_undef (destroyed_by_load Mint32 addr') e2);
+ do e3 <- add_equations args args2' e2;
+ do e4 <- track_moves mv2 e3;
+ let e5 := remove_equation (Eq kind_first_word dst (R dst1')) e4 in
+ assertion (loc_unconstrained (R dst1') e5);
+ assertion (can_undef (destroyed_by_load Mint32 addr) e5);
+ assertion (reg_unconstrained dst e5);
+ do e6 <- add_equations args args1' e5;
+ track_moves mv1 e6
+ | BSload2_1 addr args dst mv1 args' dst' mv2 s =>
+ do e1 <- track_moves mv2 e;
+ let e2 := remove_equation (Eq kind_first_word dst (R dst')) e1 in
+ assertion (reg_loc_unconstrained dst (R dst') e2);
+ assertion (can_undef (destroyed_by_load Mint32 addr) e2);
+ do e3 <- add_equations args args' e2;
+ track_moves mv1 e3
+ | BSload2_2 addr addr' args dst mv1 args' dst' mv2 s =>
+ do e1 <- track_moves mv2 e;
+ let e2 := remove_equation (Eq kind_second_word dst (R dst')) e1 in
+ assertion (reg_loc_unconstrained dst (R dst') e2);
+ assertion (can_undef (destroyed_by_load Mint32 addr') e2);
+ do e3 <- add_equations args args' e2;
+ track_moves mv1 e3
+ | BSloaddead chunk addr args dst mv s =>
+ assertion (reg_unconstrained dst e);
+ track_moves mv e
+ | BSstore chunk addr args src mv args' src' s =>
+ assertion (can_undef (destroyed_by_store chunk addr) e);
+ do e1 <- add_equations (src :: args) (src' :: args') e;
+ track_moves mv e1
+ | BSstore2 addr addr' args src mv1 args1' src1' mv2 args2' src2' s =>
+ assertion (can_undef (destroyed_by_store Mint32 addr') e);
+ do e1 <- add_equations args args2'
+ (add_equation (Eq kind_second_word src (R src2')) e);
+ do e2 <- track_moves mv2 e1;
+ assertion (can_undef (destroyed_by_store Mint32 addr) e2);
+ do e3 <- add_equations args args1'
+ (add_equation (Eq kind_first_word src (R src1')) e2);
+ track_moves mv1 e3
+ | BScall sg ros args res mv1 ros' mv2 s =>
+ let args' := loc_arguments sg in
+ let res' := map R (loc_result sg) in
+ do e1 <- track_moves mv2 e;
+ do e2 <- remove_equations_res res (sig_res sg) res' e1;
+ assertion (forallb (fun l => reg_loc_unconstrained res l e2) res');
+ assertion (no_caller_saves e2);
+ do e3 <- add_equation_ros ros ros' e2;
+ do e4 <- add_equations_args args (sig_args sg) args' e3;
+ track_moves mv1 e4
+ | BStailcall sg ros args mv1 ros' =>
+ let args' := loc_arguments sg in
+ assertion (tailcall_is_possible sg);
+ assertion (opt_typ_eq sg.(sig_res) f.(RTL.fn_sig).(sig_res));
+ assertion (ros_compatible_tailcall ros');
+ do e1 <- add_equation_ros ros ros' empty_eqs;
+ do e2 <- add_equations_args args (sig_args sg) args' e1;
+ track_moves mv1 e2
+ | BSbuiltin ef args res mv1 args' res' mv2 s =>
+ do e1 <- track_moves mv2 e;
+ let args' := map R args' in
+ let res' := map R res' in
+ do e2 <- remove_equations_res res (sig_res (ef_sig ef)) res' e1;
+ assertion (reg_unconstrained res e2);
+ assertion (forallb (fun l => loc_unconstrained l e2) res');
+ assertion (can_undef (destroyed_by_builtin ef) e2);
+ do e3 <- add_equations_args args (sig_args (ef_sig ef)) args' e2;
+ track_moves mv1 e3
+ | BSannot txt typ args res mv1 args' s =>
+ do e1 <- add_equations_args args (annot_args_typ typ) args' e;
+ track_moves mv1 e1
+ | BScond cond args mv args' s1 s2 =>
+ assertion (can_undef (destroyed_by_cond cond) e);
+ do e1 <- add_equations args args' e;
+ track_moves mv e1
+ | BSjumptable arg mv arg' tbl =>
+ assertion (can_undef destroyed_by_jumptable e);
+ track_moves mv (add_equation (Eq Full arg (R arg')) e)
+ | BSreturn None mv =>
+ track_moves mv empty_eqs
+ | BSreturn (Some arg) mv =>
+ let arg' := map R (loc_result (RTL.fn_sig f)) in
+ do e1 <- add_equations_res arg (sig_res (RTL.fn_sig f)) arg' empty_eqs;
+ track_moves mv e1
+ end.
+
+(** The main transfer function for the dataflow analysis. Like [transfer_aux],
+ it infers the equations that must hold "before" as a function of the
+ equations that must hold "after". It also handles error propagation
+ and reporting. *)
+
+Definition transfer (f: RTL.function) (shapes: PTree.t block_shape)
+ (pc: node) (after: res eqs) : res eqs :=
+ match after with
+ | Error _ => after
+ | OK e =>
+ match shapes!pc with
+ | None => Error(MSG "At PC " :: POS pc :: MSG ": unmatched block" :: nil)
+ | Some shape =>
+ match transfer_aux f shape e with
+ | None => Error(MSG "At PC " :: POS pc :: MSG ": invalid register allocation" :: nil)
+ | Some e' => OK e'
+ end
+ end
+ end.
+
+(** The semilattice for dataflow analysis. Operates on analysis results
+ of type [res eqs], that is, either a set of equations or an error
+ message. Errors correspond to [Top]. Sets of equations are ordered
+ by inclusion. *)
+
+Module LEq <: SEMILATTICE.
+
+ Definition t := res eqs.
+
+ Definition eq (x y: t) :=
+ match x, y with
+ | OK a, OK b => EqSet.Equal a b
+ | Error _, Error _ => True
+ | _, _ => False
+ end.
+
+ Lemma eq_refl: forall x, eq x x.
+ Proof.
+ intros; destruct x; simpl; auto. red; tauto.
+ Qed.
+
+ Lemma eq_sym: forall x y, eq x y -> eq y x.
+ Proof.
+ unfold eq; intros; destruct x; destruct y; auto.
+ red in H; red; intros. rewrite H; tauto.
+ Qed.
+
+ Lemma eq_trans: forall x y z, eq x y -> eq y z -> eq x z.
+ Proof.
+ unfold eq; intros. destruct x; destruct y; try contradiction; destruct z; auto.
+ red in H; red in H0; red; intros. rewrite H. auto.
+ Qed.
+
+ Definition beq (x y: t) :=
+ match x, y with
+ | OK a, OK b => EqSet.equal a b
+ | Error _, Error _ => true
+ | _, _ => false
+ end.
+
+ Lemma beq_correct: forall x y, beq x y = true -> eq x y.
+ Proof.
+ unfold beq, eq; intros. destruct x; destruct y.
+ apply EqSet.equal_2. auto.
+ discriminate.
+ discriminate.
+ auto.
+ Qed.
+
+ Definition ge (x y: t) :=
+ match x, y with
+ | OK a, OK b => EqSet.Subset b a
+ | Error _, _ => True
+ | _, Error _ => False
+ end.
+
+ Lemma ge_refl: forall x y, eq x y -> ge x y.
+ Proof.
+ unfold eq, ge, EqSet.Equal, EqSet.Subset; intros.
+ destruct x; destruct y; auto. intros; rewrite H; auto.
+ Qed.
+ Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z.
+ Proof.
+ unfold ge, EqSet.Subset; intros.
+ destruct x; auto; destruct y; try contradiction.
+ destruct z; eauto.
+ Qed.
+
+ Definition bot: t := OK empty_eqs.
+
+ Lemma ge_bot: forall x, ge x bot.
+ Proof.
+ unfold ge, bot, EqSet.Subset; simpl; intros.
+ destruct x; auto. intros. elim (EqSet.empty_1 H).
+ Qed.
+
+ Program Definition lub (x y: t) : t :=
+ match x, y return _ with
+ | OK a, OK b =>
+ OK (mkeqs (EqSet.union (eqs1 a) (eqs1 b))
+ (EqSet2.union (eqs2 a) (eqs2 b)) _)
+ | OK _, Error _ => y
+ | Error _, _ => x
+ end.
+ Next Obligation.
+ split; intros.
+ apply EqSet2.union_1 in H. destruct H; rewrite eqs_same in H.
+ apply EqSet.union_2; auto. apply EqSet.union_3; auto.
+ apply EqSet.union_1 in H. destruct H; rewrite <- eqs_same in H.
+ apply EqSet2.union_2; auto. apply EqSet2.union_3; auto.
+ Qed.
+
+ Lemma ge_lub_left: forall x y, ge (lub x y) x.
+ Proof.
+ unfold lub, ge, EqSet.Subset; intros.
+ destruct x; destruct y; auto.
+ intros; apply EqSet.union_2; auto.
+ Qed.
+
+ Lemma ge_lub_right: forall x y, ge (lub x y) y.
+ Proof.
+ unfold lub, ge, EqSet.Subset; intros.
+ destruct x; destruct y; auto.
+ intros; apply EqSet.union_3; auto.
+ Qed.
+
+End LEq.
+
+(** The backward dataflow solver is an instantiation of Kildall's algorithm. *)
+
+Module DS := Backward_Dataflow_Solver(LEq)(NodeSetBackward).
+
+(** The control-flow graph that the solver operates on is the CFG of
+ block shapes built by the structural check phase. Here is its notion
+ of successors. *)
+
+Definition successors_block_shape (bsh: block_shape) : list node :=
+ match bsh with
+ | BSnop mv s => s :: nil
+ | BSmove src dst mv s => s :: nil
+ | BSmakelong src1 src2 dst mv s => s :: nil
+ | BSlowlong src dst mv s => s :: nil
+ | BShighlong src dst mv s => s :: nil
+ | BSop op args res mv1 args' res' mv2 s => s :: nil
+ | BSopdead op args res mv s => s :: nil
+ | BSload chunk addr args dst mv1 args' dst' mv2 s => s :: nil
+ | BSload2 addr addr' args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s => s :: nil
+ | BSload2_1 addr args dst mv1 args' dst' mv2 s => s :: nil
+ | BSload2_2 addr addr' args dst mv1 args' dst' mv2 s => s :: nil
+ | BSloaddead chunk addr args dst mv s => s :: nil
+ | BSstore chunk addr args src mv1 args' src' s => s :: nil
+ | BSstore2 addr addr' args src mv1 args1' src1' mv2 args2' src2' s => s :: nil
+ | BScall sg ros args res mv1 ros' mv2 s => s :: nil
+ | BStailcall sg ros args mv1 ros' => nil
+ | BSbuiltin ef args res mv1 args' res' mv2 s => s :: nil
+ | BSannot txt typ args res mv1 args' s => s :: nil
+ | BScond cond args mv args' s1 s2 => s1 :: s2 :: nil
+ | BSjumptable arg mv arg' tbl => tbl
+ | BSreturn optarg mv => nil
+ end.
+
+Definition analyze (f: RTL.function) (bsh: PTree.t block_shape) :=
+ DS.fixpoint (PTree.map1 successors_block_shape bsh) (transfer f bsh) nil.
+
+(** * Validating and translating functions and programs *)
+
+(** Checking equations at function entry point. The RTL function receives
+ its arguments in the list [rparams] of pseudoregisters. The LTL function
+ receives them in the list [lparams] of locations dictated by the
+ calling conventions, with arguments of type [Tlong] being split in
+ two 32-bit halves. We check that the equations [e] that must hold
+ at the beginning of the functions are compatible with these calling
+ conventions, in the sense that all equations involving a pseudoreg
+ [r] from [rparams] is of the form [r = l [Full]] or [r = l [Low]]
+ or [r = l [High]], where [l] is the corresponding element of [lparams].
+
+ Note that [e] can contain additional equations [r' = l [kind]]
+ involving pseudoregs [r'] not in [rparams]: these equations are
+ automatically satisfied since the initial value of [r'] is [Vundef]. *)
+
+Function compat_entry (rparams: list reg) (tys: list typ) (lparams: list loc) (e: eqs)
+ {struct rparams} : bool :=
+ match rparams, tys, lparams with
+ | nil, nil, nil => true
+ | r1 :: rl, Tlong :: tyl, l1 :: l2 :: ll =>
+ compat_left2 r1 l1 l2 e && compat_entry rl tyl ll e
+ | r1 :: rl, (Tint|Tfloat) :: tyl, l1 :: ll =>
+ compat_left r1 l1 e && compat_entry rl tyl ll e
+ | _, _, _ => false
+ end.
+
+(** Checking the satisfiability of equations inferred at function entry
+ point. We also check that the RTL and LTL functions agree in signature
+ and stack size. *)
+
+Definition check_entrypoints_aux (rtl: RTL.function) (ltl: LTL.function) (e1: eqs) : option unit :=
+ do mv <- pair_entrypoints rtl ltl;
+ do e2 <- track_moves mv e1;
+ assertion (compat_entry (RTL.fn_params rtl)
+ (sig_args (RTL.fn_sig rtl))
+ (loc_parameters (RTL.fn_sig rtl)) e2);
+ assertion (can_undef destroyed_at_function_entry e2);
+ assertion (zeq (RTL.fn_stacksize rtl) (LTL.fn_stacksize ltl));
+ assertion (eq_signature (RTL.fn_sig rtl) (LTL.fn_sig ltl));
+ Some tt.
+
+Local Close Scope option_monad_scope.
+Local Open Scope error_monad_scope.
+
+Definition check_entrypoints (rtl: RTL.function) (ltl: LTL.function)
+ (bsh: PTree.t block_shape)
+ (a: PMap.t LEq.t): res unit :=
+ do e1 <- transfer rtl bsh (RTL.fn_entrypoint rtl) a!!(RTL.fn_entrypoint rtl);
+ match check_entrypoints_aux rtl ltl e1 with
+ | None => Error (msg "invalid register allocation at entry point")
+ | Some _ => OK tt
+ end.
+
+(** Putting it all together, this is the validation function for
+ a source RTL function and an LTL function generated by the external
+ register allocator. *)
+
+Definition check_function (rtl: RTL.function) (ltl: LTL.function) : res unit :=
+ let bsh := pair_codes rtl ltl in
+ match analyze rtl bsh with
+ | None => Error (msg "allocation analysis diverges")
+ | Some a => check_entrypoints rtl ltl bsh a
+ end.
-(** The translation of a function performs liveness analysis,
- construction and coloring of the inference graph, and per-instruction
- transformation as described above. *)
+(** [regalloc] is the external register allocator. It is written in OCaml
+ in file [backend/Regalloc.ml]. *)
-Definition live0 (f: RTL.function) (live: PMap.t Regset.t) :=
- transfer f f.(RTL.fn_entrypoint) live!!(f.(RTL.fn_entrypoint)).
+Parameter regalloc: RTL.function -> res LTL.function.
-Open Scope string_scope.
+(** Register allocation followed by validation. *)
Definition transf_function (f: RTL.function) : res LTL.function :=
match type_function f with
- | Error msg => Error msg
- | OK env =>
- match analyze f with
- | None => Error (msg "Liveness analysis failure")
- | Some live =>
- match regalloc f live (live0 f live) env with
- | None => Error (msg "Incorrect graph coloring")
- | Some assign => OK (transf_fun f live assign)
- end
- end
+ | Error m => Error m
+ | OK tyenv =>
+ match regalloc f with
+ | Error m => Error m
+ | OK tf => do x <- check_function f tf; OK tf
+ end
end.
Definition transf_fundef (fd: RTL.fundef) : res LTL.fundef :=
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index 8dfa2d4..76e9744 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -10,403 +10,1335 @@
(* *)
(* *********************************************************************)
-(** Correctness proof for the [Allocation] pass (translation from
+(** Correctness proof for the [Allocation] pass (validated translation from
RTL to LTL). *)
Require Import FSets.
Require Import Coqlib.
+Require Import Ordered.
Require Import Errors.
Require Import Maps.
+Require Import Lattice.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Memory.
Require Import Events.
-Require Import Smallstep.
Require Import Globalenvs.
+Require Import Smallstep.
Require Import Op.
Require Import Registers.
Require Import RTL.
Require Import RTLtyping.
-Require Import Liveness.
+Require Import Kildall.
Require Import Locations.
-Require Import LTL.
Require Import Conventions.
-Require Import Coloring.
-Require Import Coloringproof.
+Require Import LTL.
Require Import Allocation.
-(** * Properties of allocated locations *)
+(** * Soundness of structural checks *)
-(** We list here various properties of the locations [alloc r],
- where [r] is an RTL pseudo-register and [alloc] is the register
- assignment returned by [regalloc]. *)
+Definition expand_move (sd: loc * loc) : instruction :=
+ match sd with
+ | (R src, R dst) => Lop Omove (src::nil) dst
+ | (S sl ofs ty, R dst) => Lgetstack sl ofs ty dst
+ | (R src, S sl ofs ty) => Lsetstack src sl ofs ty
+ | (S _ _ _, S _ _ _) => Lreturn (**r should never happen *)
+ end.
-Section REGALLOC_PROPERTIES.
+Definition expand_moves (mv: moves) (k: bblock) : bblock :=
+ List.map expand_move mv ++ k.
-Variable f: RTL.function.
-Variable env: regenv.
-Variable live: PMap.t Regset.t.
-Variable alloc: reg -> loc.
-Hypothesis ALLOC: regalloc f live (live0 f live) env = Some alloc.
+Definition wf_move (sd: loc * loc) : Prop :=
+ match sd with
+ | (S _ _ _, S _ _ _) => False
+ | _ => True
+ end.
-Lemma regalloc_noteq_diff:
- forall r1 l2,
- alloc r1 <> l2 -> Loc.diff (alloc r1) l2.
+Definition wf_moves (mv: moves) : Prop :=
+ forall sd, In sd mv -> wf_move sd.
+
+Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Prop :=
+ | ebs_nop: forall mv s k,
+ wf_moves mv ->
+ expand_block_shape (BSnop mv s)
+ (Inop s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_move: forall src dst mv s k,
+ wf_moves mv ->
+ expand_block_shape (BSmove src dst mv s)
+ (Iop Omove (src :: nil) dst s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_makelong: forall src1 src2 dst mv s k,
+ wf_moves mv ->
+ expand_block_shape (BSmakelong src1 src2 dst mv s)
+ (Iop Omakelong (src1 :: src2 :: nil) dst s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_lowlong: forall src dst mv s k,
+ wf_moves mv ->
+ expand_block_shape (BSlowlong src dst mv s)
+ (Iop Olowlong (src :: nil) dst s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_highlong: forall src dst mv s k,
+ wf_moves mv ->
+ expand_block_shape (BShighlong src dst mv s)
+ (Iop Ohighlong (src :: nil) dst s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_op: forall op args res mv1 args' res' mv2 s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ expand_block_shape (BSop op args res mv1 args' res' mv2 s)
+ (Iop op args res s)
+ (expand_moves mv1
+ (Lop op args' res' :: expand_moves mv2 (Lbranch s :: k)))
+ | ebs_op_dead: forall op args res mv s k,
+ wf_moves mv ->
+ expand_block_shape (BSopdead op args res mv s)
+ (Iop op args res s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_load: forall chunk addr args dst mv1 args' dst' mv2 s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ expand_block_shape (BSload chunk addr args dst mv1 args' dst' mv2 s)
+ (Iload chunk addr args dst s)
+ (expand_moves mv1
+ (Lload chunk addr args' dst' :: expand_moves mv2 (Lbranch s :: k)))
+ | ebs_load2: forall addr addr2 args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s k,
+ wf_moves mv1 -> wf_moves mv2 -> wf_moves mv3 ->
+ offset_addressing addr (Int.repr 4) = Some addr2 ->
+ expand_block_shape (BSload2 addr addr2 args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s)
+ (Iload Mint64 addr args dst s)
+ (expand_moves mv1
+ (Lload Mint32 addr args1' dst1' ::
+ expand_moves mv2
+ (Lload Mint32 addr2 args2' dst2' ::
+ expand_moves mv3 (Lbranch s :: k))))
+ | ebs_load2_1: forall addr args dst mv1 args' dst' mv2 s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ expand_block_shape (BSload2_1 addr args dst mv1 args' dst' mv2 s)
+ (Iload Mint64 addr args dst s)
+ (expand_moves mv1
+ (Lload Mint32 addr args' dst' ::
+ expand_moves mv2 (Lbranch s :: k)))
+ | ebs_load2_2: forall addr addr2 args dst mv1 args' dst' mv2 s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ offset_addressing addr (Int.repr 4) = Some addr2 ->
+ expand_block_shape (BSload2_2 addr addr2 args dst mv1 args' dst' mv2 s)
+ (Iload Mint64 addr args dst s)
+ (expand_moves mv1
+ (Lload Mint32 addr2 args' dst' ::
+ expand_moves mv2 (Lbranch s :: k)))
+ | ebs_load_dead: forall chunk addr args dst mv s k,
+ wf_moves mv ->
+ expand_block_shape (BSloaddead chunk addr args dst mv s)
+ (Iload chunk addr args dst s)
+ (expand_moves mv (Lbranch s :: k))
+ | ebs_store: forall chunk addr args src mv1 args' src' s k,
+ wf_moves mv1 ->
+ expand_block_shape (BSstore chunk addr args src mv1 args' src' s)
+ (Istore chunk addr args src s)
+ (expand_moves mv1
+ (Lstore chunk addr args' src' :: Lbranch s :: k))
+ | ebs_store2: forall addr addr2 args src mv1 args1' src1' mv2 args2' src2' s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ offset_addressing addr (Int.repr 4) = Some addr2 ->
+ expand_block_shape (BSstore2 addr addr2 args src mv1 args1' src1' mv2 args2' src2' s)
+ (Istore Mint64 addr args src s)
+ (expand_moves mv1
+ (Lstore Mint32 addr args1' src1' ::
+ expand_moves mv2
+ (Lstore Mint32 addr2 args2' src2' ::
+ Lbranch s :: k)))
+ | ebs_call: forall sg ros args res mv1 ros' mv2 s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ expand_block_shape (BScall sg ros args res mv1 ros' mv2 s)
+ (Icall sg ros args res s)
+ (expand_moves mv1
+ (Lcall sg ros' :: expand_moves mv2 (Lbranch s :: k)))
+ | ebs_tailcall: forall sg ros args mv ros' k,
+ wf_moves mv ->
+ expand_block_shape (BStailcall sg ros args mv ros')
+ (Itailcall sg ros args)
+ (expand_moves mv (Ltailcall sg ros' :: k))
+ | ebs_builtin: forall ef args res mv1 args' res' mv2 s k,
+ wf_moves mv1 -> wf_moves mv2 ->
+ expand_block_shape (BSbuiltin ef args res mv1 args' res' mv2 s)
+ (Ibuiltin ef args res s)
+ (expand_moves mv1
+ (Lbuiltin ef args' res' :: expand_moves mv2 (Lbranch s :: k)))
+ | ebs_annot: forall txt typ args res mv args' s k,
+ wf_moves mv ->
+ expand_block_shape (BSannot txt typ args res mv args' s)
+ (Ibuiltin (EF_annot txt typ) args res s)
+ (expand_moves mv
+ (Lannot (EF_annot txt typ) args' :: Lbranch s :: k))
+ | ebs_cond: forall cond args mv args' s1 s2 k,
+ wf_moves mv ->
+ expand_block_shape (BScond cond args mv args' s1 s2)
+ (Icond cond args s1 s2)
+ (expand_moves mv (Lcond cond args' s1 s2 :: k))
+ | ebs_jumptable: forall arg mv arg' tbl k,
+ wf_moves mv ->
+ expand_block_shape (BSjumptable arg mv arg' tbl)
+ (Ijumptable arg tbl)
+ (expand_moves mv (Ljumptable arg' tbl :: k))
+ | ebs_return: forall optarg mv k,
+ wf_moves mv ->
+ expand_block_shape (BSreturn optarg mv)
+ (Ireturn optarg)
+ (expand_moves mv (Lreturn :: k)).
+
+Ltac MonadInv :=
+ match goal with
+ | [ H: match ?x with Some _ => _ | None => None end = Some _ |- _ ] =>
+ destruct x as [] eqn:? ; [MonadInv|discriminate]
+ | [ H: match ?x with left _ => _ | right _ => None end = Some _ |- _ ] =>
+ destruct x; [MonadInv|discriminate]
+ | [ H: match negb (proj_sumbool ?x) with true => _ | false => None end = Some _ |- _ ] =>
+ destruct x; [discriminate|simpl in H; MonadInv]
+ | [ H: match negb ?x with true => _ | false => None end = Some _ |- _ ] =>
+ destruct x; [discriminate|simpl in H; MonadInv]
+ | [ H: match ?x with true => _ | false => None end = Some _ |- _ ] =>
+ destruct x as [] eqn:? ; [MonadInv|discriminate]
+ | [ H: match ?x with (_, _) => _ end = Some _ |- _ ] =>
+ destruct x as [] eqn:? ; MonadInv
+ | [ H: Some _ = Some _ |- _ ] =>
+ inv H; MonadInv
+ | [ H: None = Some _ |- _ ] =>
+ discriminate
+ | _ =>
+ idtac
+ end.
+
+Lemma extract_moves_sound:
+ forall b mv b',
+ extract_moves nil b = (mv, b') ->
+ wf_moves mv /\ b = expand_moves mv b'.
Proof.
- intros. apply loc_acceptable_noteq_diff.
- eapply regalloc_acceptable; eauto.
- auto.
-Qed.
+ assert (BASE:
+ forall accu b,
+ wf_moves accu ->
+ wf_moves (List.rev accu) /\ expand_moves (List.rev accu) b = expand_moves (List.rev accu) b).
+ intros; split; auto.
+ red; intros. apply H. rewrite <- in_rev in H0; auto.
+
+ assert (IND: forall b accu mv b',
+ extract_moves accu b = (mv, b') ->
+ wf_moves accu ->
+ wf_moves mv /\ expand_moves (List.rev accu) b = expand_moves mv b').
+ induction b; simpl; intros.
+ inv H. auto.
+ destruct a; try (inv H; apply BASE; auto; fail).
+ destruct (is_move_operation op args) as [arg|] eqn:E.
+ exploit is_move_operation_correct; eauto. intros [A B]; subst.
+ (* reg-reg move *)
+ exploit IHb; eauto.
+ red; intros. destruct H1; auto. subst sd; exact I.
+ intros [P Q].
+ split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
+ rewrite app_ass. simpl. auto.
+ inv H; apply BASE; auto.
+ (* stack-reg move *)
+ exploit IHb; eauto.
+ red; intros. destruct H1; auto. subst sd; exact I.
+ intros [P Q].
+ split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
+ rewrite app_ass. simpl. auto.
+ (* reg-stack move *)
+ exploit IHb; eauto.
+ red; intros. destruct H1; auto. subst sd; exact I.
+ intros [P Q].
+ split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
+ rewrite app_ass. simpl. auto.
+
+ intros. exploit IND; eauto. red; intros. elim H0.
+Qed.
+
+Lemma check_succ_sound:
+ forall s b, check_succ s b = true -> exists k, b = Lbranch s :: k.
+Proof.
+ intros. destruct b; simpl in H; try discriminate.
+ destruct i; try discriminate.
+ destruct (peq s s0); simpl in H; inv H. exists b; auto.
+Qed.
+
+Ltac UseParsingLemmas :=
+ match goal with
+ | [ H: extract_moves nil _ = (_, _) |- _ ] =>
+ destruct (extract_moves_sound _ _ _ H); clear H; subst; UseParsingLemmas
+ | [ H: check_succ _ _ = true |- _ ] =>
+ try discriminate;
+ destruct (check_succ_sound _ _ H); clear H; subst; UseParsingLemmas
+ | _ => idtac
+ end.
-Lemma regalloc_notin_notin:
- forall r ll,
- ~(In (alloc r) ll) -> Loc.notin (alloc r) ll.
+Lemma pair_instr_block_sound:
+ forall i b bsh,
+ pair_instr_block i b = Some bsh -> expand_block_shape bsh i b.
Proof.
- intros. apply loc_acceptable_notin_notin.
- eapply regalloc_acceptable; eauto. auto.
+ intros; destruct i; simpl in H; MonadInv; UseParsingLemmas.
+(* nop *)
+ econstructor; eauto.
+(* op *)
+ destruct (classify_operation o l).
+ (* move *)
+ MonadInv; UseParsingLemmas. econstructor; eauto.
+ (* makelong *)
+ MonadInv; UseParsingLemmas. econstructor; eauto.
+ (* lowlong *)
+ MonadInv; UseParsingLemmas. econstructor; eauto.
+ (* highlong *)
+ MonadInv; UseParsingLemmas. econstructor; eauto.
+ (* other ops *)
+ MonadInv. destruct b0.
+ MonadInv; UseParsingLemmas.
+ destruct i; MonadInv; UseParsingLemmas.
+ eapply ebs_op; eauto.
+ inv H0. eapply ebs_op_dead; eauto.
+(* load *)
+ destruct b0.
+ MonadInv; UseParsingLemmas.
+ destruct i; MonadInv; UseParsingLemmas.
+ destruct (eq_chunk m Mint64).
+ MonadInv; UseParsingLemmas.
+ destruct b; MonadInv; UseParsingLemmas. destruct i; MonadInv; UseParsingLemmas.
+ eapply ebs_load2; eauto.
+ destruct (eq_addressing a addr).
+ MonadInv. inv H2. eapply ebs_load2_1; eauto.
+ MonadInv. inv H2. eapply ebs_load2_2; eauto.
+ MonadInv; UseParsingLemmas. eapply ebs_load; eauto.
+ inv H. eapply ebs_load_dead; eauto.
+(* store *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas.
+ destruct (eq_chunk m Mint64).
+ MonadInv; UseParsingLemmas.
+ destruct b; MonadInv. destruct i; MonadInv; UseParsingLemmas.
+ eapply ebs_store2; eauto.
+ MonadInv; UseParsingLemmas.
+ eapply ebs_store; eauto.
+(* call *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
+(* tailcall *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
+(* builtin *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas.
+ econstructor; eauto.
+ destruct ef; inv H. econstructor; eauto.
+(* cond *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
+(* jumptable *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
+(* return *)
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
Qed.
-Lemma regalloc_notin_notin_2:
- forall l rl,
- ~(In l (map alloc rl)) -> Loc.notin l (map alloc rl).
+Lemma matching_instr_block:
+ forall f1 f2 pc bsh i,
+ (pair_codes f1 f2)!pc = Some bsh ->
+ (RTL.fn_code f1)!pc = Some i ->
+ exists b, (LTL.fn_code f2)!pc = Some b /\ expand_block_shape bsh i b.
Proof.
- induction rl; simpl; intros. auto.
- split. apply Loc.diff_sym. apply regalloc_noteq_diff. tauto.
- apply IHrl. tauto.
+ intros. unfold pair_codes in H. rewrite PTree.gcombine in H; auto. rewrite H0 in H.
+ destruct (LTL.fn_code f2)!pc as [b|].
+ exists b; split; auto. apply pair_instr_block_sound; auto.
+ discriminate.
Qed.
-
-Lemma regalloc_norepet_norepet:
- forall rl,
- list_norepet (List.map alloc rl) ->
- Loc.norepet (List.map alloc rl).
+
+(** * Properties of equations *)
+
+Module ESF := FSetFacts.Facts(EqSet).
+Module ESP := FSetProperties.Properties(EqSet).
+Module ESD := FSetDecide.Decide(EqSet).
+
+Definition sel_val (k: equation_kind) (v: val) : val :=
+ match k with
+ | Full => v
+ | Low => Val.loword v
+ | High => Val.hiword v
+ end.
+
+(** A set of equations [e] is satisfied in a RTL pseudoreg state [rs]
+ and an LTL location state [ls] if, for every equation [r = l [k]] in [e],
+ [sel_val k (rs#r)] (the [k] fragment of [r]'s value in the RTL code)
+ is less defined than [ls l] (the value of [l] in the LTL code). *)
+
+Definition satisf (rs: regset) (ls: locset) (e: eqs) : Prop :=
+ forall q, EqSet.In q e -> Val.lessdef (sel_val (ekind q) rs#(ereg q)) (ls (eloc q)).
+
+Lemma empty_eqs_satisf:
+ forall rs ls, satisf rs ls empty_eqs.
+Proof.
+ unfold empty_eqs; intros; red; intros. ESD.fsetdec.
+Qed.
+
+Lemma satisf_incr:
+ forall rs ls (e1 e2: eqs),
+ satisf rs ls e2 -> EqSet.Subset e1 e2 -> satisf rs ls e1.
Proof.
- induction rl; simpl; intros.
- apply Loc.norepet_nil.
- inversion H.
- apply Loc.norepet_cons.
- eapply regalloc_notin_notin; eauto.
+ unfold satisf; intros. apply H. ESD.fsetdec.
+Qed.
+
+Lemma satisf_undef_reg:
+ forall rs ls e r,
+ satisf rs ls e ->
+ satisf (rs#r <- Vundef) ls e.
+Proof.
+ intros; red; intros. rewrite Regmap.gsspec. destruct (peq (ereg q) r); auto.
+ destruct (ekind q); simpl; auto.
+Qed.
+
+Lemma add_equation_lessdef:
+ forall rs ls q e,
+ satisf rs ls (add_equation q e) -> Val.lessdef (sel_val (ekind q) rs#(ereg q)) (ls (eloc q)).
+Proof.
+ intros. apply H. unfold add_equation. simpl. apply EqSet.add_1. auto.
+Qed.
+
+Lemma add_equation_satisf:
+ forall rs ls q e,
+ satisf rs ls (add_equation q e) -> satisf rs ls e.
+Proof.
+ intros. eapply satisf_incr; eauto. unfold add_equation. simpl. ESD.fsetdec.
+Qed.
+
+Lemma add_equations_satisf:
+ forall rs ls rl ml e e',
+ add_equations rl ml e = Some e' ->
+ satisf rs ls e' -> satisf rs ls e.
+Proof.
+ induction rl; destruct ml; simpl; intros; MonadInv.
auto.
+ eapply add_equation_satisf; eauto.
Qed.
-Lemma regalloc_not_temporary:
- forall (r: reg),
- Loc.notin (alloc r) temporaries.
+Lemma add_equations_lessdef:
+ forall rs ls rl ml e e',
+ add_equations rl ml e = Some e' ->
+ satisf rs ls e' ->
+ Val.lessdef_list (rs##rl) (reglist ls ml).
Proof.
- intros. apply temporaries_not_acceptable.
- eapply regalloc_acceptable; eauto.
+ induction rl; destruct ml; simpl; intros; MonadInv.
+ constructor.
+ constructor; eauto.
+ apply add_equation_lessdef with (e := e) (q := Eq Full a (R m)).
+ eapply add_equations_satisf; eauto.
Qed.
-Lemma regalloc_disj_temporaries:
- forall (rl: list reg),
- Loc.disjoint (List.map alloc rl) temporaries.
+Lemma add_equations_args_satisf:
+ forall rs ls rl tyl ll e e',
+ add_equations_args rl tyl ll e = Some e' ->
+ satisf rs ls e' -> satisf rs ls e.
Proof.
- intros.
- apply Loc.notin_disjoint. intros.
- generalize (list_in_map_inv _ _ _ H). intros [r [EQ IN]].
- subst x. apply regalloc_not_temporary; auto.
+ intros until e'. functional induction (add_equations_args rl tyl ll e); intros.
+ inv H; auto.
+ eapply add_equation_satisf. eapply add_equation_satisf. eauto.
+ eapply add_equation_satisf. eauto.
+ eapply add_equation_satisf. eauto.
+ congruence.
Qed.
-End REGALLOC_PROPERTIES.
+Lemma val_longofwords_eq:
+ forall v,
+ Val.has_type v Tlong ->
+ Val.longofwords (Val.hiword v) (Val.loword v) = v.
+Proof.
+ intros. red in H. destruct v; try contradiction.
+ reflexivity.
+ simpl. rewrite Int64.ofwords_recompose. auto.
+Qed.
-(** * Semantic agreement between RTL registers and LTL locations *)
+Lemma add_equations_args_lessdef:
+ forall rs ls rl tyl ll e e',
+ add_equations_args rl tyl ll e = Some e' ->
+ satisf rs ls e' ->
+ Val.has_type_list (rs##rl) tyl ->
+ Val.lessdef_list (rs##rl) (decode_longs tyl (map ls ll)).
+Proof.
+ intros until e'. functional induction (add_equations_args rl tyl ll e); simpl; intros.
+- inv H; auto.
+- destruct H1. constructor; auto.
+ rewrite <- (val_longofwords_eq (rs#r1)); auto. apply Val.longofwords_lessdef.
+ eapply add_equation_lessdef with (q := Eq High r1 l1).
+ eapply add_equation_satisf. eapply add_equations_args_satisf; eauto.
+ eapply add_equation_lessdef with (q := Eq Low r1 l2).
+ eapply add_equations_args_satisf; eauto.
+- destruct H1. constructor; auto.
+ eapply add_equation_lessdef with (q := Eq Full r1 l1). eapply add_equations_args_satisf; eauto.
+- destruct H1. constructor; auto.
+ eapply add_equation_lessdef with (q := Eq Full r1 l1). eapply add_equations_args_satisf; eauto.
+- discriminate.
+Qed.
-Require Import LTL.
-Module RegsetP := Properties(Regset).
-
-Section AGREE.
-
-Variable f: RTL.function.
-Variable env: regenv.
-Variable flive: PMap.t Regset.t.
-Variable assign: reg -> loc.
-Hypothesis REGALLOC: regalloc f flive (live0 f flive) env = Some assign.
-
-(** Remember the core of the code transformation performed in module
- [Allocation]: every reference to register [r] is replaced by
- a reference to location [assign r]. We will shortly prove
- the semantic equivalence between the original code and the transformed code.
- The key tool to do this is the following relation between
- a register set [rs] in the original RTL program and a location set
- [ls] in the transformed LTL program. The two sets agree if
- they assign identical values to matching registers and locations,
- that is, the value of register [r] in [rs] is the same as
- the value of location [assign r] in [ls]. However, this equality
- needs to hold only for live registers [r]. If [r] is dead at
- the current point, its value is never used later, hence the value
- of [assign r] can be arbitrary. *)
-
-Definition agree (live: Regset.t) (rs: regset) (ls: locset) : Prop :=
- forall (r: reg), Regset.In r live -> rs#r = ls (assign r).
-
-(** What follows is a long list of lemmas expressing properties
- of the [agree_live_regs] predicate that are useful for the
- semantic equivalence proof. First: two register sets that agree
- on a given set of live registers also agree on a subset of
- those live registers. *)
-
-Lemma agree_increasing:
- forall live1 live2 rs ls,
- Regset.Subset live2 live1 -> agree live1 rs ls ->
- agree live2 rs ls.
-Proof.
- unfold agree; intros.
- apply H0. apply H. auto.
-Qed.
-
-Lemma agree_succ:
- forall n s rs ls live i,
- analyze f = Some live ->
- f.(RTL.fn_code)!n = Some i ->
- In s (RTL.successors_instr i) ->
- agree live!!n rs ls ->
- agree (transfer f s live!!s) rs ls.
-Proof.
- intros.
- apply agree_increasing with (live!!n).
- eapply Liveness.analyze_solution; eauto.
+Lemma add_equation_ros_satisf:
+ forall rs ls ros mos e e',
+ add_equation_ros ros mos e = Some e' ->
+ satisf rs ls e' -> satisf rs ls e.
+Proof.
+ unfold add_equation_ros; intros. destruct ros; destruct mos; MonadInv.
+ eapply add_equation_satisf; eauto.
auto.
Qed.
-(** Some useful special cases of [agree_increasing]. *)
+Lemma remove_equation_satisf:
+ forall rs ls q e,
+ satisf rs ls e -> satisf rs ls (remove_equation q e).
+Proof.
+ intros. eapply satisf_incr; eauto. unfold remove_equation; simpl. ESD.fsetdec.
+Qed.
-Lemma agree_reg_live:
- forall r live rs ls,
- agree (reg_live r live) rs ls -> agree live rs ls.
+Lemma remove_equation_res_satisf:
+ forall rs ls r oty ll e e',
+ remove_equations_res r oty ll e = Some e' ->
+ satisf rs ls e -> satisf rs ls e'.
Proof.
- intros. apply agree_increasing with (reg_live r live); auto.
- red. apply RegsetP.subset_add_2. apply RegsetP.subset_refl.
+ intros. functional inversion H.
+ apply remove_equation_satisf. apply remove_equation_satisf; auto.
+ apply remove_equation_satisf; auto.
Qed.
-Lemma agree_reg_list_live:
- forall rl live rs ls,
- agree (reg_list_live rl live) rs ls -> agree live rs ls.
+Remark select_reg_l_monotone:
+ forall r q1 q2,
+ OrderedEquation.eq q1 q2 \/ OrderedEquation.lt q1 q2 ->
+ select_reg_l r q1 = true -> select_reg_l r q2 = true.
Proof.
- induction rl; simpl; intros.
- assumption.
- apply agree_reg_live with a. apply IHrl. assumption.
+ unfold select_reg_l; intros. destruct H.
+ red in H. congruence.
+ rewrite Pos.leb_le in *. red in H. destruct H as [A | [A B]].
+ red in A. zify; omega.
+ rewrite <- A; auto.
Qed.
-Lemma agree_reg_sum_live:
- forall ros live rs ls,
- agree (reg_sum_live ros live) rs ls -> agree live rs ls.
+Remark select_reg_h_monotone:
+ forall r q1 q2,
+ OrderedEquation.eq q1 q2 \/ OrderedEquation.lt q2 q1 ->
+ select_reg_h r q1 = true -> select_reg_h r q2 = true.
Proof.
- intros. destruct ros; simpl in H.
- apply agree_reg_live with r; auto.
- auto.
+ unfold select_reg_h; intros. destruct H.
+ red in H. congruence.
+ rewrite Pos.leb_le in *. red in H. destruct H as [A | [A B]].
+ red in A. zify; omega.
+ rewrite A; auto.
Qed.
-(** Agreement over a set of live registers just extended with [r]
- implies equality of the values of [r] and [assign r]. *)
+Remark select_reg_charact:
+ forall r q, select_reg_l r q = true /\ select_reg_h r q = true <-> ereg q = r.
+Proof.
+ unfold select_reg_l, select_reg_h; intros; split.
+ rewrite ! Pos.leb_le. unfold reg; zify; omega.
+ intros. rewrite H. rewrite ! Pos.leb_refl; auto.
+Qed.
-Lemma agree_eval_reg:
- forall r live rs ls,
- agree (reg_live r live) rs ls -> rs#r = ls (assign r).
+Lemma reg_unconstrained_sound:
+ forall r e q,
+ reg_unconstrained r e = true ->
+ EqSet.In q e ->
+ ereg q <> r.
Proof.
- intros. apply H. apply Regset.add_1. auto.
+ unfold reg_unconstrained; intros. red; intros.
+ apply select_reg_charact in H1.
+ assert (EqSet.mem_between (select_reg_l r) (select_reg_h r) e = true).
+ {
+ apply EqSet.mem_between_2 with q; auto.
+ exact (select_reg_l_monotone r).
+ exact (select_reg_h_monotone r).
+ tauto.
+ tauto.
+ }
+ rewrite H2 in H; discriminate.
Qed.
-(** Same, for a list of registers. *)
+Lemma reg_unconstrained_satisf:
+ forall r e rs ls v,
+ reg_unconstrained r e = true ->
+ satisf rs ls e ->
+ satisf (rs#r <- v) ls e.
+Proof.
+ red; intros. rewrite PMap.gso. auto. eapply reg_unconstrained_sound; eauto.
+Qed.
-Lemma agree_eval_regs:
- forall rl live rs ls,
- agree (reg_list_live rl live) rs ls ->
- rs##rl = List.map ls (List.map assign rl).
+Remark select_loc_l_monotone:
+ forall l q1 q2,
+ OrderedEquation'.eq q1 q2 \/ OrderedEquation'.lt q1 q2 ->
+ select_loc_l l q1 = true -> select_loc_l l q2 = true.
Proof.
- induction rl; simpl; intros.
+ unfold select_loc_l; intros. set (lb := OrderedLoc.diff_low_bound l) in *.
+ destruct H.
+ red in H. subst q2; auto.
+ assert (eloc q1 = eloc q2 \/ OrderedLoc.lt (eloc q1) (eloc q2)).
+ red in H. tauto.
+ destruct H1. rewrite <- H1; auto.
+ destruct (OrderedLoc.compare (eloc q2) lb); auto.
+ assert (OrderedLoc.lt (eloc q1) lb) by (eapply OrderedLoc.lt_trans; eauto).
+ destruct (OrderedLoc.compare (eloc q1) lb).
auto.
- f_equal.
- apply agree_eval_reg with live.
- apply agree_reg_list_live with rl. auto.
- eapply IHrl. eexact H.
-Qed.
-
-(** Agreement is insensitive to the current values of the temporary
- machine registers. *)
-
-Lemma agree_exten:
- forall live rs ls ls',
- agree live rs ls ->
- (forall l, Loc.notin l temporaries -> ls' l = ls l) ->
- agree live rs ls'.
-Proof.
- unfold agree; intros.
- rewrite H0. apply H. auto. eapply regalloc_not_temporary; eauto.
-Qed.
-
-Lemma agree_undef_temps:
- forall live rs ls,
- agree live rs ls -> agree live rs (undef_temps ls).
-Proof.
- intros. apply agree_exten with ls; auto.
- intros. apply Locmap.guo; auto.
-Qed.
-
-(** If a register is dead, assigning it an arbitrary value in [rs]
- and leaving [ls] unchanged preserves agreement. (This corresponds
- to an operation over a dead register in the original program
- that is turned into a no-op in the transformed program.) *)
-
-Lemma agree_assign_dead:
- forall live r rs ls v,
- ~Regset.In r live ->
- agree live rs ls ->
- agree live (rs#r <- v) ls.
-Proof.
- unfold agree; intros.
- case (Reg.eq r r0); intro.
- subst r0. contradiction.
- rewrite Regmap.gso; auto.
-Qed.
-
-(** Setting [r] to value [v] in [rs]
- and simultaneously setting [assign r] to value [v] in [ls]
- preserves agreement, provided that all live registers except [r]
- are mapped to locations other than that of [r]. *)
-
-Lemma agree_assign_live:
- forall live r rs ls v,
- (forall s,
- Regset.In s live -> s <> r -> assign s <> assign r) ->
- agree (reg_dead r live) rs ls ->
- agree live (rs#r <- v) (Locmap.set (assign r) v ls).
-Proof.
- unfold agree; intros. rewrite Regmap.gsspec.
- destruct (peq r0 r).
- subst r0. rewrite Locmap.gss. auto.
- rewrite Locmap.gso. apply H0. apply Regset.remove_2; auto.
- eapply regalloc_noteq_diff. eauto. apply sym_not_equal. apply H. auto. auto.
-Qed.
-
-(** This is a special case of the previous lemma where the value [v]
- being stored is not arbitrary, but is the value of
- another register [arg]. (This corresponds to a register-register move
- instruction.) In this case, the condition can be weakened:
- it suffices that all live registers except [arg] and [res]
- are mapped to locations other than that of [res]. *)
-
-Lemma agree_move_live:
- forall live arg res rs (ls: locset),
- (forall r,
- Regset.In r live -> r <> res -> r <> arg ->
- assign r <> assign res) ->
- agree (reg_live arg (reg_dead res live)) rs ls ->
- agree live (rs#res <- (rs#arg)) (Locmap.set (assign res) (ls (assign arg)) ls).
-Proof.
- unfold agree; intros. rewrite Regmap.gsspec. destruct (peq r res).
- subst r. rewrite Locmap.gss. apply H0.
- apply Regset.add_1; auto.
- destruct (Reg.eq r arg).
- subst r.
- replace (Locmap.set (assign res) (ls (assign arg)) ls (assign arg))
- with (ls (assign arg)).
- apply H0. apply Regset.add_1. auto.
- symmetry. destruct (Loc.eq (assign arg) (assign res)).
- rewrite <- e. apply Locmap.gss.
- apply Locmap.gso. eapply regalloc_noteq_diff; eauto.
-
- rewrite Locmap.gso. apply H0. apply Regset.add_2. apply Regset.remove_2. auto. auto.
- eapply regalloc_noteq_diff; eauto. apply sym_not_equal. apply H; auto.
-Qed.
-
-(** Yet another special case corresponding to the case of
- a redundant move. *)
-
-Lemma agree_redundant_move_live:
- forall live arg res rs (ls: locset),
- (forall r,
- Regset.In r live -> r <> res -> r <> arg ->
- assign r <> assign res) ->
- agree (reg_live arg (reg_dead res live)) rs ls ->
- assign res = assign arg ->
- agree live (rs#res <- (rs#arg)) ls.
+ eelim OrderedLoc.lt_not_eq; eauto.
+ eelim OrderedLoc.lt_not_eq. eapply OrderedLoc.lt_trans. eexact l1. eexact H2. red; auto.
+Qed.
+
+Remark select_loc_h_monotone:
+ forall l q1 q2,
+ OrderedEquation'.eq q1 q2 \/ OrderedEquation'.lt q2 q1 ->
+ select_loc_h l q1 = true -> select_loc_h l q2 = true.
Proof.
- intros.
- apply agree_exten with (Locmap.set (assign res) (ls (assign arg)) ls).
- eapply agree_move_live; eauto.
- intros. symmetry. rewrite H1. destruct (Loc.eq l (assign arg)).
- subst l. apply Locmap.gss.
- apply Locmap.gso. eapply regalloc_noteq_diff; eauto.
-Qed.
-
-(** This complicated lemma states agreement between the states after
- a function call, provided that the states before the call agree
- and that calling conventions are respected. *)
-
-Lemma agree_postcall:
- forall live args ros res rs v (ls: locset),
- (forall r,
- Regset.In r live -> r <> res ->
- ~(In (assign r) destroyed_at_call)) ->
- (forall r,
- Regset.In r live -> r <> res -> assign r <> assign res) ->
- agree (reg_list_live args (reg_sum_live ros (reg_dead res live))) rs ls ->
- agree live (rs#res <- v) (Locmap.set (assign res) v (postcall_locs ls)).
+ unfold select_loc_h; intros. set (lb := OrderedLoc.diff_high_bound l) in *.
+ destruct H.
+ red in H. subst q2; auto.
+ assert (eloc q2 = eloc q1 \/ OrderedLoc.lt (eloc q2) (eloc q1)).
+ red in H. tauto.
+ destruct H1. rewrite H1; auto.
+ destruct (OrderedLoc.compare (eloc q2) lb); auto.
+ assert (OrderedLoc.lt lb (eloc q1)) by (eapply OrderedLoc.lt_trans; eauto).
+ destruct (OrderedLoc.compare (eloc q1) lb).
+ eelim OrderedLoc.lt_not_eq. eapply OrderedLoc.lt_trans. eexact l1. eexact H2. red; auto.
+ eelim OrderedLoc.lt_not_eq. eexact H2. apply OrderedLoc.eq_sym; auto.
+ auto.
+Qed.
+
+Remark select_loc_charact:
+ forall l q,
+ select_loc_l l q = false \/ select_loc_h l q = false <-> Loc.diff l (eloc q).
Proof.
- intros.
- assert (agree (reg_dead res live) rs ls).
- apply agree_reg_sum_live with ros.
- apply agree_reg_list_live with args. assumption.
- red; intros. rewrite Regmap.gsspec. destruct (peq r res).
- subst r. rewrite Locmap.gss. auto.
- rewrite Locmap.gso. transitivity (ls (assign r)).
- apply H2. apply Regset.remove_2; auto.
- unfold postcall_locs.
- assert (~In (assign r) temporaries).
- apply Loc.notin_not_in. eapply regalloc_not_temporary; eauto.
- assert (~In (assign r) destroyed_at_call).
- apply H. auto. auto.
- caseEq (assign r); auto. intros m ASG. rewrite <- ASG.
- destruct (In_dec Loc.eq (assign r) temporaries). contradiction.
- destruct (In_dec Loc.eq (assign r) destroyed_at_call). contradiction.
+ unfold select_loc_l, select_loc_h; intros; split; intros.
+ apply OrderedLoc.outside_interval_diff.
+ destruct H.
+ left. destruct (OrderedLoc.compare (eloc q) (OrderedLoc.diff_low_bound l)); assumption || discriminate.
+ right. destruct (OrderedLoc.compare (eloc q) (OrderedLoc.diff_high_bound l)); assumption || discriminate.
+ exploit OrderedLoc.diff_outside_interval. eauto.
+ intros [A | A].
+ left. destruct (OrderedLoc.compare (eloc q) (OrderedLoc.diff_low_bound l)).
+ auto.
+ eelim OrderedLoc.lt_not_eq; eauto.
+ eelim OrderedLoc.lt_not_eq. eapply OrderedLoc.lt_trans; eauto. red; auto.
+ right. destruct (OrderedLoc.compare (eloc q) (OrderedLoc.diff_high_bound l)).
+ eelim OrderedLoc.lt_not_eq. eapply OrderedLoc.lt_trans; eauto. red; auto.
+ eelim OrderedLoc.lt_not_eq; eauto. apply OrderedLoc.eq_sym; auto.
auto.
- eapply regalloc_noteq_diff; eauto. apply sym_not_eq. auto.
Qed.
-(** Agreement between the initial register set at RTL function entry
- and the location set at LTL function entry. *)
+Lemma loc_unconstrained_sound:
+ forall l e q,
+ loc_unconstrained l e = true ->
+ EqSet.In q e ->
+ Loc.diff l (eloc q).
+Proof.
+ unfold loc_unconstrained; intros.
+ destruct (select_loc_l l q) eqn:SL.
+ destruct (select_loc_h l q) eqn:SH.
+ assert (EqSet2.mem_between (select_loc_l l) (select_loc_h l) (eqs2 e) = true).
+ {
+ apply EqSet2.mem_between_2 with q; auto.
+ exact (select_loc_l_monotone l).
+ exact (select_loc_h_monotone l).
+ apply eqs_same. auto.
+ }
+ rewrite H1 in H; discriminate.
+ apply select_loc_charact; auto.
+ apply select_loc_charact; auto.
+Qed.
+
+Lemma loc_unconstrained_satisf:
+ forall rs ls k r l e v,
+ satisf rs ls (remove_equation (Eq k r l) e) ->
+ loc_unconstrained l (remove_equation (Eq k r l) e) = true ->
+ Val.lessdef (sel_val k rs#r) v ->
+ satisf rs (Locmap.set l v ls) e.
+Proof.
+ intros; red; intros.
+ destruct (OrderedEquation.eq_dec q (Eq k r l)).
+ subst q; simpl. rewrite Locmap.gss. auto.
+ assert (EqSet.In q (remove_equation (Eq k r l) e)).
+ simpl. ESD.fsetdec.
+ rewrite Locmap.gso. apply H; auto. eapply loc_unconstrained_sound; eauto.
+Qed.
+
+Lemma reg_loc_unconstrained_sound:
+ forall r l e q,
+ reg_loc_unconstrained r l e = true ->
+ EqSet.In q e ->
+ ereg q <> r /\ Loc.diff l (eloc q).
+Proof.
+ intros. destruct (andb_prop _ _ H).
+ split. eapply reg_unconstrained_sound; eauto. eapply loc_unconstrained_sound; eauto.
+Qed.
+
+Lemma parallel_assignment_satisf:
+ forall k r l e rs ls v v',
+ Val.lessdef (sel_val k v) v' ->
+ reg_loc_unconstrained r l (remove_equation (Eq k r l) e) = true ->
+ satisf rs ls (remove_equation (Eq k r l) e) ->
+ satisf (rs#r <- v) (Locmap.set l v' ls) e.
+Proof.
+ intros; red; intros.
+ destruct (OrderedEquation.eq_dec q (Eq k r l)).
+ subst q; simpl. rewrite Regmap.gss; rewrite Locmap.gss; auto.
+ assert (EqSet.In q (remove_equation {| ekind := k; ereg := r; eloc := l |} e)).
+ simpl. ESD.fsetdec.
+ exploit reg_loc_unconstrained_sound; eauto. intros [A B].
+ rewrite Regmap.gso; auto. rewrite Locmap.gso; auto.
+Qed.
+
+Lemma parallel_assignment_satisf_2:
+ forall rs ls res res' oty e e' v v',
+ remove_equations_res res oty res' e = Some e' ->
+ satisf rs ls e' ->
+ reg_unconstrained res e' = true ->
+ forallb (fun l => loc_unconstrained l e') res' = true ->
+ Val.lessdef v v' ->
+ satisf (rs#res <- v) (Locmap.setlist res' (encode_long oty v') ls) e.
+Proof.
+ intros; red; intros.
+ functional inversion H.
+- (* Two 32-bit halves *)
+ subst.
+ set (e' := remove_equation {| ekind := Low; ereg := res; eloc := l2 |}
+ (remove_equation {| ekind := High; ereg := res; eloc := l1 |} e)) in *.
+ simpl in H2. InvBooleans. simpl.
+ destruct (OrderedEquation.eq_dec q (Eq Low res l2)).
+ subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gss.
+ apply Val.loword_lessdef; auto.
+ destruct (OrderedEquation.eq_dec q (Eq High res l1)).
+ subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gso by auto. rewrite Locmap.gss.
+ apply Val.hiword_lessdef; auto.
+ assert (EqSet.In q e'). unfold e', remove_equation; simpl; ESD.fsetdec.
+ rewrite Regmap.gso. rewrite ! Locmap.gso. auto.
+ eapply loc_unconstrained_sound; eauto.
+ eapply loc_unconstrained_sound; eauto.
+ eapply reg_unconstrained_sound; eauto.
+- (* One location *)
+ subst. simpl in H2. InvBooleans.
+ replace (encode_long oty v') with (v' :: nil).
+ set (e' := remove_equation {| ekind := Full; ereg := res; eloc := l1 |} e) in *.
+ destruct (OrderedEquation.eq_dec q (Eq Full res l1)).
+ subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gss. auto.
+ assert (EqSet.In q e'). unfold e', remove_equation; simpl. ESD.fsetdec.
+ simpl. rewrite Regmap.gso. rewrite Locmap.gso. auto.
+ eapply loc_unconstrained_sound; eauto.
+ eapply reg_unconstrained_sound; eauto.
+ destruct oty as [[]|]; reflexivity || contradiction.
+Qed.
+
+Lemma in_subst_reg:
+ forall r1 r2 q (e: eqs),
+ EqSet.In q e ->
+ ereg q = r1 /\ EqSet.In (Eq (ekind q) r2 (eloc q)) (subst_reg r1 r2 e)
+ \/ ereg q <> r1 /\ EqSet.In q (subst_reg r1 r2 e).
+Proof.
+ intros r1 r2 q e0 IN0. unfold subst_reg.
+ set (f := fun (q: EqSet.elt) e => add_equation (Eq (ekind q) r2 (eloc q)) (remove_equation q e)).
+ set (elt := EqSet.elements_between (select_reg_l r1) (select_reg_h r1) e0).
+ assert (IN_ELT: forall q, EqSet.In q elt <-> EqSet.In q e0 /\ ereg q = r1).
+ {
+ intros. unfold elt. rewrite EqSet.elements_between_iff.
+ rewrite select_reg_charact. tauto.
+ exact (select_reg_l_monotone r1).
+ exact (select_reg_h_monotone r1).
+ }
+ set (P := fun e1 e2 =>
+ EqSet.In q e1 ->
+ EqSet.In (Eq (ekind q) r2 (eloc q)) e2).
+ assert (P elt (EqSet.fold f elt e0)).
+ {
+ apply ESP.fold_rec; unfold P; intros.
+ - ESD.fsetdec.
+ - simpl. red in H1. apply H1 in H3. destruct H3.
+ + subst x. ESD.fsetdec.
+ + rewrite ESF.add_iff. rewrite ESF.remove_iff.
+ destruct (OrderedEquation.eq_dec x {| ekind := ekind q; ereg := r2; eloc := eloc q |}); auto.
+ left. subst x; auto.
+ }
+ set (Q := fun e1 e2 =>
+ ~EqSet.In q e1 ->
+ EqSet.In q e2).
+ assert (Q elt (EqSet.fold f elt e0)).
+ {
+ apply ESP.fold_rec; unfold Q; intros.
+ - auto.
+ - simpl. red in H2. rewrite H2 in H4.
+ rewrite ESF.add_iff. rewrite ESF.remove_iff.
+ right. split. apply H3. tauto. tauto.
+ }
+ destruct (ESP.In_dec q elt).
+ left. split. apply IN_ELT. auto. apply H. auto.
+ right. split. red; intros. elim n. rewrite IN_ELT. auto. apply H0. auto.
+Qed.
+
+Lemma subst_reg_satisf:
+ forall src dst rs ls e,
+ satisf rs ls (subst_reg dst src e) ->
+ satisf (rs#dst <- (rs#src)) ls e.
+Proof.
+ intros; red; intros.
+ destruct (in_subst_reg dst src q e H0) as [[A B] | [A B]].
+ subst dst. rewrite Regmap.gss. exploit H; eauto.
+ rewrite Regmap.gso; auto.
+Qed.
+
+Lemma in_subst_reg_kind:
+ forall r1 k1 r2 k2 q (e: eqs),
+ EqSet.In q e ->
+ (ereg q, ekind q) = (r1, k1) /\ EqSet.In (Eq k2 r2 (eloc q)) (subst_reg_kind r1 k1 r2 k2 e)
+ \/ EqSet.In q (subst_reg_kind r1 k1 r2 k2 e).
+Proof.
+ intros r1 k1 r2 k2 q e0 IN0. unfold subst_reg.
+ set (f := fun (q: EqSet.elt) e =>
+ if IndexedEqKind.eq (ekind q) k1
+ then add_equation (Eq k2 r2 (eloc q)) (remove_equation q e)
+ else e).
+ set (elt := EqSet.elements_between (select_reg_l r1) (select_reg_h r1) e0).
+ assert (IN_ELT: forall q, EqSet.In q elt <-> EqSet.In q e0 /\ ereg q = r1).
+ {
+ intros. unfold elt. rewrite EqSet.elements_between_iff.
+ rewrite select_reg_charact. tauto.
+ exact (select_reg_l_monotone r1).
+ exact (select_reg_h_monotone r1).
+ }
+ set (P := fun e1 e2 =>
+ EqSet.In q e1 -> ekind q = k1 ->
+ EqSet.In (Eq k2 r2 (eloc q)) e2).
+ assert (P elt (EqSet.fold f elt e0)).
+ {
+ intros; apply ESP.fold_rec; unfold P; intros.
+ - ESD.fsetdec.
+ - simpl. red in H1. apply H1 in H3. destruct H3.
+ + subst x. unfold f. destruct (IndexedEqKind.eq (ekind q) k1).
+ simpl. ESD.fsetdec. contradiction.
+ + unfold f. destruct (IndexedEqKind.eq (ekind x) k1).
+ simpl. rewrite ESF.add_iff. rewrite ESF.remove_iff.
+ destruct (OrderedEquation.eq_dec x {| ekind := k2; ereg := r2; eloc := eloc q |}); auto.
+ left. subst x; auto.
+ auto.
+ }
+ set (Q := fun e1 e2 =>
+ ~EqSet.In q e1 \/ ekind q <> k1 ->
+ EqSet.In q e2).
+ assert (Q elt (EqSet.fold f elt e0)).
+ {
+ apply ESP.fold_rec; unfold Q; intros.
+ - auto.
+ - unfold f. red in H2. rewrite H2 in H4.
+ destruct (IndexedEqKind.eq (ekind x) k1).
+ simpl. rewrite ESF.add_iff. rewrite ESF.remove_iff.
+ right. split. apply H3. tauto. intuition congruence.
+ apply H3. intuition.
+ }
+ destruct (ESP.In_dec q elt).
+ destruct (IndexedEqKind.eq (ekind q) k1).
+ left. split. f_equal. apply IN_ELT. auto. auto. apply H. auto. auto.
+ right. apply H0. auto.
+ right. apply H0. auto.
+Qed.
+
+Lemma subst_reg_kind_satisf_makelong:
+ forall src1 src2 dst rs ls e,
+ let e1 := subst_reg_kind dst High src1 Full e in
+ let e2 := subst_reg_kind dst Low src2 Full e1 in
+ reg_unconstrained dst e2 = true ->
+ satisf rs ls e2 ->
+ satisf (rs#dst <- (Val.longofwords rs#src1 rs#src2)) ls e.
+Proof.
+ intros; red; intros.
+ destruct (in_subst_reg_kind dst High src1 Full q e H1) as [[A B] | B]; fold e1 in B.
+ destruct (in_subst_reg_kind dst Low src2 Full _ e1 B) as [[C D] | D]; fold e2 in D.
+ simpl in C; simpl in D. inv C.
+ inversion A. rewrite H3; rewrite H4. rewrite Regmap.gss.
+ apply Val.lessdef_trans with (rs#src1).
+ simpl. destruct (rs#src1); simpl; auto. destruct (rs#src2); simpl; auto.
+ rewrite Int64.hi_ofwords. auto.
+ exploit H0. eexact D. simpl. auto.
+ destruct (in_subst_reg_kind dst Low src2 Full q e1 B) as [[C D] | D]; fold e2 in D.
+ inversion C. rewrite H3; rewrite H4. rewrite Regmap.gss.
+ apply Val.lessdef_trans with (rs#src2).
+ simpl. destruct (rs#src1); simpl; auto. destruct (rs#src2); simpl; auto.
+ rewrite Int64.lo_ofwords. auto.
+ exploit H0. eexact D. simpl. auto.
+ rewrite Regmap.gso. apply H0; auto. eapply reg_unconstrained_sound; eauto.
+Qed.
+
+Lemma subst_reg_kind_satisf_lowlong:
+ forall src dst rs ls e,
+ let e1 := subst_reg_kind dst Full src Low e in
+ reg_unconstrained dst e1 = true ->
+ satisf rs ls e1 ->
+ satisf (rs#dst <- (Val.loword rs#src)) ls e.
+Proof.
+ intros; red; intros.
+ destruct (in_subst_reg_kind dst Full src Low q e H1) as [[A B] | B]; fold e1 in B.
+ inversion A. rewrite H3; rewrite H4. simpl. rewrite Regmap.gss.
+ exploit H0. eexact B. simpl. auto.
+ rewrite Regmap.gso. apply H0; auto. eapply reg_unconstrained_sound; eauto.
+Qed.
+
+Lemma subst_reg_kind_satisf_highlong:
+ forall src dst rs ls e,
+ let e1 := subst_reg_kind dst Full src High e in
+ reg_unconstrained dst e1 = true ->
+ satisf rs ls e1 ->
+ satisf (rs#dst <- (Val.hiword rs#src)) ls e.
+Proof.
+ intros; red; intros.
+ destruct (in_subst_reg_kind dst Full src High q e H1) as [[A B] | B]; fold e1 in B.
+ inversion A. rewrite H3; rewrite H4. simpl. rewrite Regmap.gss.
+ exploit H0. eexact B. simpl. auto.
+ rewrite Regmap.gso. apply H0; auto. eapply reg_unconstrained_sound; eauto.
+Qed.
+
+Module ESF2 := FSetFacts.Facts(EqSet2).
+Module ESP2 := FSetProperties.Properties(EqSet2).
+Module ESD2 := FSetDecide.Decide(EqSet2).
+
+Lemma in_subst_loc:
+ forall l1 l2 q (e e': eqs),
+ EqSet.In q e ->
+ subst_loc l1 l2 e = Some e' ->
+ (eloc q = l1 /\ EqSet.In (Eq (ekind q) (ereg q) l2) e') \/ (Loc.diff l1 (eloc q) /\ EqSet.In q e').
+Proof.
+ intros l1 l2 q e0 e0'.
+ unfold subst_loc.
+ set (f := fun (q0 : EqSet2.elt) (opte : option eqs) =>
+ match opte with
+ | Some e =>
+ if Loc.eq l1 (eloc q0)
+ then
+ Some
+ (add_equation {| ekind := ekind q0; ereg := ereg q0; eloc := l2 |}
+ (remove_equation q0 e))
+ else None
+ | None => None
+ end).
+ set (elt := EqSet2.elements_between (select_loc_l l1) (select_loc_h l1) (eqs2 e0)).
+ intros IN SUBST.
+ set (P := fun e1 (opte: option eqs) =>
+ match opte with
+ | None => True
+ | Some e2 =>
+ EqSet2.In q e1 ->
+ eloc q = l1 /\ EqSet.In (Eq (ekind q) (ereg q) l2) e2
+ end).
+ assert (P elt (EqSet2.fold f elt (Some e0))).
+ {
+ apply ESP2.fold_rec; unfold P; intros.
+ - ESD2.fsetdec.
+ - destruct a as [e2|]; simpl; auto.
+ destruct (Loc.eq l1 (eloc x)); auto.
+ unfold add_equation, remove_equation; simpl.
+ red in H1. rewrite H1. intros [A|A].
+ + subst x. split. auto. ESD.fsetdec.
+ + exploit H2; eauto. intros [B C]. split. auto.
+ rewrite ESF.add_iff. rewrite ESF.remove_iff.
+ destruct (OrderedEquation.eq_dec x {| ekind := ekind q; ereg := ereg q; eloc := l2 |}).
+ left. rewrite e1; auto.
+ right; auto.
+ }
+ set (Q := fun e1 (opte: option eqs) =>
+ match opte with
+ | None => True
+ | Some e2 => ~EqSet2.In q e1 -> EqSet.In q e2
+ end).
+ assert (Q elt (EqSet2.fold f elt (Some e0))).
+ {
+ apply ESP2.fold_rec; unfold Q; intros.
+ - auto.
+ - destruct a as [e2|]; simpl; auto.
+ destruct (Loc.eq l1 (eloc x)); auto.
+ red in H2. rewrite H2; intros.
+ unfold add_equation, remove_equation; simpl.
+ rewrite ESF.add_iff. rewrite ESF.remove_iff.
+ right; split. apply H3. tauto. tauto.
+ }
+ rewrite SUBST in H; rewrite SUBST in H0; simpl in *.
+ destruct (ESP2.In_dec q elt).
+ left. apply H; auto.
+ right. split; auto.
+ rewrite <- select_loc_charact.
+ destruct (select_loc_l l1 q) eqn: LL; auto.
+ destruct (select_loc_h l1 q) eqn: LH; auto.
+ elim n. eapply EqSet2.elements_between_iff.
+ exact (select_loc_l_monotone l1).
+ exact (select_loc_h_monotone l1).
+ split. apply eqs_same; auto. auto.
+Qed.
+
+Lemma subst_loc_satisf:
+ forall src dst rs ls e e',
+ subst_loc dst src e = Some e' ->
+ satisf rs ls e' ->
+ satisf rs (Locmap.set dst (ls src) ls) e.
+Proof.
+ intros; red; intros.
+ exploit in_subst_loc; eauto. intros [[A B] | [A B]].
+ subst dst. rewrite Locmap.gss. apply (H0 _ B).
+ rewrite Locmap.gso; auto.
+Qed.
+
+Lemma can_undef_sound:
+ forall e ml q,
+ can_undef ml e = true -> EqSet.In q e -> Loc.notin (eloc q) (map R ml).
+Proof.
+ induction ml; simpl; intros.
+ tauto.
+ InvBooleans. split.
+ apply Loc.diff_sym. eapply loc_unconstrained_sound; eauto.
+ eauto.
+Qed.
+
+Lemma undef_regs_outside:
+ forall ml ls l,
+ Loc.notin l (map R ml) -> undef_regs ml ls l = ls l.
+Proof.
+ induction ml; simpl; intros. auto.
+ rewrite Locmap.gso. apply IHml. tauto. apply Loc.diff_sym. tauto.
+Qed.
+
+Lemma can_undef_satisf:
+ forall ml e rs ls,
+ can_undef ml e = true ->
+ satisf rs ls e ->
+ satisf rs (undef_regs ml ls) e.
+Proof.
+ intros; red; intros. rewrite undef_regs_outside. eauto.
+ eapply can_undef_sound; eauto.
+Qed.
+
+Lemma can_undef_except_sound:
+ forall lx e ml q,
+ can_undef_except lx ml e = true -> EqSet.In q e -> Loc.diff (eloc q) lx -> Loc.notin (eloc q) (map R ml).
+Proof.
+ induction ml; simpl; intros.
+ tauto.
+ InvBooleans. split.
+ destruct (orb_true_elim _ _ H2).
+ apply proj_sumbool_true in e0. congruence.
+ apply Loc.diff_sym. eapply loc_unconstrained_sound; eauto.
+ eapply IHml; eauto.
+Qed.
+
+Lemma subst_loc_undef_satisf:
+ forall src dst rs ls ml e e',
+ subst_loc dst src e = Some e' ->
+ can_undef_except dst ml e = true ->
+ satisf rs ls e' ->
+ satisf rs (Locmap.set dst (ls src) (undef_regs ml ls)) e.
+Proof.
+ intros; red; intros.
+ exploit in_subst_loc; eauto. intros [[A B] | [A B]].
+ rewrite A. rewrite Locmap.gss. apply (H1 _ B).
+ rewrite Locmap.gso; auto. rewrite undef_regs_outside. eauto.
+ eapply can_undef_except_sound; eauto. apply Loc.diff_sym; auto.
+Qed.
+
+Lemma transfer_use_def_satisf:
+ forall args res args' res' und e e' rs ls,
+ transfer_use_def args res args' res' und e = Some e' ->
+ satisf rs ls e' ->
+ Val.lessdef_list rs##args (reglist ls args') /\
+ (forall v v', Val.lessdef v v' ->
+ satisf (rs#res <- v) (Locmap.set (R res') v' (undef_regs und ls)) e).
+Proof.
+ unfold transfer_use_def; intros. MonadInv.
+ split. eapply add_equations_lessdef; eauto.
+ intros. eapply parallel_assignment_satisf; eauto. assumption.
+ eapply can_undef_satisf; eauto.
+ eapply add_equations_satisf; eauto.
+Qed.
+
+Lemma add_equations_res_lessdef:
+ forall r oty ll e e' rs ls,
+ add_equations_res r oty ll e = Some e' ->
+ satisf rs ls e' ->
+ Val.lessdef_list (encode_long oty rs#r) (map ls ll).
+Proof.
+ intros. functional inversion H.
+- subst. simpl. constructor.
+ eapply add_equation_lessdef with (q := Eq High r l1).
+ eapply add_equation_satisf. eauto.
+ constructor.
+ eapply add_equation_lessdef with (q := Eq Low r l2). eauto.
+ constructor.
+- subst. replace (encode_long oty rs#r) with (rs#r :: nil). simpl. constructor; auto.
+ eapply add_equation_lessdef with (q := Eq Full r l1); eauto.
+ destruct oty as [[]|]; reflexivity || contradiction.
+Qed.
+
+Definition callee_save_loc (l: loc) :=
+ match l with
+ | R r => ~(In r destroyed_at_call)
+ | S sl ofs ty => sl <> Outgoing
+ end.
+
+Definition agree_callee_save (ls1 ls2: locset) : Prop :=
+ forall l, callee_save_loc l -> ls1 l = ls2 l.
+
+Lemma return_regs_agree_callee_save:
+ forall caller callee,
+ agree_callee_save caller (return_regs caller callee).
+Proof.
+ intros; red; intros. unfold return_regs. red in H.
+ destruct l.
+ rewrite pred_dec_false; auto.
+ destruct sl; auto || congruence.
+Qed.
+
+Lemma no_caller_saves_sound:
+ forall e q,
+ no_caller_saves e = true ->
+ EqSet.In q e ->
+ callee_save_loc (eloc q).
+Proof.
+ unfold no_caller_saves, callee_save_loc; intros.
+ exploit EqSet.for_all_2; eauto.
+ hnf. intros. simpl in H1. rewrite H1. auto.
+ lazy beta. destruct (eloc q).
+ intros; red; intros. destruct (orb_true_elim _ _ H1); InvBooleans.
+ eapply int_callee_save_not_destroyed; eauto.
+ apply index_int_callee_save_pos2. omega.
+ eapply float_callee_save_not_destroyed; eauto.
+ apply index_float_callee_save_pos2. omega.
+ destruct sl; congruence.
+Qed.
+
+Lemma function_return_satisf:
+ forall rs ls_before ls_after res res' sg e e' v,
+ res' = map R (loc_result sg) ->
+ remove_equations_res res (sig_res sg) res' e = Some e' ->
+ satisf rs ls_before e' ->
+ forallb (fun l => reg_loc_unconstrained res l e') res' = true ->
+ no_caller_saves e' = true ->
+ Val.lessdef_list (encode_long (sig_res sg) v) (map ls_after res') ->
+ agree_callee_save ls_before ls_after ->
+ satisf (rs#res <- v) ls_after e.
+Proof.
+ intros; red; intros.
+ functional inversion H0.
+- subst. rewrite <- H11 in *. unfold encode_long in H4. rewrite <- H7 in H4.
+ simpl in H4. inv H4. inv H14.
+ set (e' := remove_equation {| ekind := Low; ereg := res; eloc := l2 |}
+ (remove_equation {| ekind := High; ereg := res; eloc := l1 |} e)) in *.
+ simpl in H2. InvBooleans.
+ destruct (OrderedEquation.eq_dec q (Eq Low res l2)).
+ subst q; simpl. rewrite Regmap.gss. auto.
+ destruct (OrderedEquation.eq_dec q (Eq High res l1)).
+ subst q; simpl. rewrite Regmap.gss. auto.
+ assert (EqSet.In q e'). unfold e', remove_equation; simpl; ESD.fsetdec.
+ exploit reg_loc_unconstrained_sound. eexact H. eauto. intros [A B].
+ exploit reg_loc_unconstrained_sound. eexact H2. eauto. intros [C D].
+ rewrite Regmap.gso; auto.
+ exploit no_caller_saves_sound; eauto. intros.
+ red in H5. rewrite <- H5; auto.
+- subst. rewrite <- H11 in *.
+ replace (encode_long (sig_res sg) v) with (v :: nil) in H4.
+ simpl in H4. inv H4.
+ simpl in H2. InvBooleans.
+ set (e' := remove_equation {| ekind := Full; ereg := res; eloc := l1 |} e) in *.
+ destruct (OrderedEquation.eq_dec q (Eq Full res l1)).
+ subst q; simpl. rewrite Regmap.gss; auto.
+ assert (EqSet.In q e'). unfold e', remove_equation; simpl. ESD.fsetdec.
+ exploit reg_loc_unconstrained_sound; eauto. intros [A B].
+ rewrite Regmap.gso; auto.
+ exploit no_caller_saves_sound; eauto. intros.
+ red in H5. rewrite <- H5; auto.
+ destruct (sig_res sg) as [[]|]; reflexivity || contradiction.
+Qed.
-Lemma agree_init_regs:
- forall live rl vl,
- (forall r1 r2,
- In r1 rl -> Regset.In r2 live -> r1 <> r2 ->
- assign r1 <> assign r2) ->
- agree live (RTL.init_regs vl rl)
- (LTL.init_locs vl (List.map assign rl)).
+Lemma compat_left_sound:
+ forall r l e q,
+ compat_left r l e = true -> EqSet.In q e -> ereg q = r -> ekind q = Full /\ eloc q = l.
Proof.
- intro live.
- assert (agree live (Regmap.init Vundef) (Locmap.init Vundef)).
- red; intros. rewrite Regmap.gi. auto.
- induction rl; simpl; intros.
+ unfold compat_left; intros.
+ rewrite EqSet.for_all_between_iff in H.
+ apply select_reg_charact in H1. destruct H1.
+ exploit H; eauto. intros.
+ destruct (ekind q); try discriminate.
+ destruct (Loc.eq l (eloc q)); try discriminate.
auto.
- destruct vl. auto.
- assert (agree live (init_regs vl rl) (init_locs vl (map assign rl))).
- apply IHrl. intros. apply H0. tauto. auto. auto.
- red; intros. rewrite Regmap.gsspec. destruct (peq r a).
- subst r. rewrite Locmap.gss. auto.
- rewrite Locmap.gso. apply H1; auto.
- eapply regalloc_noteq_diff; eauto.
+ intros. subst x2. auto.
+ exact (select_reg_l_monotone r).
+ exact (select_reg_h_monotone r).
Qed.
-Lemma agree_parameters:
- forall vl,
- let params := f.(RTL.fn_params) in
- agree (live0 f flive)
- (RTL.init_regs vl params)
- (LTL.init_locs vl (List.map assign params)).
+Lemma compat_left2_sound:
+ forall r l1 l2 e q,
+ compat_left2 r l1 l2 e = true -> EqSet.In q e -> ereg q = r ->
+ (ekind q = High /\ eloc q = l1) \/ (ekind q = Low /\ eloc q = l2).
Proof.
- intros. apply agree_init_regs.
- intros. eapply regalloc_correct_3; eauto.
+ unfold compat_left2; intros.
+ rewrite EqSet.for_all_between_iff in H.
+ apply select_reg_charact in H1. destruct H1.
+ exploit H; eauto. intros.
+ destruct (ekind q); try discriminate.
+ InvBooleans. auto.
+ InvBooleans. auto.
+ intros. subst x2. auto.
+ exact (select_reg_l_monotone r).
+ exact (select_reg_h_monotone r).
Qed.
-End AGREE.
+Lemma val_hiword_longofwords:
+ forall v1 v2, Val.lessdef (Val.hiword (Val.longofwords v1 v2)) v1.
+Proof.
+ intros. destruct v1; simpl; auto. destruct v2; auto. unfold Val.hiword.
+ rewrite Int64.hi_ofwords. auto.
+Qed.
-(** * Preservation of semantics *)
+Lemma val_loword_longofwords:
+ forall v1 v2, Val.lessdef (Val.loword (Val.longofwords v1 v2)) v2.
+Proof.
+ intros. destruct v1; simpl; auto. destruct v2; auto. unfold Val.loword.
+ rewrite Int64.lo_ofwords. auto.
+Qed.
-(** We now show that the LTL code reflecting register allocation has
- the same semantics as the original RTL code. We start with
- standard properties of translated functions and
- global environments in the original and translated code. *)
+Lemma compat_entry_satisf:
+ forall rl tyl ll e,
+ compat_entry rl tyl ll e = true ->
+ forall vl ls,
+ Val.lessdef_list vl (decode_longs tyl (map ls ll)) ->
+ satisf (init_regs vl rl) ls e.
+Proof.
+ intros until e. functional induction (compat_entry rl tyl ll e); intros.
+- (* no params *)
+ simpl. red; intros. rewrite Regmap.gi. destruct (ekind q); simpl; auto.
+- (* a param of type Tlong *)
+ InvBooleans. simpl in H0. inv H0. simpl.
+ red; intros. rewrite Regmap.gsspec. destruct (peq (ereg q) r1).
+ exploit compat_left2_sound; eauto.
+ intros [[A B] | [A B]]; rewrite A; rewrite B; simpl.
+ apply Val.lessdef_trans with (Val.hiword (Val.longofwords (ls l1) (ls l2))).
+ apply Val.hiword_lessdef; auto. apply val_hiword_longofwords.
+ apply Val.lessdef_trans with (Val.loword (Val.longofwords (ls l1) (ls l2))).
+ apply Val.loword_lessdef; auto. apply val_loword_longofwords.
+ eapply IHb; eauto.
+- (* a param of type Tint *)
+ InvBooleans. simpl in H0. inv H0. simpl.
+ red; intros. rewrite Regmap.gsspec. destruct (peq (ereg q) r1).
+ exploit compat_left_sound; eauto. intros [A B]. rewrite A; rewrite B; auto.
+ eapply IHb; eauto.
+- (* a param of type Tfloat *)
+ InvBooleans. simpl in H0. inv H0. simpl.
+ red; intros. rewrite Regmap.gsspec. destruct (peq (ereg q) r1).
+ exploit compat_left_sound; eauto. intros [A B]. rewrite A; rewrite B; auto.
+ eapply IHb; eauto.
+- (* error case *)
+ discriminate.
+Qed.
+
+Lemma call_regs_param_values:
+ forall sg ls,
+ map (call_regs ls) (loc_parameters sg) = map ls (loc_arguments sg).
+Proof.
+ intros. unfold loc_parameters. rewrite list_map_compose.
+ apply list_map_exten; intros. unfold call_regs, parameter_of_argument.
+ exploit loc_arguments_acceptable; eauto. unfold loc_argument_acceptable.
+ destruct x; auto. destruct sl; tauto.
+Qed.
+
+Lemma return_regs_arg_values:
+ forall sg ls1 ls2,
+ tailcall_is_possible sg = true ->
+ map (return_regs ls1 ls2) (loc_arguments sg) = map ls2 (loc_arguments sg).
+Proof.
+ intros. apply list_map_exten; intros.
+ exploit loc_arguments_acceptable; eauto.
+ exploit tailcall_is_possible_correct; eauto.
+ unfold loc_argument_acceptable, return_regs.
+ destruct x; intros.
+ rewrite pred_dec_true; auto.
+ contradiction.
+Qed.
+
+Lemma find_function_tailcall:
+ forall tge ros ls1 ls2,
+ ros_compatible_tailcall ros = true ->
+ find_function tge ros (return_regs ls1 ls2) = find_function tge ros ls2.
+Proof.
+ unfold ros_compatible_tailcall, find_function; intros.
+ destruct ros as [r|id]; auto.
+ unfold return_regs. destruct (in_dec mreg_eq r destroyed_at_call); simpl in H.
+ auto. congruence.
+Qed.
+
+(** * Properties of the dataflow analysis *)
+
+Lemma analyze_successors:
+ forall f bsh an pc bs s e,
+ analyze f bsh = Some an ->
+ bsh!pc = Some bs ->
+ In s (successors_block_shape bs) ->
+ an!!pc = OK e ->
+ exists e', transfer f bsh s an!!s = OK e' /\ EqSet.Subset e' e.
+Proof.
+ unfold analyze; intros. exploit DS.fixpoint_solution; eauto.
+ instantiate (1 := pc). instantiate (1 := s).
+ unfold successors_list. rewrite PTree.gmap1. rewrite H0. simpl. auto.
+ rewrite H2. unfold DS.L.ge. destruct (transfer f bsh s an#s); intros.
+ exists e0; auto.
+ contradiction.
+Qed.
+
+Lemma satisf_successors:
+ forall f bsh an pc bs s e rs ls,
+ analyze f bsh = Some an ->
+ bsh!pc = Some bs ->
+ In s (successors_block_shape bs) ->
+ an!!pc = OK e ->
+ satisf rs ls e ->
+ exists e', transfer f bsh s an!!s = OK e' /\ satisf rs ls e'.
+Proof.
+ intros. exploit analyze_successors; eauto. intros [e' [A B]].
+ exists e'; split; auto. eapply satisf_incr; eauto.
+Qed.
+
+(** Inversion on [transf_function] *)
+
+Inductive transf_function_spec (f: RTL.function) (tf: LTL.function) : Prop :=
+ | transf_function_spec_intro:
+ forall tyenv an mv k e1 e2,
+ wt_function f tyenv ->
+ analyze f (pair_codes f tf) = Some an ->
+ (LTL.fn_code tf)!(LTL.fn_entrypoint tf) = Some(expand_moves mv (Lbranch (RTL.fn_entrypoint f) :: k)) ->
+ wf_moves mv ->
+ transfer f (pair_codes f tf) (RTL.fn_entrypoint f) an!!(RTL.fn_entrypoint f) = OK e1 ->
+ track_moves mv e1 = Some e2 ->
+ compat_entry (RTL.fn_params f) (sig_args (RTL.fn_sig f)) (loc_parameters (fn_sig tf)) e2 = true ->
+ can_undef destroyed_at_function_entry e2 = true ->
+ RTL.fn_stacksize f = LTL.fn_stacksize tf ->
+ RTL.fn_sig f = LTL.fn_sig tf ->
+ transf_function_spec f tf.
+
+Lemma transf_function_inv:
+ forall f tf,
+ transf_function f = OK tf ->
+ transf_function_spec f tf.
+Proof.
+ unfold transf_function; intros.
+ destruct (type_function f) as [tyenv|] eqn:TY; try discriminate.
+ destruct (regalloc f); try discriminate.
+ destruct (check_function f f0) as [] eqn:?; inv H.
+ unfold check_function in Heqr.
+ destruct (analyze f (pair_codes f tf)) as [an|] eqn:?; try discriminate.
+ monadInv Heqr.
+ destruct (check_entrypoints_aux f tf x) as [y|] eqn:?; try discriminate.
+ unfold check_entrypoints_aux, pair_entrypoints in Heqo0. MonadInv.
+ exploit extract_moves_sound; eauto. intros [A B]. subst b.
+ exploit check_succ_sound; eauto. intros [k EQ1]. subst b0.
+ econstructor; eauto. eapply type_function_correct; eauto. congruence.
+Qed.
+
+Lemma invert_code:
+ forall f tf pc i opte e,
+ (RTL.fn_code f)!pc = Some i ->
+ transfer f (pair_codes f tf) pc opte = OK e ->
+ exists eafter, exists bsh, exists bb,
+ opte = OK eafter /\
+ (pair_codes f tf)!pc = Some bsh /\
+ (LTL.fn_code tf)!pc = Some bb /\
+ expand_block_shape bsh i bb /\
+ transfer_aux f bsh eafter = Some e.
+Proof.
+ intros. destruct opte as [eafter|]; simpl in H0; try discriminate. exists eafter.
+ destruct (pair_codes f tf)!pc as [bsh|] eqn:?; try discriminate. exists bsh.
+ exploit matching_instr_block; eauto. intros [bb [A B]].
+ destruct (transfer_aux f bsh eafter) as [e1|] eqn:?; inv H0.
+ exists bb. auto.
+Qed.
+
+(** * Semantic preservation *)
Section PRESERVATION.
@@ -452,348 +1384,726 @@ Lemma sig_function_translated:
transf_fundef f = OK tf ->
LTL.funsig tf = RTL.funsig f.
Proof.
- intros f tf. destruct f; simpl.
- unfold transf_function.
- destruct (type_function f).
- destruct (analyze f).
- destruct (regalloc f t).
- intro. monadInv H. inv EQ. auto.
- simpl; congruence. simpl; congruence. simpl; congruence.
- intro EQ; inv EQ. auto.
-Qed.
-
-(** The proof of semantic preservation is a simulation argument
- based on diagrams of the following form:
-<<
- st1 --------------- st2
- | |
- t| |t
- | |
- v v
- st1'--------------- st2'
->>
- Hypotheses: the left vertical arrow represents a transition in the
- original RTL code. The top horizontal bar is the [match_states]
- relation defined below. It implies agreement between
- the RTL register map [rs] and the LTL location map [ls]
- over the pseudo-registers live before the RTL instruction at [pc].
-
- Conclusions: the right vertical arrow is an [exec_instrs] transition
- in the LTL code generated by translation of the current function.
- The bottom horizontal bar is the [match_states] relation.
-*)
-
-Inductive match_stackframes: list RTL.stackframe -> list LTL.stackframe -> Prop :=
- | match_stackframes_nil:
- match_stackframes nil nil
+ intros; destruct f; monadInv H.
+ destruct (transf_function_inv _ _ EQ). simpl; auto.
+ auto.
+Qed.
+
+Lemma find_function_translated:
+ forall ros rs fd ros' e e' ls,
+ RTL.find_function ge ros rs = Some fd ->
+ add_equation_ros ros ros' e = Some e' ->
+ satisf rs ls e' ->
+ exists tfd,
+ LTL.find_function tge ros' ls = Some tfd /\ transf_fundef fd = OK tfd.
+Proof.
+ unfold RTL.find_function, LTL.find_function; intros.
+ destruct ros as [r|id]; destruct ros' as [r'|id']; simpl in H0; MonadInv.
+ (* two regs *)
+ exploit add_equation_lessdef; eauto. intros LD. inv LD.
+ eapply functions_translated; eauto.
+ rewrite <- H2 in H. simpl in H. congruence.
+ (* two symbols *)
+ rewrite symbols_preserved. rewrite Heqo.
+ eapply function_ptr_translated; eauto.
+Qed.
+
+Lemma exec_moves:
+ forall mv rs s f sp bb m e e' ls,
+ track_moves mv e = Some e' ->
+ wf_moves mv ->
+ satisf rs ls e' ->
+ exists ls',
+ star step tge (Block s f sp (expand_moves mv bb) ls m)
+ E0 (Block s f sp bb ls' m)
+ /\ satisf rs ls' e.
+Proof.
+Opaque destroyed_by_op.
+ induction mv; simpl; intros.
+ (* base *)
+- unfold expand_moves; simpl. inv H. exists ls; split. apply star_refl. auto.
+ (* step *)
+- destruct a as [src dst]. unfold expand_moves. simpl.
+ destruct (track_moves mv e) as [e1|] eqn:?; MonadInv.
+ assert (wf_moves mv). red; intros. apply H0; auto with coqlib.
+ destruct src as [rsrc | ssrc]; destruct dst as [rdst | sdst].
+ (* reg-reg *)
++ exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
+ intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto.
+ econstructor. simpl. eauto. auto. auto.
+ (* reg->stack *)
++ exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
+ intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto.
+ econstructor. simpl. eauto. auto.
+ (* stack->reg *)
++ simpl in Heqb. exploit IHmv; eauto. eapply subst_loc_undef_satisf; eauto.
+ intros [ls' [A B]]. exists ls'; split; auto. eapply star_left; eauto.
+ econstructor. auto. auto.
+ (* stack->stack *)
++ exploit H0; auto with coqlib. unfold wf_move. tauto.
+Qed.
+
+(** The simulation relation *)
+
+Inductive match_stackframes: list RTL.stackframe -> list LTL.stackframe -> signature -> Prop :=
+ | match_stackframes_nil: forall sg,
+ sg.(sig_res) = Some Tint ->
+ match_stackframes nil nil sg
| match_stackframes_cons:
- forall s ts res f sp pc rs ls env live assign,
- match_stackframes s ts ->
- wt_function f env ->
- analyze f = Some live ->
- regalloc f live (live0 f live) env = Some assign ->
- (forall rv,
- agree assign (transfer f pc live!!pc)
- (rs#res <- rv)
- (Locmap.set (assign res) rv ls)) ->
+ forall res f sp pc rs s tf bb ls ts sg an e tyenv
+ (STACKS: match_stackframes s ts (fn_sig tf))
+ (FUN: transf_function f = OK tf)
+ (ANL: analyze f (pair_codes f tf) = Some an)
+ (EQ: transfer f (pair_codes f tf) pc an!!pc = OK e)
+ (WTF: wt_function f tyenv)
+ (WTRS: wt_regset tyenv rs)
+ (WTRES: tyenv res = proj_sig_res sg)
+ (STEPS: forall v ls1 m,
+ Val.lessdef_list (encode_long (sig_res sg) v) (map ls1 (map R (loc_result sg))) ->
+ agree_callee_save ls ls1 ->
+ exists ls2,
+ star LTL.step tge (Block ts tf sp bb ls1 m)
+ E0 (State ts tf sp pc ls2 m)
+ /\ satisf (rs#res <- v) ls2 e),
match_stackframes
(RTL.Stackframe res f sp pc rs :: s)
- (LTL.Stackframe (assign res) (transf_fun f live assign) sp ls pc :: ts).
+ (LTL.Stackframe tf sp ls bb :: ts)
+ sg.
Inductive match_states: RTL.state -> LTL.state -> Prop :=
| match_states_intro:
- forall s f sp pc rs m ts ls live assign env
- (STACKS: match_stackframes s ts)
- (WT: wt_function f env)
- (ANL: analyze f = Some live)
- (ASG: regalloc f live (live0 f live) env = Some assign)
- (AG: agree assign (transfer f pc live!!pc) rs ls),
+ forall s f sp pc rs m ts tf ls m' an e tyenv
+ (STACKS: match_stackframes s ts (fn_sig tf))
+ (FUN: transf_function f = OK tf)
+ (ANL: analyze f (pair_codes f tf) = Some an)
+ (EQ: transfer f (pair_codes f tf) pc an!!pc = OK e)
+ (SAT: satisf rs ls e)
+ (MEM: Mem.extends m m')
+ (WTF: wt_function f tyenv)
+ (WTRS: wt_regset tyenv rs),
match_states (RTL.State s f sp pc rs m)
- (LTL.State ts (transf_fun f live assign) sp pc ls m)
+ (LTL.State ts tf sp pc ls m')
| match_states_call:
- forall s f args m ts tf,
- match_stackframes s ts ->
- transf_fundef f = OK tf ->
+ forall s f args m ts tf ls m'
+ (STACKS: match_stackframes s ts (funsig tf))
+ (FUN: transf_fundef f = OK tf)
+ (ARGS: Val.lessdef_list args (decode_longs (sig_args (funsig tf)) (map ls (loc_arguments (funsig tf)))))
+ (AG: agree_callee_save (parent_locset ts) ls)
+ (MEM: Mem.extends m m')
+ (WTARGS: Val.has_type_list args (sig_args (funsig tf))),
match_states (RTL.Callstate s f args m)
- (LTL.Callstate ts tf args m)
+ (LTL.Callstate ts tf ls m')
| match_states_return:
- forall s v m ts,
- match_stackframes s ts ->
- match_states (RTL.Returnstate s v m)
- (LTL.Returnstate ts v m).
-
-(** The simulation proof is by case analysis over the RTL transition
- taken in the source program. *)
-
-Ltac CleanupHyps :=
- match goal with
- | H: (match_states _ _) |- _ =>
- inv H; CleanupHyps
- | H1: (PTree.get _ _ = Some _),
- H2: (agree _ (transfer _ _ _) _ _) |- _ =>
- unfold transfer in H2; rewrite H1 in H2; simpl in H2; CleanupHyps
- | _ => idtac
- end.
-
-Ltac WellTypedHyp :=
- match goal with
- | H1: (PTree.get _ _ = Some _),
- H2: (wt_function _ _) |- _ =>
- let R := fresh "WTI" in (
- generalize (wt_instrs _ _ H2 _ _ H1); intro R)
- | _ => idtac
- end.
+ forall s res m ts ls m' sg
+ (STACKS: match_stackframes s ts sg)
+ (RES: Val.lessdef_list (encode_long (sig_res sg) res) (map ls (map R (loc_result sg))))
+ (AG: agree_callee_save (parent_locset ts) ls)
+ (MEM: Mem.extends m m')
+ (WTRES: Val.has_type res (proj_sig_res sg)),
+ match_states (RTL.Returnstate s res m)
+ (LTL.Returnstate ts ls m').
+
+Lemma match_stackframes_change_sig:
+ forall s ts sg sg',
+ match_stackframes s ts sg ->
+ sg'.(sig_res) = sg.(sig_res) ->
+ match_stackframes s ts sg'.
+Proof.
+ intros. inv H.
+ constructor. congruence.
+ econstructor; eauto.
+ unfold proj_sig_res in *. rewrite H0; auto.
+ intros. unfold loc_result in H; rewrite H0 in H; eauto.
+Qed.
-Ltac TranslInstr :=
+Ltac UseShape :=
match goal with
- | H: (PTree.get _ _ = Some _) |- _ =>
- simpl in H; simpl; rewrite PTree.gmap; rewrite H; simpl; auto
+ | [ CODE: (RTL.fn_code _)!_ = Some _, EQ: transfer _ _ _ _ = OK _ |- _ ] =>
+ destruct (invert_code _ _ _ _ _ _ CODE EQ) as (eafter & bsh & bb & AFTER & BSH & TCODE & EBS & TR);
+ inv EBS; unfold transfer_aux in TR; MonadInv
end.
-Ltac MatchStates :=
+Ltac UseType :=
match goal with
- | |- match_states (RTL.State _ _ _ _ _ _) (LTL.State _ _ _ _ _ _) =>
- eapply match_states_intro; eauto; MatchStates
- | H: (PTree.get ?pc _ = Some _) |- agree _ _ _ _ =>
- eapply agree_succ with (n := pc); eauto; MatchStates
- | |- In _ (RTL.successors_instr _) =>
- unfold RTL.successors_instr; auto with coqlib
- | _ => idtac
+ | [ CODE: (RTL.fn_code _)!_ = Some _, WTF: wt_function _ _ |- _ ] =>
+ generalize (wt_instrs _ _ WTF _ _ CODE); intro WTI
end.
-Lemma transl_find_function:
- forall ros f args lv rs ls alloc,
- RTL.find_function ge ros rs = Some f ->
- agree alloc (reg_list_live args (reg_sum_live ros lv)) rs ls ->
- exists tf,
- LTL.find_function tge (sum_left_map alloc ros) ls = Some tf /\
- transf_fundef f = OK tf.
+Remark addressing_not_long:
+ forall env f addr args dst s r,
+ wt_instr env f (Iload Mint64 addr args dst s) ->
+ In r args -> r <> dst.
Proof.
- intros; destruct ros; simpl in *.
- assert (rs#r = ls (alloc r)).
- eapply agree_eval_reg. eapply agree_reg_list_live; eauto.
- rewrite <- H1. apply functions_translated. auto.
- rewrite symbols_preserved. destruct (Genv.find_symbol ge i).
- apply function_ptr_translated. auto. discriminate.
+ intros.
+ assert (forall ty, In ty (type_of_addressing addr) -> ty = Tint).
+ { intros. destruct addr; simpl in H1; intuition. }
+ inv H.
+ assert (env r = Tint).
+ { apply H1. rewrite <- H5. apply in_map; auto. }
+ simpl in H8; congruence.
Qed.
-Theorem transl_step_correct:
- forall s1 t s2, RTL.step ge s1 t s2 ->
- forall s1', match_states s1 s1' ->
- exists s2', LTL.step tge s1' t s2' /\ match_states s2 s2'.
+(** The proof of semantic preservation is a simulation argument of the
+ "plus" kind. *)
+
+Lemma step_simulation:
+ forall S1 t S2, RTL.step ge S1 t S2 ->
+ forall S1', match_states S1 S1' ->
+ exists S2', plus LTL.step tge S1' t S2' /\ match_states S2 S2'.
Proof.
- induction 1; intros; CleanupHyps; WellTypedHyp.
+ induction 1; intros S1' MS; inv MS; try UseType; try UseShape.
+
+(* nop *)
+ exploit exec_moves; eauto. intros [ls1 [X Y]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
+
+(* op move *)
+- simpl in H0. inv H0.
+ exploit (exec_moves mv); eauto. intros [ls1 [X Y]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. eapply subst_reg_satisf; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+ inv WTI. apply wt_regset_assign; auto. rewrite <- H2. apply WTRS. congruence.
- (* Inop *)
+(* op makelong *)
+- generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'.
+ simpl in H0. inv H0.
+ exploit (exec_moves mv); eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply exec_Lnop. TranslInstr. MatchStates.
-
- (* Iop *)
- generalize (PTree.gmap (transf_instr f live assign) pc (RTL.fn_code f)).
- rewrite H. simpl.
- caseEq (Regset.mem res live!!pc); intro LV;
- rewrite LV in AG.
- generalize (Regset.mem_2 LV). intro LV'.
- generalize (regalloc_correct_1 f env live _ _ _ _ ASG H).
- unfold correct_alloc_instr, is_redundant_move.
- caseEq (is_move_operation op args).
- (* Special case for moves *)
- intros arg IMO CORR.
- generalize (is_move_operation_correct _ _ IMO).
- intros [EQ1 EQ2]. subst op; subst args.
- injection H0; intro.
- destruct (Loc.eq (assign arg) (assign res)); intro CODE.
- (* sub-case: redundant move *)
- econstructor; split. eapply exec_Lnop; eauto.
- MatchStates.
- rewrite <- H1. eapply agree_redundant_move_live; eauto.
- (* sub-case: non-redundant move *)
- econstructor; split. eapply exec_Lop; eauto. simpl. eauto.
- MatchStates.
- rewrite <- H1. set (ls1 := undef_temps ls).
- replace (ls (assign arg)) with (ls1 (assign arg)).
- eapply agree_move_live; eauto.
- unfold ls1. eapply agree_undef_temps; eauto.
- unfold ls1. simpl. apply Locmap.guo. eapply regalloc_not_temporary; eauto.
- (* Not a move *)
- intros INMO CORR CODE.
- assert (eval_operation tge sp op (map ls (map assign args)) m = Some v).
- replace (map ls (map assign args)) with (rs##args).
- rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved.
- eapply agree_eval_regs; eauto.
- econstructor; split. eapply exec_Lop; eauto. MatchStates.
- apply agree_assign_live with f env live; auto.
- eapply agree_undef_temps; eauto.
- eapply agree_reg_list_live; eauto.
- (* Result is not live, instruction turned into a nop *)
- intro CODE. econstructor; split. eapply exec_Lnop; eauto.
- MatchStates. apply agree_assign_dead; auto.
- red; intro. exploit Regset.mem_1; eauto. congruence.
-
- (* Iload *)
- caseEq (Regset.mem dst live!!pc); intro LV;
- rewrite LV in AG.
- (* dst is live *)
- exploit Regset.mem_2; eauto. intro LV'.
- assert (eval_addressing tge sp addr (map ls (map assign args)) = Some a).
- replace (map ls (map assign args)) with (rs##args).
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
- eapply agree_eval_regs; eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto.
+ eapply subst_reg_kind_satisf_makelong. eauto. eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+
+(* op lowlong *)
+- generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'.
+ simpl in H0. inv H0.
+ exploit (exec_moves mv); eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply exec_Lload; eauto. TranslInstr. rewrite LV; auto.
- generalize (regalloc_correct_1 f env live _ _ _ _ ASG H).
- unfold correct_alloc_instr. intro CORR.
- MatchStates.
- eapply agree_assign_live; eauto.
- eapply agree_undef_temps; eauto.
- eapply agree_reg_list_live; eauto.
- (* dst is dead *)
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto.
+ eapply subst_reg_kind_satisf_lowlong. eauto. eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+
+(* op highlong *)
+- generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'.
+ simpl in H0. inv H0.
+ exploit (exec_moves mv); eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply exec_Lnop. TranslInstr. rewrite LV; auto.
- MatchStates. apply agree_assign_dead; auto.
- red; intro; exploit Regset.mem_1; eauto. congruence.
-
- (* Istore *)
- assert (eval_addressing tge sp addr (map ls (map assign args)) = Some a).
- replace (map ls (map assign args)) with (rs##args).
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
- eapply agree_eval_regs; eauto.
- assert (ESRC: rs#src = ls (assign src)).
- eapply agree_eval_reg. eapply agree_reg_list_live. eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto.
+ eapply subst_reg_kind_satisf_highlong. eauto. eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+
+(* op regular *)
+- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ exploit transfer_use_def_satisf; eauto. intros [X Y].
+ exploit eval_operation_lessdef; eauto. intros [v' [F G]].
+ exploit (exec_moves mv2); eauto. intros [ls2 [A2 B2]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor. instantiate (1 := v'). rewrite <- F.
+ apply eval_operation_preserved. exact symbols_preserved.
+ eauto. eapply star_right. eexact A2. constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
+ eapply wt_exec_Iop; eauto.
+
+(* op dead *)
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply exec_Lstore; eauto. TranslInstr.
- rewrite <- ESRC. eauto.
- MatchStates.
- eapply agree_undef_temps; eauto.
- eapply agree_reg_live. eapply agree_reg_list_live. eauto.
-
- (* Icall *)
- exploit transl_find_function; eauto. intros [tf [TFIND TF]].
- generalize (regalloc_correct_1 f env live _ _ _ _ ASG H). unfold correct_alloc_instr. intros [CORR1 [CORR2 CORR3]].
- assert (rs##args = map ls (map assign args)).
- eapply agree_eval_regs; eauto.
- econstructor; split.
- eapply exec_Lcall; eauto. TranslInstr.
- rewrite (sig_function_translated _ _ TF). eauto.
- rewrite H1.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors. eauto. eauto. simpl; eauto. eauto.
+ eapply reg_unconstrained_satisf; eauto.
+ intros [enext [U V]].
econstructor; eauto.
+ eapply wt_exec_Iop; eauto.
+
+(* load regular *)
+- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ exploit transfer_use_def_satisf; eauto. intros [X Y].
+ exploit eval_addressing_lessdef; eauto. intros [a' [F G]].
+ exploit Mem.loadv_extends; eauto. intros [v' [P Q]].
+ exploit (exec_moves mv2); eauto. intros [ls2 [A2 B2]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor. instantiate (1 := a'). rewrite <- F.
+ apply eval_addressing_preserved. exact symbols_preserved. eauto. eauto.
+ eapply star_right. eexact A2. constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
+ eapply wt_exec_Iload; eauto.
+
+(* load pair *)
+- exploit Mem.loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V12).
+ set (v2' := if big_endian then v2 else v1) in *.
+ set (v1' := if big_endian then v1 else v2) in *.
+ exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ assert (LD1: Val.lessdef_list rs##args (reglist ls1 args1')).
+ { eapply add_equations_lessdef; eauto. }
+ exploit eval_addressing_lessdef. eexact LD1. eauto. intros [a1' [F1 G1]].
+ exploit Mem.loadv_extends. eauto. eexact LOAD1. eexact G1. intros (v1'' & LOAD1' & LD2).
+ set (ls2 := Locmap.set (R dst1') v1'' (undef_regs (destroyed_by_load Mint32 addr) ls1)).
+ assert (SAT2: satisf (rs#dst <- v) ls2 e2).
+ { eapply loc_unconstrained_satisf. eapply can_undef_satisf; eauto.
+ eapply reg_unconstrained_satisf. eauto.
+ eapply add_equations_satisf; eauto. assumption.
+ rewrite Regmap.gss. apply Val.lessdef_trans with v1'; auto.
+ subst v. unfold v1', kind_first_word.
+ destruct big_endian; apply val_hiword_longofwords || apply val_loword_longofwords.
+ }
+ exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
+ assert (LD3: Val.lessdef_list rs##args (reglist ls3 args2')).
+ { replace (rs##args) with ((rs#dst<-v)##args).
+ eapply add_equations_lessdef; eauto.
+ apply list_map_exten; intros. rewrite Regmap.gso; auto.
+ eapply addressing_not_long; eauto.
+ }
+ exploit eval_addressing_lessdef. eexact LD3.
+ eapply eval_offset_addressing; eauto. intros [a2' [F2 G2]].
+ exploit Mem.loadv_extends. eauto. eexact LOAD2. eexact G2. intros (v2'' & LOAD2' & LD4).
+ set (ls4 := Locmap.set (R dst2') v2'' (undef_regs (destroyed_by_load Mint32 addr2) ls3)).
+ assert (SAT4: satisf (rs#dst <- v) ls4 e0).
+ { eapply loc_unconstrained_satisf. eapply can_undef_satisf; eauto.
+ eapply add_equations_satisf; eauto. assumption.
+ apply Val.lessdef_trans with v2'; auto.
+ rewrite Regmap.gss. subst v. unfold v2', kind_second_word.
+ destruct big_endian; apply val_hiword_longofwords || apply val_loword_longofwords.
+ }
+ exploit (exec_moves mv3); eauto. intros [ls5 [A5 B5]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor.
+ instantiate (1 := a1'). rewrite <- F1. apply eval_addressing_preserved. exact symbols_preserved.
+ eexact LOAD1'. instantiate (1 := ls2); auto.
+ eapply star_trans. eexact A3.
+ eapply star_left. econstructor.
+ instantiate (1 := a2'). rewrite <- F2. apply eval_addressing_preserved. exact symbols_preserved.
+ eexact LOAD2'. instantiate (1 := ls4); auto.
+ eapply star_right. eexact A5.
+ constructor.
+ eauto. eauto. eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
econstructor; eauto.
- intros. eapply agree_succ with (n := pc); eauto.
- simpl. auto.
- eapply agree_postcall; eauto.
-
- (* Itailcall *)
- exploit transl_find_function; eauto. intros [tf [TFIND TF]].
- assert (rs##args = map ls (map assign args)).
- eapply agree_eval_regs; eauto.
- econstructor; split.
- eapply exec_Ltailcall; eauto. TranslInstr.
- rewrite (sig_function_translated _ _ TF). eauto.
- rewrite H1. econstructor; eauto.
-
- (* Ibuiltin *)
+ eapply wt_exec_Iload; eauto.
+
+(* load first word of a pair *)
+- exploit Mem.loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V12).
+ set (v2' := if big_endian then v2 else v1) in *.
+ set (v1' := if big_endian then v1 else v2) in *.
+ exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ assert (LD1: Val.lessdef_list rs##args (reglist ls1 args')).
+ { eapply add_equations_lessdef; eauto. }
+ exploit eval_addressing_lessdef. eexact LD1. eauto. intros [a1' [F1 G1]].
+ exploit Mem.loadv_extends. eauto. eexact LOAD1. eexact G1. intros (v1'' & LOAD1' & LD2).
+ set (ls2 := Locmap.set (R dst') v1'' (undef_regs (destroyed_by_load Mint32 addr) ls1)).
+ assert (SAT2: satisf (rs#dst <- v) ls2 e0).
+ { eapply parallel_assignment_satisf; eauto.
+ apply Val.lessdef_trans with v1'; auto.
+ subst v. unfold v1', kind_first_word.
+ destruct big_endian; apply val_hiword_longofwords || apply val_loword_longofwords.
+ eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto.
+ }
+ exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
econstructor; split.
- eapply exec_Lbuiltin; eauto. TranslInstr.
- replace (map ls (@map reg loc assign args)) with (rs##args).
- eapply external_call_symbols_preserved; eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor.
+ instantiate (1 := a1'). rewrite <- F1. apply eval_addressing_preserved. exact symbols_preserved.
+ eexact LOAD1'. instantiate (1 := ls2); auto.
+ eapply star_right. eexact A3.
+ constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
+ econstructor; eauto.
+ eapply wt_exec_Iload; eauto.
+
+(* load second word of a pair *)
+- exploit Mem.loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V12).
+ set (v2' := if big_endian then v2 else v1) in *.
+ set (v1' := if big_endian then v1 else v2) in *.
+ exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ assert (LD1: Val.lessdef_list rs##args (reglist ls1 args')).
+ { eapply add_equations_lessdef; eauto. }
+ exploit eval_addressing_lessdef. eexact LD1.
+ eapply eval_offset_addressing; eauto. intros [a1' [F1 G1]].
+ exploit Mem.loadv_extends. eauto. eexact LOAD2. eexact G1. intros (v2'' & LOAD2' & LD2).
+ set (ls2 := Locmap.set (R dst') v2'' (undef_regs (destroyed_by_load Mint32 addr2) ls1)).
+ assert (SAT2: satisf (rs#dst <- v) ls2 e0).
+ { eapply parallel_assignment_satisf; eauto.
+ apply Val.lessdef_trans with v2'; auto.
+ subst v. unfold v2', kind_second_word.
+ destruct big_endian; apply val_hiword_longofwords || apply val_loword_longofwords.
+ eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto.
+ }
+ exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor.
+ instantiate (1 := a1'). rewrite <- F1. apply eval_addressing_preserved. exact symbols_preserved.
+ eexact LOAD2'. instantiate (1 := ls2); auto.
+ eapply star_right. eexact A3.
+ constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
+ econstructor; eauto.
+ eapply wt_exec_Iload; eauto.
+
+(* load dead *)
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors. eauto. eauto. simpl; eauto. eauto.
+ eapply reg_unconstrained_satisf; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+ eapply wt_exec_Iload; eauto.
+
+(* store *)
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
+ exploit add_equations_lessdef; eauto. intros LD. simpl in LD. inv LD.
+ exploit eval_addressing_lessdef; eauto. intros [a' [F G]].
+ exploit Mem.storev_extends; eauto. intros [m'' [P Q]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact X.
+ eapply star_two. econstructor. instantiate (1 := a'). rewrite <- F.
+ apply eval_addressing_preserved. exact symbols_preserved. eauto. eauto.
+ constructor. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto.
+ eapply can_undef_satisf; eauto. eapply add_equations_satisf; eauto. intros [enext [U V]].
+ econstructor; eauto.
+
+(* store 2 *)
+- exploit Mem.storev_int64_split; eauto.
+ replace (if big_endian then Val.hiword rs#src else Val.loword rs#src)
+ with (sel_val kind_first_word rs#src)
+ by (unfold kind_first_word; destruct big_endian; reflexivity).
+ replace (if big_endian then Val.loword rs#src else Val.hiword rs#src)
+ with (sel_val kind_second_word rs#src)
+ by (unfold kind_second_word; destruct big_endian; reflexivity).
+ intros [m1 [STORE1 STORE2]].
+ exploit (exec_moves mv1); eauto. intros [ls1 [X Y]].
+ exploit add_equations_lessdef. eexact Heqo1. eexact Y. intros LD1.
+ exploit add_equation_lessdef. eapply add_equations_satisf. eexact Heqo1. eexact Y.
+ simpl. intros LD2.
+ set (ls2 := undef_regs (destroyed_by_store Mint32 addr) ls1).
+ assert (SAT2: satisf rs ls2 e1).
+ eapply can_undef_satisf. eauto.
+ eapply add_equation_satisf. eapply add_equations_satisf; eauto.
+ exploit eval_addressing_lessdef. eexact LD1. eauto. intros [a1' [F1 G1]].
+ assert (F1': eval_addressing tge sp addr (reglist ls1 args1') = Some a1').
+ rewrite <- F1. apply eval_addressing_preserved. exact symbols_preserved.
+ exploit Mem.storev_extends. eauto. eexact STORE1. eexact G1. eauto.
+ intros [m1' [STORE1' EXT1]].
+ exploit (exec_moves mv2); eauto. intros [ls3 [U V]].
+ exploit add_equations_lessdef. eexact Heqo. eexact V. intros LD3.
+ exploit add_equation_lessdef. eapply add_equations_satisf. eexact Heqo. eexact V.
+ simpl. intros LD4.
+ exploit eval_addressing_lessdef. eexact LD3. eauto. intros [a2' [F2 G2]].
+ assert (F2': eval_addressing tge sp addr (reglist ls3 args2') = Some a2').
+ rewrite <- F2. apply eval_addressing_preserved. exact symbols_preserved.
+ exploit eval_offset_addressing. eauto. eexact F2'. intros F2''.
+ exploit Mem.storev_extends. eexact EXT1. eexact STORE2.
+ apply Val.add_lessdef. eexact G2. eauto. eauto.
+ intros [m2' [STORE2' EXT2]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact X.
+ eapply star_left.
+ econstructor. eexact F1'. eexact STORE1'. instantiate (1 := ls2). auto.
+ eapply star_trans. eexact U.
+ eapply star_two.
+ econstructor. eexact F2''. eexact STORE2'. eauto.
+ constructor. eauto. eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto.
+ eapply can_undef_satisf. eauto.
+ eapply add_equation_satisf. eapply add_equations_satisf; eauto.
+ intros [enext [P Q]].
+ econstructor; eauto.
+
+(* call *)
+- set (sg := RTL.funsig fd) in *.
+ set (args' := loc_arguments sg) in *.
+ set (res' := map R (loc_result sg)) in *.
+ exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ exploit find_function_translated. eauto. eauto. eapply add_equations_args_satisf; eauto.
+ intros [tfd [E F]].
+ assert (SIG: funsig tfd = sg). eapply sig_function_translated; eauto.
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact A1. econstructor; eauto.
+ eauto. traceEq.
+ exploit analyze_successors; eauto. simpl. left; eauto. intros [enext [U V]].
+ econstructor; eauto.
+ econstructor; eauto.
+ inv WTI. rewrite SIG. auto.
+ intros. exploit (exec_moves mv2); eauto.
+ eapply function_return_satisf with (v := v) (ls_before := ls1) (ls_after := ls0); eauto.
+ eapply add_equation_ros_satisf; eauto.
+ eapply add_equations_args_satisf; eauto.
+ congruence. intros [ls2 [A2 B2]].
+ exists ls2; split.
+ eapply star_right. eexact A2. constructor. traceEq.
+ apply satisf_incr with eafter; auto.
+ rewrite SIG. eapply add_equations_args_lessdef; eauto.
+ inv WTI. rewrite <- H7. apply wt_regset_list; auto.
+ simpl. red; auto.
+ rewrite SIG. inv WTI. rewrite <- H7. apply wt_regset_list; auto.
+
+(* tailcall *)
+- set (sg := RTL.funsig fd) in *.
+ set (args' := loc_arguments sg) in *.
+ exploit Mem.free_parallel_extends; eauto. intros [m'' [P Q]].
+ exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
+ exploit find_function_translated. eauto. eauto. eapply add_equations_args_satisf; eauto.
+ intros [tfd [E F]].
+ assert (SIG: funsig tfd = sg). eapply sig_function_translated; eauto.
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact A1. econstructor; eauto.
+ rewrite <- E. apply find_function_tailcall; auto.
+ replace (fn_stacksize tf) with (RTL.fn_stacksize f); eauto.
+ destruct (transf_function_inv _ _ FUN); auto.
+ eauto. traceEq.
+ econstructor; eauto.
+ eapply match_stackframes_change_sig; eauto. rewrite SIG. rewrite e0. decEq.
+ destruct (transf_function_inv _ _ FUN); auto.
+ rewrite SIG. rewrite return_regs_arg_values; auto. eapply add_equations_args_lessdef; eauto.
+ inv WTI. rewrite <- H8. apply wt_regset_list; auto.
+ apply return_regs_agree_callee_save.
+ rewrite SIG. inv WTI. rewrite <- H8. apply wt_regset_list; auto.
+
+(* builtin *)
+- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ exploit external_call_mem_extends; eauto.
+ eapply add_equations_args_lessdef; eauto.
+ inv WTI. rewrite <- H4. apply wt_regset_list; auto.
+ intros [v' [m'' [F [G [J K]]]]].
+ assert (E: map ls1 (map R args') = reglist ls1 args').
+ { unfold reglist. rewrite list_map_compose. auto. }
+ rewrite E in F. clear E.
+ set (vl' := encode_long (sig_res (ef_sig ef)) v').
+ set (ls2 := Locmap.setlist (map R res') vl' (undef_regs (destroyed_by_builtin ef) ls1)).
+ assert (satisf (rs#res <- v) ls2 e0).
+ { eapply parallel_assignment_satisf_2; eauto.
+ eapply can_undef_satisf; eauto.
+ eapply add_equations_args_satisf; eauto. }
+ exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor.
+ econstructor. unfold reglist. eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact varinfo_preserved.
- eapply agree_eval_regs; eauto.
- generalize (regalloc_correct_1 f env live _ _ _ _ ASG H).
- unfold correct_alloc_instr. intro CORR.
- MatchStates.
- eapply agree_assign_live; eauto.
- eapply agree_reg_list_live; eauto.
-
- (* Icond *)
- assert (COND: eval_condition cond (map ls (map assign args)) m = Some b).
- replace (map ls (map assign args)) with (rs##args). auto.
- eapply agree_eval_regs; eauto.
+ instantiate (1 := vl'); auto.
+ instantiate (1 := ls2); auto.
+ eapply star_right. eexact A3.
+ econstructor.
+ reflexivity. reflexivity. reflexivity. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+ inv WTI. apply wt_regset_assign; auto. rewrite H9.
+ eapply external_call_well_typed; eauto.
+
+(* annot *)
+- exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
+ exploit external_call_mem_extends; eauto. eapply add_equations_args_lessdef; eauto.
+ inv WTI. simpl in H4. rewrite <- H4. apply wt_regset_list; auto.
+ intros [v' [m'' [F [G [J K]]]]].
+ assert (v = Vundef). red in H0; inv H0. auto.
econstructor; split.
- eapply exec_Lcond; eauto. TranslInstr.
- MatchStates. destruct b; simpl; auto.
- eapply agree_undef_temps; eauto.
- eapply agree_reg_list_live. eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_two. econstructor.
+ eapply external_call_symbols_preserved' with (ge1 := ge).
+ econstructor; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ eauto. constructor. eauto. eauto. traceEq.
+ exploit satisf_successors. eauto. eauto. simpl; eauto. eauto.
+ eapply satisf_undef_reg with (r := res).
+ eapply add_equations_args_satisf; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+ change (destroyed_by_builtin (EF_annot txt typ)) with (@nil mreg).
+ simpl. subst v. assumption.
+ apply wt_regset_assign; auto. subst v. constructor.
- (* Ijumptable *)
- assert (rs#arg = ls (assign arg)). apply AG. apply Regset.add_1. auto.
+(* cond *)
+- exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
econstructor; split.
- eapply exec_Ljumptable; eauto. TranslInstr. congruence.
- MatchStates. eapply list_nth_z_in; eauto.
- eapply agree_undef_temps; eauto.
- eapply agree_reg_live; eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact A1.
+ econstructor. eapply eval_condition_lessdef; eauto. eapply add_equations_lessdef; eauto.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto.
+ instantiate (1 := if b then ifso else ifnot). simpl. destruct b; auto.
+ eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
- (* Ireturn *)
+(* jumptable *)
+- exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
+ assert (Val.lessdef (Vint n) (ls1 (R arg'))).
+ rewrite <- H0. eapply add_equation_lessdef with (q := Eq Full arg (R arg')); eauto.
+ inv H2.
econstructor; split.
- eapply exec_Lreturn; eauto. TranslInstr; eauto.
- replace (regmap_optget or Vundef rs)
- with (locmap_optget (option_map assign or) Vundef ls).
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact A1.
+ econstructor. eauto. eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto.
+ instantiate (1 := pc'). simpl. eapply list_nth_z_in; eauto.
+ eapply can_undef_satisf. eauto. eapply add_equation_satisf; eauto.
+ intros [enext [U V]].
econstructor; eauto.
- inv WTI. destruct or; simpl in *.
- symmetry; eapply agree_eval_reg; eauto.
- auto.
- (* internal function *)
- generalize H7. simpl. unfold transf_function.
- caseEq (type_function f); simpl; try congruence. intros env TYP.
- assert (WTF: wt_function f env). apply type_function_correct; auto.
- caseEq (analyze f); simpl; try congruence. intros live ANL.
- caseEq (regalloc f live (live0 f live) env); simpl; try congruence.
- intros alloc ALLOC EQ. inv EQ. simpl in *.
+(* return *)
+- destruct (transf_function_inv _ _ FUN).
+ exploit Mem.free_parallel_extends; eauto. rewrite H10. intros [m'' [P Q]].
+ destruct or as [r|]; MonadInv.
+ (* with an argument *)
++ exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
econstructor; split.
- eapply exec_function_internal; simpl; eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact A1.
+ econstructor. eauto. eauto. traceEq.
+ simpl. econstructor; eauto. rewrite <- H11.
+ replace (map (return_regs (parent_locset ts) ls1) (map R (loc_result (RTL.fn_sig f))))
+ with (map ls1 (map R (loc_result (RTL.fn_sig f)))).
+ eapply add_equations_res_lessdef; eauto.
+ rewrite !list_map_compose. apply list_map_exten; intros.
+ unfold return_regs. apply pred_dec_true. eapply loc_result_caller_save; eauto.
+ apply return_regs_agree_callee_save.
+ inv WTI. simpl in H13. unfold proj_sig_res. rewrite <- H11; rewrite <- H13. apply WTRS.
+ (* without an argument *)
++ exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact A1.
+ econstructor. eauto. eauto. traceEq.
simpl. econstructor; eauto.
- change (transfer f (RTL.fn_entrypoint f) live !! (RTL.fn_entrypoint f))
- with (live0 f live).
- eapply agree_parameters; eauto.
+ unfold encode_long, loc_result. destruct (sig_res (fn_sig tf)) as [[]|]; simpl; auto.
+ apply return_regs_agree_callee_save.
+ constructor.
+
+(* internal function *)
+- monadInv FUN. simpl in *.
+ destruct (transf_function_inv _ _ EQ).
+ exploit Mem.alloc_extends; eauto. apply Zle_refl. rewrite H8; apply Zle_refl.
+ intros [m'' [U V]].
+ exploit (exec_moves mv). eauto. eauto.
+ eapply can_undef_satisf; eauto. eapply compat_entry_satisf; eauto.
+ rewrite call_regs_param_values. rewrite H9. eexact ARGS.
+ intros [ls1 [A B]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_left. econstructor; eauto.
+ eapply star_right. eexact A.
+ econstructor; eauto.
+ eauto. eauto. traceEq.
+ econstructor; eauto.
+ apply wt_init_regs. inv H0. rewrite wt_params. congruence.
- (* external function *)
- injection H7; intro EQ; inv EQ.
+(* external function *)
+- exploit external_call_mem_extends; eauto. intros [v' [m'' [F [G [J K]]]]].
+ simpl in FUN; inv FUN.
econstructor; split.
- eapply exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
+ apply plus_one. econstructor; eauto.
+ eapply external_call_symbols_preserved' with (ge1 := ge).
+ econstructor; eauto.
exact symbols_preserved. exact varinfo_preserved.
- eapply match_states_return; eauto.
-
- (* return *)
- inv H4.
+ econstructor; eauto. simpl.
+ replace (map
+ (Locmap.setlist (map R (loc_result (ef_sig ef)))
+ (encode_long (sig_res (ef_sig ef)) v') ls)
+ (map R (loc_result (ef_sig ef))))
+ with (encode_long (sig_res (ef_sig ef)) v').
+ apply encode_long_lessdef; auto.
+ unfold encode_long, loc_result.
+ destruct (sig_res (ef_sig ef)) as [[]|]; simpl; symmetry; f_equal; auto.
+ red; intros. rewrite Locmap.gsetlisto. apply AG; auto.
+ apply Loc.notin_iff. intros.
+ exploit list_in_map_inv; eauto. intros [r [A B]]; subst l'.
+ destruct l; simpl; auto. red; intros; subst r0; elim H0.
+ eapply loc_result_caller_save; eauto.
+ simpl. eapply external_call_well_typed; eauto.
+
+(* return *)
+- inv STACKS.
+ exploit STEPS; eauto. intros [ls2 [A B]].
econstructor; split.
- eapply exec_return; eauto.
+ eapply plus_left. constructor. eexact A. traceEq.
econstructor; eauto.
+ apply wt_regset_assign; auto. congruence.
Qed.
-(** The semantic equivalence between the original and transformed programs
- follows easily. *)
-
-Lemma transf_initial_states:
+Lemma initial_states_simulation:
forall st1, RTL.initial_state prog st1 ->
exists st2, LTL.initial_state tprog st2 /\ match_states st1 st2.
Proof.
- intros. inversion H.
+ intros. inv H.
exploit function_ptr_translated; eauto. intros [tf [FIND TR]].
- exists (LTL.Callstate nil tf nil m0); split.
+ exploit sig_function_translated; eauto. intros SIG.
+ exists (LTL.Callstate nil tf (Locmap.init Vundef) m0); split.
econstructor; eauto.
eapply Genv.init_mem_transf_partial; eauto.
rewrite symbols_preserved.
rewrite (transform_partial_program_main _ _ TRANSF). auto.
- rewrite <- H3. apply sig_function_translated; auto.
- constructor; auto. constructor.
+ congruence.
+ constructor; auto.
+ constructor. rewrite SIG; rewrite H3; auto.
+ rewrite SIG; rewrite H3; simpl; auto.
+ red; auto.
+ apply Mem.extends_refl.
+ rewrite SIG; rewrite H3; simpl; auto.
Qed.
-Lemma transf_final_states:
+Lemma final_states_simulation:
forall st1 st2 r,
match_states st1 st2 -> RTL.final_state st1 r -> LTL.final_state st2 r.
Proof.
- intros. inv H0. inv H. inv H4. econstructor.
+ intros. inv H0. inv H. inv STACKS.
+ econstructor. simpl; reflexivity.
+ unfold loc_result in RES; rewrite H in RES. simpl in RES. inv RES. inv H3; auto.
Qed.
Theorem transf_program_correct:
forward_simulation (RTL.semantics prog) (LTL.semantics tprog).
Proof.
- eapply forward_simulation_step.
+ eapply forward_simulation_plus.
eexact symbols_preserved.
- eexact transf_initial_states.
- eexact transf_final_states.
- exact transl_step_correct.
+ eexact initial_states_simulation.
+ eexact final_states_simulation.
+ exact step_simulation.
Qed.
End PRESERVATION.
diff --git a/backend/Alloctyping.v b/backend/Alloctyping.v
deleted file mode 100644
index 59bf621..0000000
--- a/backend/Alloctyping.v
+++ /dev/null
@@ -1,205 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** 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 Liveness.
-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 ->
- f.(RTL.fn_code)!pc = Some 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 *)
- exploit regalloc_correct_1; eauto. unfold correct_alloc_instr.
- intros [A1 [A2 A3]].
- WT.
- destruct ros; simpl; auto.
- split. autorewrite with allocty; auto.
- split. auto with allocty. auto.
- (* tailcall *)
- exploit regalloc_correct_1; eauto. unfold correct_alloc_instr.
- intro A1.
- WT.
- destruct ros; simpl; auto.
- split. autorewrite with allocty; auto.
- split. auto with allocty. auto.
- rewrite transf_unroll; auto.
- (* builtin *)
- WT.
- (* cond *)
- WT.
- (* jumptable *)
- 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.
diff --git a/backend/Asmgenproof0.v b/backend/Asmgenproof0.v
index 72de80a..f74fba8 100644
--- a/backend/Asmgenproof0.v
+++ b/backend/Asmgenproof0.v
@@ -80,29 +80,6 @@ Qed.
Hint Resolve preg_of_not_SP preg_of_not_PC: asmgen.
-Lemma nontemp_diff:
- forall r r',
- nontemp_preg r = true -> nontemp_preg r' = false -> r <> r'.
-Proof.
- congruence.
-Qed.
-Hint Resolve nontemp_diff: asmgen.
-
-Lemma temporaries_temp_preg:
- forall r, In r temporary_regs -> nontemp_preg (preg_of r) = false.
-Proof.
- assert (List.forallb (fun r => negb (nontemp_preg (preg_of r))) temporary_regs = true) by reflexivity.
- rewrite List.forallb_forall in H. intros. generalize (H r H0).
- destruct (nontemp_preg (preg_of r)); simpl; congruence.
-Qed.
-
-Lemma nontemp_data_preg:
- forall r, nontemp_preg r = true -> data_preg r = true.
-Proof.
- destruct r; try (destruct i); try (destruct f); simpl; congruence.
-Qed.
-Hint Resolve nontemp_data_preg: asmgen.
-
Lemma nextinstr_pc:
forall rs, (nextinstr rs)#PC = Val.add rs#PC Vone.
Proof.
@@ -121,12 +98,6 @@ Proof.
intros. apply nextinstr_inv. red; intro; subst; discriminate.
Qed.
-Lemma nextinstr_inv2:
- forall r rs, nontemp_preg r = true -> (nextinstr rs)#r = rs#r.
-Proof.
- intros. apply nextinstr_inv1; auto with asmgen.
-Qed.
-
Lemma nextinstr_set_preg:
forall rs m v,
(nextinstr (rs#(preg_of m) <- v))#PC = Val.add rs#PC Vone.
@@ -135,6 +106,58 @@ Proof.
rewrite Pregmap.gso. auto. apply sym_not_eq. apply preg_of_not_PC.
Qed.
+Lemma undef_regs_other:
+ forall r rl rs,
+ (forall r', In r' rl -> r <> r') ->
+ undef_regs rl rs r = rs r.
+Proof.
+ induction rl; simpl; intros. auto.
+ rewrite IHrl by auto. rewrite Pregmap.gso; auto.
+Qed.
+
+Fixpoint preg_notin (r: preg) (rl: list mreg) : Prop :=
+ match rl with
+ | nil => True
+ | r1 :: nil => r <> preg_of r1
+ | r1 :: rl => r <> preg_of r1 /\ preg_notin r rl
+ end.
+
+Remark preg_notin_charact:
+ forall r rl,
+ preg_notin r rl <-> (forall mr, In mr rl -> r <> preg_of mr).
+Proof.
+ induction rl; simpl; intros.
+ tauto.
+ destruct rl.
+ simpl. split. intros. intuition congruence. auto.
+ rewrite IHrl. split.
+ intros [A B]. intros. destruct H. congruence. auto.
+ auto.
+Qed.
+
+Lemma undef_regs_other_2:
+ forall r rl rs,
+ preg_notin r rl ->
+ undef_regs (map preg_of rl) rs r = rs r.
+Proof.
+ intros. apply undef_regs_other. intros.
+ exploit list_in_map_inv; eauto. intros [mr [A B]]. subst.
+ rewrite preg_notin_charact in H. auto.
+Qed.
+
+Lemma set_pregs_other_2:
+ forall r rl vl rs,
+ preg_notin r rl ->
+ set_regs (map preg_of rl) vl rs r = rs r.
+Proof.
+ induction rl; simpl; intros.
+ auto.
+ destruct vl; auto.
+ assert (r <> preg_of a) by (destruct rl; tauto).
+ assert (preg_notin r rl) by (destruct rl; simpl; tauto).
+ rewrite IHrl by auto. apply Pregmap.gso; auto.
+Qed.
+
(** * Agreement between Mach registers and processor registers *)
Record agree (ms: Mach.regset) (sp: val) (rs: Asm.regset) : Prop := mkagree {
@@ -224,7 +247,21 @@ Proof.
intros. unfold nextinstr. apply agree_set_other. auto. auto.
Qed.
-Lemma agree_undef_regs:
+Lemma agree_set_mregs:
+ forall sp rl vl vl' ms rs,
+ agree ms sp rs ->
+ Val.lessdef_list vl vl' ->
+ agree (Mach.set_regs rl vl ms) sp (set_regs (map preg_of rl) vl' rs).
+Proof.
+ induction rl; simpl; intros.
+ auto.
+ inv H0. auto. apply IHrl; auto.
+ eapply agree_set_mreg. eexact H.
+ rewrite Pregmap.gss. auto.
+ intros. apply Pregmap.gso; auto.
+Qed.
+
+Lemma agree_undef_nondata_regs:
forall ms sp rl rs,
agree ms sp rs ->
(forall r, In r rl -> data_preg r = false) ->
@@ -237,31 +274,33 @@ Proof.
intros. apply H0; auto.
Qed.
-Lemma agree_exten_temps:
- forall ms sp rs rs',
+Lemma agree_undef_regs:
+ forall ms sp rl rs rs',
agree ms sp rs ->
- (forall r, nontemp_preg r = true -> rs'#r = rs#r) ->
- agree (undef_temps ms) sp rs'.
+ (forall r', data_preg r' = true -> preg_notin r' rl -> rs'#r' = rs#r') ->
+ agree (Mach.undef_regs rl ms) sp rs'.
Proof.
intros. destruct H. split; auto.
- rewrite H0; auto. auto.
- intros. unfold undef_temps.
- destruct (In_dec mreg_eq r temporary_regs).
- rewrite Mach.undef_regs_same; auto.
- rewrite Mach.undef_regs_other; auto. rewrite H0; auto.
- simpl in n. destruct r; auto; intuition.
+ rewrite <- agree_sp0. apply H0; auto.
+ rewrite preg_notin_charact. intros. apply not_eq_sym. apply preg_of_not_SP.
+ intros. destruct (In_dec mreg_eq r rl).
+ rewrite Mach.undef_regs_same; auto.
+ rewrite Mach.undef_regs_other; auto. rewrite H0; auto.
+ apply preg_of_data.
+ rewrite preg_notin_charact. intros; red; intros. elim n.
+ exploit preg_of_injective; eauto. congruence.
Qed.
Lemma agree_set_undef_mreg:
- forall ms sp rs r v rs',
+ forall ms sp rs r v rl rs',
agree ms sp rs ->
Val.lessdef v (rs'#(preg_of r)) ->
- (forall r', nontemp_preg r' = true -> r' <> preg_of r -> rs'#r' = rs#r') ->
- agree (Regmap.set r v (undef_temps ms)) sp rs'.
+ (forall r', data_preg r' = true -> r' <> preg_of r -> preg_notin r' rl -> rs'#r' = rs#r') ->
+ agree (Regmap.set r v (Mach.undef_regs rl ms)) sp rs'.
Proof.
intros. apply agree_set_mreg with (rs'#(preg_of r) <- (rs#(preg_of r))); auto.
- eapply agree_exten_temps; eauto.
- intros. unfold Pregmap.set. destruct (PregEq.eq r0 (preg_of r)).
+ apply agree_undef_regs with rs; auto.
+ intros. unfold Pregmap.set. destruct (PregEq.eq r' (preg_of r)).
congruence. auto.
intros. rewrite Pregmap.gso; auto.
Qed.
@@ -291,9 +330,7 @@ Proof.
exploit Mem.loadv_extends; eauto. intros [v' [A B]].
rewrite (sp_val _ _ _ H) in A.
exists v'; split; auto.
- destruct ty; econstructor.
- reflexivity. assumption.
- reflexivity. assumption.
+ econstructor. eauto. assumption.
Qed.
Lemma extcall_args_match:
@@ -667,6 +704,7 @@ Ltac TailNoLabel :=
match goal with
| [ |- tail_nolabel _ (_ :: _) ] => apply tail_nolabel_cons; [auto; exact I | TailNoLabel]
| [ H: Error _ = OK _ |- _ ] => discriminate
+ | [ H: assertion_failed = OK _ |- _ ] => discriminate
| [ H: OK _ = OK _ |- _ ] => inv H; TailNoLabel
| [ H: bind _ _ = OK _ |- _ ] => monadInv H; TailNoLabel
| [ H: (if ?x then _ else _) = OK _ |- _ ] => destruct x; TailNoLabel
diff --git a/backend/Bounds.v b/backend/Bounds.v
index ef78b2e..bcd2848 100644
--- a/backend/Bounds.v
+++ b/backend/Bounds.v
@@ -17,7 +17,6 @@ Require Import AST.
Require Import Op.
Require Import Locations.
Require Import Linear.
-Require Import Lineartyping.
Require Import Conventions.
(** * Resource bounds for a function *)
@@ -30,14 +29,12 @@ Require Import Conventions.
the activation record. *)
Record bounds : Type := mkbounds {
- bound_int_local: Z;
- bound_float_local: Z;
+ bound_local: Z;
bound_int_callee_save: Z;
bound_float_callee_save: Z;
bound_outgoing: Z;
bound_stack_data: Z;
- bound_int_local_pos: bound_int_local >= 0;
- bound_float_local_pos: bound_float_local >= 0;
+ bound_local_pos: bound_local >= 0;
bound_int_callee_save_pos: bound_int_callee_save >= 0;
bound_float_callee_save_pos: bound_float_callee_save >= 0;
bound_outgoing_pos: bound_outgoing >= 0;
@@ -47,41 +44,39 @@ Record bounds : Type := mkbounds {
(** The following predicates define the correctness of a set of bounds
for the code of a function. *)
-Section BELOW.
+Section WITHIN_BOUNDS.
-Variable funct: function.
Variable b: bounds.
Definition mreg_within_bounds (r: mreg) :=
- match mreg_type r with
- | Tint => index_int_callee_save r < bound_int_callee_save b
- | Tfloat => index_float_callee_save r < bound_float_callee_save b
- end.
-
-Definition slot_within_bounds (s: slot) :=
- match s with
- | Local ofs Tint => 0 <= ofs < bound_int_local b
- | Local ofs Tfloat => 0 <= ofs < bound_float_local b
- | Outgoing ofs ty => 0 <= ofs /\ ofs + typesize ty <= bound_outgoing b
- | Incoming ofs ty => In (S s) (loc_parameters funct.(fn_sig))
+ index_int_callee_save r < bound_int_callee_save b
+ /\ index_float_callee_save r < bound_float_callee_save b.
+
+Definition slot_within_bounds (sl: slot) (ofs: Z) (ty: typ) :=
+ match sl with
+ | Local => ofs + typesize ty <= bound_local b
+ | Outgoing => ofs + typesize ty <= bound_outgoing b
+ | Incoming => True
end.
Definition instr_within_bounds (i: instruction) :=
match i with
- | Lgetstack s r => slot_within_bounds s /\ mreg_within_bounds r
- | Lsetstack r s => slot_within_bounds s
+ | Lgetstack sl ofs ty r => slot_within_bounds sl ofs ty /\ mreg_within_bounds r
+ | Lsetstack r sl ofs ty => slot_within_bounds sl ofs ty
| Lop op args res => mreg_within_bounds res
| Lload chunk addr args dst => mreg_within_bounds dst
| Lcall sig ros => size_arguments sig <= bound_outgoing b
- | Lbuiltin ef args res => mreg_within_bounds res
- | Lannot ef args => forall s, In (S s) args -> slot_within_bounds s
+ | Lbuiltin ef args res =>
+ forall r, In r res \/ In r (destroyed_by_builtin ef) -> mreg_within_bounds r
+ | Lannot ef args =>
+ forall sl ofs ty, In (S sl ofs ty) args -> slot_within_bounds sl ofs ty
| _ => True
end.
-End BELOW.
+End WITHIN_BOUNDS.
Definition function_within_bounds (f: function) (b: bounds) : Prop :=
- forall instr, In instr f.(fn_code) -> instr_within_bounds f b instr.
+ forall instr, In instr f.(fn_code) -> instr_within_bounds b instr.
(** * Inference of resource bounds for a function *)
@@ -99,14 +94,14 @@ Variable f: function.
Definition regs_of_instr (i: instruction) : list mreg :=
match i with
- | Lgetstack s r => r :: nil
- | Lsetstack r s => r :: nil
+ | Lgetstack sl ofs ty r => r :: nil
+ | Lsetstack r sl ofs ty => r :: nil
| Lop op args res => res :: nil
| Lload chunk addr args dst => dst :: nil
| Lstore chunk addr args src => nil
| Lcall sig ros => nil
| Ltailcall sig ros => nil
- | Lbuiltin ef args res => res :: nil
+ | Lbuiltin ef args res => res ++ destroyed_by_builtin ef
| Lannot ef args => nil
| Llabel lbl => nil
| Lgoto lbl => nil
@@ -115,58 +110,55 @@ Definition regs_of_instr (i: instruction) : list mreg :=
| Lreturn => nil
end.
-Fixpoint slots_of_locs (l: list loc) : list slot :=
+Fixpoint slots_of_locs (l: list loc) : list (slot * Z * typ) :=
match l with
| nil => nil
- | S s :: l' => s :: slots_of_locs l'
+ | S sl ofs ty :: l' => (sl, ofs, ty) :: slots_of_locs l'
| R r :: l' => slots_of_locs l'
end.
-Definition slots_of_instr (i: instruction) : list slot :=
+Definition slots_of_instr (i: instruction) : list (slot * Z * typ) :=
match i with
- | Lgetstack s r => s :: nil
- | Lsetstack r s => s :: nil
+ | Lgetstack sl ofs ty r => (sl, ofs, ty) :: nil
+ | Lsetstack r sl ofs ty => (sl, ofs, ty) :: nil
| Lannot ef args => slots_of_locs args
| _ => nil
end.
-Definition max_over_list (A: Type) (valu: A -> Z) (l: list A) : Z :=
+Definition max_over_list {A: Type} (valu: A -> Z) (l: list A) : Z :=
List.fold_left (fun m l => Zmax m (valu l)) l 0.
Definition max_over_instrs (valu: instruction -> Z) : Z :=
- max_over_list instruction valu f.(fn_code).
+ max_over_list valu f.(fn_code).
Definition max_over_regs_of_instr (valu: mreg -> Z) (i: instruction) : Z :=
- max_over_list mreg valu (regs_of_instr i).
+ max_over_list valu (regs_of_instr i).
-Definition max_over_slots_of_instr (valu: slot -> Z) (i: instruction) : Z :=
- max_over_list slot valu (slots_of_instr i).
+Definition max_over_slots_of_instr (valu: slot * Z * typ -> Z) (i: instruction) : Z :=
+ max_over_list valu (slots_of_instr i).
Definition max_over_regs_of_funct (valu: mreg -> Z) : Z :=
max_over_instrs (max_over_regs_of_instr valu).
-Definition max_over_slots_of_funct (valu: slot -> Z) : Z :=
+Definition max_over_slots_of_funct (valu: slot * Z * typ -> Z) : Z :=
max_over_instrs (max_over_slots_of_instr valu).
Definition int_callee_save (r: mreg) := 1 + index_int_callee_save r.
Definition float_callee_save (r: mreg) := 1 + index_float_callee_save r.
-Definition int_local (s: slot) :=
- match s with Local ofs Tint => 1 + ofs | _ => 0 end.
-
-Definition float_local (s: slot) :=
- match s with Local ofs Tfloat => 1 + ofs | _ => 0 end.
+Definition local_slot (s: slot * Z * typ) :=
+ match s with (Local, ofs, ty) => ofs + typesize ty | _ => 0 end.
-Definition outgoing_slot (s: slot) :=
- match s with Outgoing ofs ty => ofs + typesize ty | _ => 0 end.
+Definition outgoing_slot (s: slot * Z * typ) :=
+ match s with (Outgoing, ofs, ty) => ofs + typesize ty | _ => 0 end.
Definition outgoing_space (i: instruction) :=
match i with Lcall sig _ => size_arguments sig | _ => 0 end.
Lemma max_over_list_pos:
forall (A: Type) (valu: A -> Z) (l: list A),
- max_over_list A valu l >= 0.
+ max_over_list valu l >= 0.
Proof.
intros until valu. unfold max_over_list.
assert (forall l z, fold_left (fun x y => Zmax x (valu y)) l z >= z).
@@ -176,7 +168,7 @@ Proof.
Qed.
Lemma max_over_slots_of_funct_pos:
- forall (valu: slot -> Z), max_over_slots_of_funct valu >= 0.
+ forall (valu: slot * Z * typ -> Z), max_over_slots_of_funct valu >= 0.
Proof.
intros. unfold max_over_slots_of_funct.
unfold max_over_instrs. apply max_over_list_pos.
@@ -188,18 +180,16 @@ Proof.
intros. unfold max_over_regs_of_funct.
unfold max_over_instrs. apply max_over_list_pos.
Qed.
-
+
Program Definition function_bounds :=
mkbounds
- (max_over_slots_of_funct int_local)
- (max_over_slots_of_funct float_local)
+ (max_over_slots_of_funct local_slot)
(max_over_regs_of_funct int_callee_save)
(max_over_regs_of_funct float_callee_save)
(Zmax (max_over_instrs outgoing_space)
(max_over_slots_of_funct outgoing_slot))
(Zmax f.(fn_stacksize) 0)
- (max_over_slots_of_funct_pos int_local)
- (max_over_slots_of_funct_pos float_local)
+ (max_over_slots_of_funct_pos local_slot)
(max_over_regs_of_funct_pos int_callee_save)
(max_over_regs_of_funct_pos float_callee_save)
_ _.
@@ -215,7 +205,7 @@ Qed.
Lemma max_over_list_bound:
forall (A: Type) (valu: A -> Z) (l: list A) (x: A),
- In x l -> valu x <= max_over_list A valu l.
+ In x l -> valu x <= max_over_list valu l.
Proof.
intros until x. unfold max_over_list.
assert (forall c z,
@@ -250,7 +240,7 @@ Proof.
Qed.
Lemma max_over_slots_of_funct_bound:
- forall (valu: slot -> Z) i s,
+ forall (valu: slot * Z * typ -> Z) i s,
In i f.(fn_code) -> In s (slots_of_instr i) ->
valu s <= max_over_slots_of_funct valu.
Proof.
@@ -282,34 +272,23 @@ Proof.
eapply max_over_regs_of_funct_bound; eauto.
Qed.
-Lemma int_local_slot_bound:
- forall i ofs,
- In i f.(fn_code) -> In (Local ofs Tint) (slots_of_instr i) ->
- ofs < bound_int_local function_bounds.
-Proof.
- intros. apply Zlt_le_trans with (int_local (Local ofs Tint)).
- unfold int_local. omega.
- unfold function_bounds, bound_int_local.
- eapply max_over_slots_of_funct_bound; eauto.
-Qed.
-
-Lemma float_local_slot_bound:
- forall i ofs,
- In i f.(fn_code) -> In (Local ofs Tfloat) (slots_of_instr i) ->
- ofs < bound_float_local function_bounds.
+Lemma local_slot_bound:
+ forall i ofs ty,
+ In i f.(fn_code) -> In (Local, ofs, ty) (slots_of_instr i) ->
+ ofs + typesize ty <= bound_local function_bounds.
Proof.
- intros. apply Zlt_le_trans with (float_local (Local ofs Tfloat)).
- unfold float_local. omega.
- unfold function_bounds, bound_float_local.
+ intros.
+ unfold function_bounds, bound_local.
+ change (ofs + typesize ty) with (local_slot (Local, ofs, ty)).
eapply max_over_slots_of_funct_bound; eauto.
Qed.
Lemma outgoing_slot_bound:
forall i ofs ty,
- In i f.(fn_code) -> In (Outgoing ofs ty) (slots_of_instr i) ->
+ In i f.(fn_code) -> In (Outgoing, ofs, ty) (slots_of_instr i) ->
ofs + typesize ty <= bound_outgoing function_bounds.
Proof.
- intros. change (ofs + typesize ty) with (outgoing_slot (Outgoing ofs ty)).
+ intros. change (ofs + typesize ty) with (outgoing_slot (Outgoing, ofs, ty)).
unfold function_bounds, bound_outgoing.
apply Zmax_bound_r. eapply max_over_slots_of_funct_bound; eauto.
Qed.
@@ -332,28 +311,25 @@ Lemma mreg_is_within_bounds:
forall r, In r (regs_of_instr i) ->
mreg_within_bounds function_bounds r.
Proof.
- intros. unfold mreg_within_bounds.
- case (mreg_type r).
+ intros. unfold mreg_within_bounds. split.
eapply int_callee_save_bound; eauto.
eapply float_callee_save_bound; eauto.
Qed.
Lemma slot_is_within_bounds:
forall i, In i f.(fn_code) ->
- forall s, In s (slots_of_instr i) -> Lineartyping.slot_valid f s ->
- slot_within_bounds f function_bounds s.
+ forall sl ty ofs, In (sl, ofs, ty) (slots_of_instr i) ->
+ slot_within_bounds function_bounds sl ofs ty.
Proof.
intros. unfold slot_within_bounds.
- destruct s.
- destruct t.
- split. exact H1. eapply int_local_slot_bound; eauto.
- split. exact H1. eapply float_local_slot_bound; eauto.
- exact H1.
- split. simpl in H1. exact H1. eapply outgoing_slot_bound; eauto.
+ destruct sl.
+ eapply local_slot_bound; eauto.
+ auto.
+ eapply outgoing_slot_bound; eauto.
Qed.
Lemma slots_of_locs_charact:
- forall s l, In s (slots_of_locs l) <-> In (S s) l.
+ forall sl ofs ty l, In (sl, ofs, ty) (slots_of_locs l) <-> In (S sl ofs ty) l.
Proof.
induction l; simpl; intros.
tauto.
@@ -366,27 +342,21 @@ Qed.
Lemma instr_is_within_bounds:
forall i,
In i f.(fn_code) ->
- Lineartyping.wt_instr f i ->
- instr_within_bounds f function_bounds i.
+ instr_within_bounds function_bounds i.
Proof.
intros;
destruct i;
generalize (mreg_is_within_bounds _ H); generalize (slot_is_within_bounds _ H);
simpl; intros; auto.
-(* getstack *)
- inv H0. split; auto.
-(* setstack *)
- inv H0; auto.
(* call *)
eapply size_arguments_bound; eauto.
+(* builtin *)
+ apply H1. apply in_or_app; auto.
(* annot *)
- inv H0. apply H1. rewrite slots_of_locs_charact; auto.
- generalize (H8 _ H3). unfold loc_acceptable, slot_valid.
- destruct s; (contradiction || omega).
+ apply H0. rewrite slots_of_locs_charact; auto.
Qed.
Lemma function_is_within_bounds:
- Lineartyping.wt_code f f.(fn_code) ->
function_within_bounds f function_bounds.
Proof.
intros; red; intros. apply instr_is_within_bounds; auto.
diff --git a/backend/CMtypecheck.ml b/backend/CMtypecheck.ml
index 39e8c51..1a19f71 100644
--- a/backend/CMtypecheck.ml
+++ b/backend/CMtypecheck.ml
@@ -24,15 +24,16 @@ open Cminor
exception Error of string
-let name_of_typ = function Tint -> "int" | Tfloat -> "float"
+let name_of_typ = function Tint -> "int" | Tfloat -> "float" | Tlong -> "long"
type ty = Base of typ | Var of ty option ref
let newvar () = Var (ref None)
let tint = Base Tint
let tfloat = Base Tfloat
+let tlong = Base Tlong
-let ty_of_typ = function Tint -> tint | Tfloat -> tfloat
+let ty_of_typ = function Tint -> tint | Tfloat -> tfloat | Tlong -> tlong
let ty_of_sig_args tyl = List.map ty_of_typ tyl
@@ -81,6 +82,7 @@ let name_of_comparison = function
let type_constant = function
| Ointconst _ -> tint
| Ofloatconst _ -> tfloat
+ | Olongconst _ -> tlong
| Oaddrsymbol _ -> tint
| Oaddrstack _ -> tint
@@ -98,6 +100,15 @@ let type_unary_operation = function
| Ointuoffloat -> tfloat, tint
| Ofloatofint -> tint, tfloat
| Ofloatofintu -> tint, tfloat
+ | Onegl -> tlong, tlong
+ | Onotl -> tlong, tlong
+ | Ointoflong -> tlong, tint
+ | Olongofint -> tint, tlong
+ | Olongofintu -> tint, tlong
+ | Olongoffloat -> tfloat, tlong
+ | Olonguoffloat -> tfloat, tlong
+ | Ofloatoflong -> tlong, tfloat
+ | Ofloatoflongu -> tlong, tfloat
let type_binary_operation = function
| Oadd -> tint, tint, tint
@@ -117,52 +128,28 @@ let type_binary_operation = function
| Osubf -> tfloat, tfloat, tfloat
| Omulf -> tfloat, tfloat, tfloat
| Odivf -> tfloat, tfloat, tfloat
+ | Oaddl -> tlong, tlong, tlong
+ | Osubl -> tlong, tlong, tlong
+ | Omull -> tlong, tlong, tlong
+ | Odivl -> tlong, tlong, tlong
+ | Odivlu -> tlong, tlong, tlong
+ | Omodl -> tlong, tlong, tlong
+ | Omodlu -> tlong, tlong, tlong
+ | Oandl -> tlong, tlong, tlong
+ | Oorl -> tlong, tlong, tlong
+ | Oxorl -> tlong, tlong, tlong
+ | Oshll -> tlong, tint, tlong
+ | Oshrl -> tlong, tint, tlong
+ | Oshrlu -> tlong, tint, tlong
| Ocmp _ -> tint, tint, tint
| Ocmpu _ -> tint, tint, tint
| Ocmpf _ -> tfloat, tfloat, tint
+ | Ocmpl _ -> tlong, tlong, tint
+ | Ocmplu _ -> tlong, tlong, tint
-let name_of_constant = function
- | Ointconst n -> sprintf "intconst %ld" (camlint_of_coqint n)
- | Ofloatconst n -> sprintf "floatconst %g" (camlfloat_of_coqfloat n)
- | Oaddrsymbol (s, ofs) -> sprintf "addrsymbol %s %ld" (extern_atom s) (camlint_of_coqint ofs)
- | Oaddrstack n -> sprintf "addrstack %ld" (camlint_of_coqint n)
+let name_of_unary_operation = PrintCminor.name_of_unop
-let name_of_unary_operation = function
- | Ocast8signed -> "cast8signed"
- | Ocast16signed -> "cast16signed"
- | Ocast8unsigned -> "cast8unsigned"
- | Ocast16unsigned -> "cast16unsigned"
- | Onegint -> "negint"
- | Onotint -> "notint"
- | Onegf -> "negf"
- | Oabsf -> "absf"
- | Osingleoffloat -> "singleoffloat"
- | Ointoffloat -> "intoffloat"
- | Ointuoffloat -> "intuoffloat"
- | Ofloatofint -> "floatofint"
- | Ofloatofintu -> "floatofintu"
-
-let name_of_binary_operation = function
- | Oadd -> "add"
- | Osub -> "sub"
- | Omul -> "mul"
- | Odiv -> "div"
- | Odivu -> "divu"
- | Omod -> "mod"
- | Omodu -> "modu"
- | Oand -> "and"
- | Oor -> "or"
- | Oxor -> "xor"
- | Oshl -> "shl"
- | Oshr -> "shr"
- | Oshru -> "shru"
- | Oaddf -> "addf"
- | Osubf -> "subf"
- | Omulf -> "mulf"
- | Odivf -> "divf"
- | Ocmp c -> sprintf "cmp %s" (name_of_comparison c)
- | Ocmpu c -> sprintf "cmpu %s" (name_of_comparison c)
- | Ocmpf c -> sprintf "cmpf %s" (name_of_comparison c)
+let name_of_binary_operation = PrintCminor.name_of_binop
let type_chunk = function
| Mint8signed -> tint
@@ -170,19 +157,12 @@ let type_chunk = function
| Mint16signed -> tint
| Mint16unsigned -> tint
| Mint32 -> tint
+ | Mint64 -> tlong
| Mfloat32 -> tfloat
| Mfloat64 -> tfloat
| Mfloat64al32 -> tfloat
-let name_of_chunk = function
- | Mint8signed -> "int8signed"
- | Mint8unsigned -> "int8unsigned"
- | Mint16signed -> "int16signed"
- | Mint16unsigned -> "int16unsigned"
- | Mint32 -> "int32"
- | Mfloat32 -> "float32"
- | Mfloat64 -> "float64"
- | Mfloat64al32 -> "float64al32"
+let name_of_chunk = PrintAST.name_of_chunk
let rec type_expr env lenv e =
match e with
diff --git a/backend/CSEproof.v b/backend/CSEproof.v
index 8fc9407..1e269f8 100644
--- a/backend/CSEproof.v
+++ b/backend/CSEproof.v
@@ -996,12 +996,7 @@ Proof.
rewrite <- RES. apply eval_operation_preserved. exact symbols_preserved.
(* state matching *)
econstructor; eauto.
- apply wt_regset_assign; auto.
- generalize (wt_instrs _ _ WTF pc _ H); intro WTI; inv WTI.
- simpl in H0. inv H0. rewrite <- H3. apply WTREGS.
- replace (tyenv res) with (snd (type_of_operation op)).
- eapply type_of_operation_sound; eauto.
- rewrite <- H6. reflexivity.
+ eapply wt_exec_Iop; eauto. eapply wt_instrs; eauto.
eapply analysis_correct_1; eauto. simpl; auto.
unfold transfer; rewrite H.
eapply add_op_satisfiable; eauto. eapply wf_analyze; eauto.
@@ -1028,8 +1023,7 @@ Proof.
eapply exec_Iload; eauto.
(* state matching *)
econstructor; eauto.
- generalize (wt_instrs _ _ WTF pc _ H); intro WTI; inv WTI.
- apply wt_regset_assign. auto. rewrite H8. eapply type_of_chunk_correct; eauto.
+ eapply wt_exec_Iload; eauto. eapply wt_instrs; eauto.
eapply analysis_correct_1; eauto. simpl; auto.
unfold transfer; rewrite H.
eapply add_load_satisfiable; eauto. eapply wf_analyze; eauto.
diff --git a/backend/CleanupLabels.v b/backend/CleanupLabels.v
index 8db871e..5eaa81e 100644
--- a/backend/CleanupLabels.v
+++ b/backend/CleanupLabels.v
@@ -16,15 +16,15 @@
control-flow graph. Many of these labels are never branched to,
which can complicate further optimizations over linearized code.
(There are no such optimizations yet.) In preparation for these
- further optimizations, and to make the generated LTLin code
+ further optimizations, and to make the generated Linear code
better-looking, the present pass removes labels that cannot be
branched to. *)
-Require Import Coqlib.
-Require Import Ordered.
Require Import FSets.
Require FSetAVL.
-Require Import LTLin.
+Require Import Coqlib.
+Require Import Ordered.
+Require Import Linear.
Module Labelset := FSetAVL.Make(OrderedPositive).
@@ -63,7 +63,6 @@ Definition cleanup_labels (c: code) :=
Definition transf_function (f: function) : function :=
mkfunction
(fn_sig f)
- (fn_params f)
(fn_stacksize f)
(cleanup_labels (fn_code f)).
diff --git a/backend/CleanupLabelsproof.v b/backend/CleanupLabelsproof.v
index 70f0eb3..65ba61c 100644
--- a/backend/CleanupLabelsproof.v
+++ b/backend/CleanupLabelsproof.v
@@ -24,7 +24,7 @@ Require Import Globalenvs.
Require Import Smallstep.
Require Import Op.
Require Import Locations.
-Require Import LTLin.
+Require Import Linear.
Require Import CleanupLabels.
Module LabelsetFacts := FSetFacts.Facts(Labelset).
@@ -103,7 +103,7 @@ Proof.
intros; red; intros; destruct i; simpl; auto.
apply Labelset.add_2; auto.
apply Labelset.add_2; auto.
- revert H; induction l0; simpl. auto. intros; apply Labelset.add_2; auto.
+ revert H; induction l; simpl. auto. intros; apply Labelset.add_2; auto.
Qed.
Remark add_label_branched_to_contains:
@@ -114,7 +114,7 @@ Proof.
destruct i; simpl; intros; try contradiction.
apply Labelset.add_1; auto.
apply Labelset.add_1; auto.
- revert H. induction l0; simpl; intros.
+ revert H. induction l; simpl; intros.
contradiction.
destruct H. apply Labelset.add_1; auto. apply Labelset.add_2; auto.
Qed.
@@ -200,11 +200,11 @@ Qed.
Inductive match_stackframes: stackframe -> stackframe -> Prop :=
| match_stackframe_intro:
- forall res f sp ls c,
+ forall f sp ls c,
incl c f.(fn_code) ->
match_stackframes
- (Stackframe res f sp ls c)
- (Stackframe res (transf_function f) sp ls
+ (Stackframe f sp ls c)
+ (Stackframe (transf_function f) sp ls
(remove_unused_labels (labels_branched_to f.(fn_code)) c)).
Inductive match_states: state -> state -> Prop :=
@@ -231,6 +231,14 @@ Definition measure (st: state) : nat :=
| _ => O
end.
+Lemma match_parent_locset:
+ forall s ts,
+ list_forall2 match_stackframes s ts ->
+ parent_locset ts = parent_locset s.
+Proof.
+ induction 1; simpl. auto. inv H; auto.
+Qed.
+
Theorem transf_step_correct:
forall s1 t s2, step ge s1 t s2 ->
forall s1' (MS: match_states s1 s1'),
@@ -238,19 +246,27 @@ Theorem transf_step_correct:
\/ (measure s2 < measure s1 /\ t = E0 /\ match_states s2 s1')%nat.
Proof.
induction 1; intros; inv MS; try rewrite remove_unused_labels_cons.
+(* Lgetstack *)
+ left; econstructor; split.
+ econstructor; eauto.
+ econstructor; eauto with coqlib.
+(* Lsetstack *)
+ left; econstructor; split.
+ econstructor; eauto.
+ econstructor; eauto with coqlib.
(* Lop *)
left; econstructor; split.
econstructor; eauto. instantiate (1 := v). rewrite <- H.
apply eval_operation_preserved. exact symbols_preserved.
econstructor; eauto with coqlib.
(* Lload *)
- assert (eval_addressing tge sp addr (map rs args) = Some a).
+ assert (eval_addressing tge sp addr (LTL.reglist rs args) = Some a).
rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
left; econstructor; split.
econstructor; eauto.
econstructor; eauto with coqlib.
(* Lstore *)
- assert (eval_addressing tge sp addr (map rs args) = Some a).
+ assert (eval_addressing tge sp addr (LTL.reglist rs args) = Some a).
rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
left; econstructor; split.
econstructor; eauto.
@@ -262,13 +278,18 @@ Proof.
econstructor; eauto. constructor; auto. constructor; eauto with coqlib.
(* Ltailcall *)
left; econstructor; split.
- econstructor. eapply find_function_translated; eauto.
+ econstructor. erewrite match_parent_locset; eauto. eapply find_function_translated; eauto.
symmetry; apply sig_function_translated.
simpl. eauto.
econstructor; eauto.
(* Lbuiltin *)
left; econstructor; split.
- econstructor; eauto. eapply external_call_symbols_preserved; eauto.
+ econstructor; eauto. eapply external_call_symbols_preserved'; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ econstructor; eauto with coqlib.
+(* Lannot *)
+ left; econstructor; split.
+ econstructor; eauto. eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto with coqlib.
(* Llabel *)
@@ -285,7 +306,7 @@ Proof.
econstructor; eauto. eapply find_label_incl; eauto.
(* Lcond taken *)
left; econstructor; split.
- econstructor. auto. eapply find_label_translated; eauto. red; auto.
+ econstructor. auto. eauto. eapply find_label_translated; eauto. red; auto.
econstructor; eauto. eapply find_label_incl; eauto.
(* Lcond not taken *)
left; econstructor; split.
@@ -294,11 +315,12 @@ Proof.
(* Ljumptable *)
left; econstructor; split.
econstructor. eauto. eauto. eapply find_label_translated; eauto.
- red. eapply list_nth_z_in; eauto.
+ red. eapply list_nth_z_in; eauto. eauto.
econstructor; eauto. eapply find_label_incl; eauto.
(* Lreturn *)
left; econstructor; split.
econstructor; eauto.
+ erewrite <- match_parent_locset; eauto.
econstructor; eauto with coqlib.
(* internal function *)
left; econstructor; split.
@@ -306,7 +328,7 @@ Proof.
econstructor; eauto with coqlib.
(* external function *)
left; econstructor; split.
- econstructor; eauto. eapply external_call_symbols_preserved; eauto.
+ econstructor; eauto. eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto with coqlib.
(* return *)
@@ -333,11 +355,11 @@ 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 H6. econstructor; eauto.
Qed.
Theorem transf_program_correct:
- forward_simulation (LTLin.semantics prog) (LTLin.semantics tprog).
+ forward_simulation (Linear.semantics prog) (Linear.semantics tprog).
Proof.
eapply forward_simulation_opt.
eexact symbols_preserved.
diff --git a/backend/CleanupLabelstyping.v b/backend/CleanupLabelstyping.v
deleted file mode 100644
index 11b516f..0000000
--- a/backend/CleanupLabelstyping.v
+++ /dev/null
@@ -1,59 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Type preservation for the CleanupLabels pass *)
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Op.
-Require Import Locations.
-Require Import LTLin.
-Require Import CleanupLabels.
-Require Import LTLintyping.
-
-Lemma in_remove_unused_labels:
- forall bto i c, In i (remove_unused_labels bto c) -> In i c.
-Proof.
- unfold remove_unused_labels, remove_unused. induction c; simpl.
- auto.
- rewrite list_fold_right_eq. destruct a; simpl; intuition.
- destruct (Labelset.mem l bto); simpl in H; intuition.
-Qed.
-
-Lemma wt_transf_function:
- forall f,
- wt_function f ->
- wt_function (transf_function f).
-Proof.
- intros. inv H. constructor; simpl; auto.
- unfold cleanup_labels; red; intros.
- apply wt_instrs. eapply in_remove_unused_labels; eauto.
-Qed.
-
-Lemma wt_transf_fundef:
- forall f,
- wt_fundef f ->
- wt_fundef (transf_fundef f).
-Proof.
- induction 1. constructor. constructor. apply wt_transf_function; auto.
-Qed.
-
-Lemma program_typing_preserved:
- forall p,
- wt_program p ->
- wt_program (transf_program p).
-Proof.
- intros; red; intros.
- exploit transform_program_function; eauto. intros [f1 [A B]]. subst f.
- apply wt_transf_fundef. eapply H; eauto.
-Qed.
diff --git a/backend/Cminor.v b/backend/Cminor.v
index 3d177e4..3963e76 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -36,6 +36,7 @@ Require Import Switch.
Inductive constant : Type :=
| Ointconst: int -> constant (**r integer constant *)
| Ofloatconst: float -> constant (**r floating-point constant *)
+ | Olongconst: int64 -> constant (**r long integer constant *)
| Oaddrsymbol: ident -> int -> constant (**r address of the symbol plus the offset *)
| Oaddrstack: int -> constant. (**r stack pointer plus the given offset *)
@@ -52,7 +53,16 @@ Inductive unary_operation : Type :=
| Ointoffloat: unary_operation (**r signed integer to float *)
| Ointuoffloat: unary_operation (**r unsigned integer to float *)
| Ofloatofint: unary_operation (**r float to signed integer *)
- | Ofloatofintu: unary_operation. (**r float to unsigned integer *)
+ | Ofloatofintu: unary_operation (**r float to unsigned integer *)
+ | Onegl: unary_operation (**r long integer opposite *)
+ | Onotl: unary_operation (**r long bitwise complement *)
+ | Ointoflong: unary_operation (**r long to int *)
+ | Olongofint: unary_operation (**r signed int to long *)
+ | Olongofintu: unary_operation (**r unsigned int to long *)
+ | Olongoffloat: unary_operation (**r signed long to float *)
+ | Olonguoffloat: unary_operation (**r unsigned long to float *)
+ | Ofloatoflong: unary_operation (**r float to signed long *)
+ | Ofloatoflongu: unary_operation. (**r float to unsigned long *)
Inductive binary_operation : Type :=
| Oadd: binary_operation (**r integer addition *)
@@ -62,19 +72,34 @@ Inductive binary_operation : Type :=
| Odivu: binary_operation (**r integer unsigned division *)
| Omod: binary_operation (**r integer signed modulus *)
| Omodu: binary_operation (**r integer unsigned modulus *)
- | Oand: binary_operation (**r bitwise ``and'' *)
- | Oor: binary_operation (**r bitwise ``or'' *)
- | Oxor: binary_operation (**r bitwise ``xor'' *)
- | Oshl: binary_operation (**r left shift *)
- | Oshr: binary_operation (**r right signed shift *)
- | Oshru: binary_operation (**r right unsigned shift *)
+ | Oand: binary_operation (**r integer bitwise ``and'' *)
+ | Oor: binary_operation (**r integer bitwise ``or'' *)
+ | Oxor: binary_operation (**r integer bitwise ``xor'' *)
+ | Oshl: binary_operation (**r integer left shift *)
+ | Oshr: binary_operation (**r integer right signed shift *)
+ | Oshru: binary_operation (**r integer right unsigned shift *)
| Oaddf: binary_operation (**r float addition *)
| Osubf: binary_operation (**r float subtraction *)
| Omulf: binary_operation (**r float multiplication *)
| Odivf: binary_operation (**r float division *)
+ | Oaddl: binary_operation (**r long addition *)
+ | Osubl: binary_operation (**r long subtraction *)
+ | Omull: binary_operation (**r long multiplication *)
+ | Odivl: binary_operation (**r long signed division *)
+ | Odivlu: binary_operation (**r long unsigned division *)
+ | Omodl: binary_operation (**r long signed modulus *)
+ | Omodlu: binary_operation (**r long unsigned modulus *)
+ | Oandl: binary_operation (**r long bitwise ``and'' *)
+ | Oorl: binary_operation (**r long bitwise ``or'' *)
+ | Oxorl: binary_operation (**r long bitwise ``xor'' *)
+ | Oshll: binary_operation (**r long left shift *)
+ | Oshrl: binary_operation (**r long right signed shift *)
+ | Oshrlu: binary_operation (**r long right unsigned shift *)
| Ocmp: comparison -> binary_operation (**r integer signed comparison *)
| Ocmpu: comparison -> binary_operation (**r integer unsigned comparison *)
- | Ocmpf: comparison -> binary_operation. (**r float comparison *)
+ | Ocmpf: comparison -> binary_operation (**r float comparison *)
+ | Ocmpl: comparison -> binary_operation (**r long signed comparison *)
+ | Ocmplu: comparison -> binary_operation. (**r long unsigned comparison *)
(** Expressions include reading local variables, constants and
arithmetic operations, reading store locations, and conditional
@@ -214,6 +239,7 @@ Definition eval_constant (sp: val) (cst: constant) : option val :=
match cst with
| Ointconst n => Some (Vint n)
| Ofloatconst n => Some (Vfloat n)
+ | Olongconst n => Some (Vlong n)
| Oaddrsymbol s ofs =>
Some(match Genv.find_symbol ge s with
| None => Vundef
@@ -236,6 +262,15 @@ Definition eval_unop (op: unary_operation) (arg: val) : option val :=
| Ointuoffloat => Val.intuoffloat arg
| Ofloatofint => Val.floatofint arg
| Ofloatofintu => Val.floatofintu arg
+ | Onegl => Some (Val.negl arg)
+ | Onotl => Some (Val.notl arg)
+ | Ointoflong => Some (Val.loword arg)
+ | Olongofint => Some (Val.longofint arg)
+ | Olongofintu => Some (Val.longofintu arg)
+ | Olongoffloat => Val.longoffloat arg
+ | Olonguoffloat => Val.longuoffloat arg
+ | Ofloatoflong => Val.floatoflong arg
+ | Ofloatoflongu => Val.floatoflongu arg
end.
Definition eval_binop
@@ -258,9 +293,24 @@ Definition eval_binop
| Osubf => Some (Val.subf arg1 arg2)
| Omulf => Some (Val.mulf arg1 arg2)
| Odivf => Some (Val.divf arg1 arg2)
+ | Oaddl => Some (Val.addl arg1 arg2)
+ | Osubl => Some (Val.subl arg1 arg2)
+ | Omull => Some (Val.mull arg1 arg2)
+ | Odivl => Val.divls arg1 arg2
+ | Odivlu => Val.divlu arg1 arg2
+ | Omodl => Val.modls arg1 arg2
+ | Omodlu => Val.modlu arg1 arg2
+ | Oandl => Some (Val.andl arg1 arg2)
+ | Oorl => Some (Val.orl arg1 arg2)
+ | Oxorl => Some (Val.xorl arg1 arg2)
+ | Oshll => Some (Val.shll arg1 arg2)
+ | Oshrl => Some (Val.shrl arg1 arg2)
+ | Oshrlu => Some (Val.shrlu arg1 arg2)
| Ocmp c => Some (Val.cmp c arg1 arg2)
| Ocmpu c => Some (Val.cmpu (Mem.valid_pointer m) c arg1 arg2)
| Ocmpf c => Some (Val.cmpf c arg1 arg2)
+ | Ocmpl c => Some (Val.cmpl c arg1 arg2)
+ | Ocmplu c => Some (Val.cmplu c arg1 arg2)
end.
(** Evaluation of an expression: [eval_expr ge sp e m a v]
diff --git a/backend/CminorSel.v b/backend/CminorSel.v
index b5a0d39..3538dda 100644
--- a/backend/CminorSel.v
+++ b/backend/CminorSel.v
@@ -43,6 +43,8 @@ Inductive expr : Type :=
| Econdition : condition -> exprlist -> expr -> expr -> expr
| Elet : expr -> expr -> expr
| Eletvar : nat -> expr
+ | Ebuiltin : external_function -> exprlist -> expr
+ | Eexternal : ident -> signature -> exprlist -> expr
with exprlist : Type :=
| Enil: exprlist
@@ -171,6 +173,17 @@ Inductive eval_expr: letenv -> expr -> val -> Prop :=
| eval_Eletvar: forall le n v,
nth_error le n = Some v ->
eval_expr le (Eletvar n) v
+ | eval_Ebuiltin: forall le ef al vl v,
+ eval_exprlist le al vl ->
+ external_call ef ge vl m E0 v m ->
+ eval_expr le (Ebuiltin ef al) v
+ | eval_Eexternal: forall le id sg al b ef vl v,
+ Genv.find_symbol ge id = Some b ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ ef_sig ef = sg ->
+ eval_exprlist le al vl ->
+ external_call ef ge vl m E0 v m ->
+ eval_expr le (Eexternal id sg al) v
with eval_exprlist: letenv -> exprlist -> list val -> Prop :=
| eval_Enil: forall le,
@@ -388,6 +401,8 @@ Fixpoint lift_expr (p: nat) (a: expr) {struct a}: expr :=
| Elet b c => Elet (lift_expr p b) (lift_expr (S p) c)
| Eletvar n =>
if le_gt_dec p n then Eletvar (S n) else Eletvar n
+ | Ebuiltin ef bl => Ebuiltin ef (lift_exprlist p bl)
+ | Eexternal id sg bl => Eexternal id sg (lift_exprlist p bl)
end
with lift_exprlist (p: nat) (a: exprlist) {struct a}: exprlist :=
diff --git a/backend/Constprop.v b/backend/Constprop.v
index f85405d..0575079 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -47,6 +47,7 @@ Module Approx <: SEMILATTICE_WITH_TOP.
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.
@@ -139,6 +140,10 @@ Fixpoint eval_load_init (chunk: memory_chunk) (pos: Z) (il: list init_data): app
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
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index 580d551..a6385f4 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -186,6 +186,11 @@ Proof.
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.
diff --git a/backend/Conventions.v b/backend/Conventions.v
index c11bf47..abfe4ee 100644
--- a/backend/Conventions.v
+++ b/backend/Conventions.v
@@ -22,110 +22,6 @@ Require Export Conventions1.
[arch/abi/Conventions1.v]. This file adds various processor-independent
definitions and lemmas. *)
-(** * Acceptable locations for register allocation *)
-
-(** The following predicate describes the locations that can be assigned
- to an RTL pseudo-register during register allocation: a non-temporary
- machine register or a [Local] stack slot are acceptable. *)
-
-Definition loc_acceptable (l: loc) : Prop :=
- match l with
- | R r => ~(In l temporaries)
- | S (Local ofs ty) => ofs >= 0
- | S (Incoming _ _) => False
- | S (Outgoing _ _) => False
- end.
-
-Definition locs_acceptable (ll: list loc) : Prop :=
- forall l, In l ll -> loc_acceptable l.
-
-Lemma temporaries_not_acceptable:
- forall l, loc_acceptable l -> Loc.notin l temporaries.
-Proof.
- unfold loc_acceptable; destruct l.
- simpl. intuition congruence.
- destruct s; try contradiction.
- intro. simpl. tauto.
-Qed.
-Hint Resolve temporaries_not_acceptable: locs.
-
-Lemma locs_acceptable_disj_temporaries:
- forall ll, locs_acceptable ll -> Loc.disjoint ll temporaries.
-Proof.
- intros. apply Loc.notin_disjoint. intros.
- apply temporaries_not_acceptable. auto.
-Qed.
-
-Lemma loc_acceptable_noteq_diff:
- forall l1 l2,
- loc_acceptable l1 -> l1 <> l2 -> Loc.diff l1 l2.
-Proof.
- unfold loc_acceptable, Loc.diff; destruct l1; destruct l2;
- try (destruct s); try (destruct s0); intros; auto; try congruence.
- case (zeq z z0); intro.
- compare t t0; intro.
- subst z0; subst t0; tauto.
- tauto. tauto.
- contradiction. contradiction.
-Qed.
-
-Lemma loc_acceptable_notin_notin:
- forall r ll,
- loc_acceptable r ->
- ~(In r ll) -> Loc.notin r ll.
-Proof.
- induction ll; simpl; intros.
- auto.
- split. apply loc_acceptable_noteq_diff. assumption.
- apply sym_not_equal. tauto.
- apply IHll. assumption. tauto.
-Qed.
-
-(** * Additional properties of result and argument locations *)
-
-(** The result location is not a callee-save register. *)
-
-Lemma loc_result_not_callee_save:
- forall (s: signature),
- ~(In (loc_result s) int_callee_save_regs \/ In (loc_result s) float_callee_save_regs).
-Proof.
- intros. generalize (loc_result_caller_save s).
- generalize (int_callee_save_not_destroyed (loc_result s)).
- generalize (float_callee_save_not_destroyed (loc_result s)).
- tauto.
-Qed.
-
-(** Callee-save registers do not overlap with argument locations. *)
-
-Lemma arguments_not_preserved:
- forall sig l,
- Loc.notin l destroyed_at_call -> loc_acceptable l ->
- Loc.notin l (loc_arguments sig).
-Proof.
- intros. destruct l; red in H0.
- apply Loc.reg_notin. red; intros.
- exploit Loc.notin_not_in; eauto. eapply arguments_caller_save; eauto.
- destruct s; try contradiction.
- unfold loc_arguments. apply loc_arguments_rec_notin_local.
-Qed.
-
-(** There is no partial overlap between arguments and acceptable locations. *)
-
-Lemma no_overlap_arguments:
- forall args sg,
- locs_acceptable args ->
- Loc.no_overlap args (loc_arguments sg).
-Proof.
- unfold Loc.no_overlap; intros.
- generalize (H r H0).
- generalize (loc_arguments_acceptable _ _ H1).
- destruct s; destruct r; simpl.
- intros. case (mreg_eq m0 m); intro. left; congruence. tauto.
- intros. right; destruct s; auto.
- intros. right. auto.
- destruct s; try tauto. destruct s0; tauto.
-Qed.
-
(** ** Location of function parameters *)
(** A function finds the values of its parameter in the same locations
@@ -135,86 +31,39 @@ Qed.
Definition parameter_of_argument (l: loc) : loc :=
match l with
- | S (Outgoing n ty) => S (Incoming n ty)
+ | S Outgoing n ty => S Incoming n ty
| _ => l
end.
Definition loc_parameters (s: signature) :=
List.map parameter_of_argument (loc_arguments s).
-Lemma loc_parameters_type:
- forall sig, List.map Loc.type (loc_parameters sig) = sig.(sig_args).
-Proof.
- intros. unfold loc_parameters.
- rewrite list_map_compose.
- rewrite <- loc_arguments_type.
- apply list_map_exten.
- intros. destruct x; simpl. auto.
- destruct s; reflexivity.
-Qed.
-
-Lemma loc_parameters_length:
- forall sg, List.length (loc_parameters sg) = List.length sg.(sig_args).
-Proof.
- intros. unfold loc_parameters. rewrite list_length_map.
- apply loc_arguments_length.
-Qed.
-
-Lemma loc_parameters_not_temporaries:
- forall sig, Loc.disjoint (loc_parameters sig) temporaries.
-Proof.
- intro; red; intros.
- unfold loc_parameters in H.
- elim (list_in_map_inv _ _ _ H). intros y [EQ IN].
- generalize (loc_arguments_not_temporaries sig y x2 IN H0).
- subst x1. destruct x2.
- destruct y; simpl. auto. destruct s; auto.
- byContradiction. generalize H0. simpl. NotOrEq.
-Qed.
-
-Lemma no_overlap_parameters:
- forall params sg,
- locs_acceptable params ->
- Loc.no_overlap (loc_parameters sg) params.
-Proof.
- unfold Loc.no_overlap; intros.
- unfold loc_parameters in H0.
- elim (list_in_map_inv _ _ _ H0). intros t [EQ IN].
- rewrite EQ.
- generalize (loc_arguments_acceptable _ _ IN).
- generalize (H s H1).
- destruct s; destruct t; simpl.
- intros. case (mreg_eq m0 m); intro. left; congruence. tauto.
- intros. right; destruct s; simpl; auto.
- intros; right; auto.
- destruct s; try tauto. destruct s0; try tauto.
- intros; simpl. tauto.
-Qed.
-
Lemma incoming_slot_in_parameters:
forall ofs ty sg,
- In (S (Incoming ofs ty)) (loc_parameters sg) ->
- In (S (Outgoing ofs ty)) (loc_arguments sg).
+ In (S Incoming ofs ty) (loc_parameters sg) ->
+ In (S Outgoing ofs ty) (loc_arguments sg).
Proof.
intros.
unfold loc_parameters in H.
- change (S (Incoming ofs ty)) with (parameter_of_argument (S (Outgoing ofs ty))) in H.
+ change (S Incoming ofs ty) with (parameter_of_argument (S Outgoing ofs ty)) in H.
exploit list_in_map_inv. eexact H. intros [x [A B]]. simpl in A.
exploit loc_arguments_acceptable; eauto. unfold loc_argument_acceptable; intros.
destruct x; simpl in A; try discriminate.
- destruct s; try contradiction.
+ destruct sl; try contradiction.
inv A. auto.
Qed.
-
(** * Tail calls *)
(** A tail-call is possible for a signature if the corresponding
arguments are all passed in registers. *)
+(** A tail-call is possible for a signature if the corresponding
+ arguments are all passed in registers. *)
+
Definition tailcall_possible (s: signature) : Prop :=
forall l, In l (loc_arguments s) ->
- match l with R _ => True | S _ => False end.
+ match l with R _ => True | S _ _ _ => False end.
(** Decide whether a tailcall is possible. *)
@@ -223,7 +72,7 @@ Definition tailcall_is_possible (sg: signature) : bool :=
match l with
| nil => true
| R _ :: l' => tcisp l'
- | S _ :: l' => false
+ | S _ _ _ :: l' => false
end
in tcisp (loc_arguments sg).
@@ -237,31 +86,12 @@ Proof.
destruct H0. subst l0. auto. apply IHl. auto. auto. discriminate.
Qed.
-(** * Counting temporaries *)
-
-(** Given a list [tys] of types representing arguments to an operator,
- [arity_ok tys] returns [true] if there are enough temporaries to
- reload all arguments into temporaries. *)
-
-Fixpoint arity_ok_rec (tys: list typ) (itmps ftmps: list mreg)
- {struct tys} : bool :=
- match tys with
- | nil => true
- | Tint :: ts =>
- match itmps with
- | nil => false
- | it1 :: its => arity_ok_rec ts its ftmps
- end
- | Tfloat :: ts =>
- match ftmps with
- | nil => false
- | ft1 :: fts => arity_ok_rec ts itmps fts
- end
- end.
-
-Definition arity_ok (tys: list typ) :=
- arity_ok_rec tys int_temporaries float_temporaries.
-
-
-
-
+Lemma zero_size_arguments_tailcall_possible:
+ forall sg, size_arguments sg = 0 -> tailcall_possible sg.
+Proof.
+ intros; red; intros. exploit loc_arguments_acceptable; eauto.
+ unfold loc_argument_acceptable.
+ destruct l; intros. auto. destruct sl; try contradiction. destruct H1.
+ generalize (loc_arguments_bounded _ _ _ H0).
+ generalize (typesize_pos ty). omega.
+Qed.
diff --git a/backend/IRC.ml b/backend/IRC.ml
new file mode 100644
index 0000000..573c3d7
--- /dev/null
+++ b/backend/IRC.ml
@@ -0,0 +1,894 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+open Printf
+open Camlcoq
+open Datatypes
+open AST
+open Registers
+open Machregs
+open Locations
+open Conventions1
+open Conventions
+open XTL
+
+(* Iterated Register Coalescing: George and Appel's graph coloring algorithm *)
+
+type var_stats = {
+ mutable cost: int; (* estimated cost of a spill *)
+ mutable usedefs: int (* number of uses and defs *)
+}
+
+(* Representation of the interference graph. Each node of the graph
+ (i.e. each variable) is represented as follows. *)
+
+type node =
+ { ident: int; (*r unique identifier *)
+ typ: typ; (*r its type *)
+ var: var; (*r the XTL variable it comes from *)
+ regclass: int; (*r identifier of register class *)
+ mutable accesses: int; (*r number of defs and uses *)
+ mutable spillcost: float; (*r estimated cost of spilling *)
+ mutable adjlist: node list; (*r all nodes it interferes with *)
+ mutable degree: int; (*r number of adjacent nodes *)
+ mutable movelist: move list; (*r list of moves it is involved in *)
+ mutable extra_adj: node list; (*r extra interferences (see below) *)
+ mutable extra_pref: move list; (*r extra preferences (see below) *)
+ mutable alias: node option; (*r [Some n] if coalesced with [n] *)
+ mutable color: loc option; (*r chosen color *)
+ mutable nstate: nodestate; (*r in which set of nodes it is *)
+ mutable nprev: node; (*r for double linking *)
+ mutable nnext: node (*r for double linking *)
+ }
+
+(* These are the possible states for nodes. *)
+
+and nodestate =
+ | Colored
+ | Initial
+ | SimplifyWorklist
+ | FreezeWorklist
+ | SpillWorklist
+ | CoalescedNodes
+ | SelectStack
+
+(* Each move (i.e. wish to be put in the same location) is represented
+ as follows. *)
+
+and move =
+ { src: node; (*r source of the move *)
+ dst: node; (*r destination of the move *)
+ mutable mstate: movestate; (*r in which set of moves it is *)
+ mutable mprev: move; (*r for double linking *)
+ mutable mnext: move (*r for double linking *)
+ }
+
+(* These are the possible states for moves *)
+
+and movestate =
+ | CoalescedMoves
+ | ConstrainedMoves
+ | FrozenMoves
+ | WorklistMoves
+ | ActiveMoves
+
+(* Note on "precolored" nodes and how they are handled:
+
+The register allocator can express interferences and preferences between
+any two values of type [var]: either pseudoregisters, to be colored by IRC,
+or fixed, "precolored" locations.
+
+I and P between two pseudoregisters are recorded in the graph that IRC
+modifies, via the [adjlist] and [movelist] fields.
+
+I and P between a pseudoregister and a machine register are also
+recorded in the IRC graph, but only in the [adjlist] and [movelist]
+fields of the pseudoregister. This is the special case described
+in George and Appel's papers.
+
+I and P between a pseudoregister and a stack slot
+are omitted from the IRC graph, as they contribute nothing to the
+simplification and coalescing process. We record them in the
+[extra_adj] and [extra_pref] fields, where they can be honored
+after IRC elimination, when assigning a stack slot to a spilled variable. *)
+
+let name_of_loc = function
+ | R r ->
+ begin match Machregsaux.name_of_register r with
+ | None -> "fixed-reg"
+ | Some s -> s
+ end
+ | S (Local, ofs, ty) ->
+ sprintf "L%c%ld" (PrintXTL.short_name_of_type ty) (camlint_of_coqint ofs)
+ | S (Incoming, ofs, ty) ->
+ sprintf "I%c%ld" (PrintXTL.short_name_of_type ty) (camlint_of_coqint ofs)
+ | S (Outgoing, ofs, ty) ->
+ sprintf "O%c%ld" (PrintXTL.short_name_of_type ty) (camlint_of_coqint ofs)
+
+let name_of_node n =
+ match n.var with
+ | V(r, ty) -> sprintf "x%ld" (P.to_int32 r)
+ | L l -> name_of_loc l
+
+(* The algorithm manipulates partitions of the nodes and of the moves
+ according to their states, frequently moving a node or a move from
+ a state to another, and frequently enumerating all nodes or all moves
+ of a given state. To support these operations efficiently,
+ nodes or moves having the same state are put into imperative doubly-linked
+ lists, allowing for constant-time insertion and removal, and linear-time
+ scanning. We now define the operations over these doubly-linked lists. *)
+
+module DLinkNode = struct
+ type t = node
+ let make state =
+ let rec empty =
+ { ident = 0; typ = Tint; var = V(P.one, Tint); regclass = 0;
+ adjlist = []; degree = 0; accesses = 0; spillcost = 0.0;
+ movelist = []; extra_adj = []; extra_pref = [];
+ alias = None; color = None;
+ nstate = state; nprev = empty; nnext = empty }
+ in empty
+ let dummy = make Colored
+ let clear dl = dl.nnext <- dl; dl.nprev <- dl
+ let notempty dl = dl.nnext != dl
+ let insert n dl =
+ n.nstate <- dl.nstate;
+ n.nnext <- dl.nnext; n.nprev <- dl;
+ dl.nnext.nprev <- n; dl.nnext <- n
+ let remove n dl =
+ assert (n.nstate = dl.nstate);
+ n.nnext.nprev <- n.nprev; n.nprev.nnext <- n.nnext
+ let move n dl1 dl2 =
+ remove n dl1; insert n dl2
+ let pick dl =
+ let n = dl.nnext in remove n dl; n
+ let iter f dl =
+ let rec iter n = if n != dl then (f n; iter n.nnext)
+ in iter dl.nnext
+ let fold f dl accu =
+ let rec fold n accu = if n == dl then accu else fold n.nnext (f n accu)
+ in fold dl.nnext accu
+end
+
+module DLinkMove = struct
+ type t = move
+ let make state =
+ let rec empty =
+ { src = DLinkNode.dummy; dst = DLinkNode.dummy;
+ mstate = state; mprev = empty; mnext = empty }
+ in empty
+ let dummy = make CoalescedMoves
+ let clear dl = dl.mnext <- dl; dl.mprev <- dl
+ let notempty dl = dl.mnext != dl
+ let insert m dl =
+ m.mstate <- dl.mstate;
+ m.mnext <- dl.mnext; m.mprev <- dl;
+ dl.mnext.mprev <- m; dl.mnext <- m
+ let remove m dl =
+ assert (m.mstate = dl.mstate);
+ m.mnext.mprev <- m.mprev; m.mprev.mnext <- m.mnext
+ let move m dl1 dl2 =
+ remove m dl1; insert m dl2
+ let pick dl =
+ let m = dl.mnext in remove m dl; m
+ let iter f dl =
+ let rec iter m = if m != dl then (f m; iter m.mnext)
+ in iter dl.mnext
+ let fold f dl accu =
+ let rec fold m accu = if m == dl then accu else fold m.mnext (f m accu)
+ in fold dl.mnext accu
+end
+
+(* Auxiliary data structures *)
+
+module IntSet = Set.Make(struct
+ type t = int
+ let compare (x:int) (y:int) = compare x y
+end)
+
+module IntPairSet = Set.Make(struct
+ type t = int * int
+ let compare ((x1, y1): (int * int)) (x2, y2) =
+ if x1 < x2 then -1 else
+ if x1 > x2 then 1 else
+ if y1 < y2 then -1 else
+ if y1 > y2 then 1 else
+ 0
+ end)
+
+(* The global state of the algorithm *)
+
+type graph = {
+ (* Machine registers available for allocation *)
+ caller_save_registers: mreg array array;
+ callee_save_registers: mreg array array;
+ num_available_registers: int array;
+ start_points: int array;
+ allocatable_registers: mreg list;
+ (* Costs for pseudo-registers *)
+ stats_of_reg: reg -> var_stats;
+ (* Mapping from XTL variables to nodes *)
+ varTable: (var, node) Hashtbl.t;
+ mutable nextIdent: int;
+ (* The adjacency set *)
+ mutable adjSet: IntPairSet.t;
+ (* Low-degree, non-move-related nodes *)
+ simplifyWorklist: DLinkNode.t;
+ (* Low-degree, move-related nodes *)
+ freezeWorklist: DLinkNode.t;
+ (* High-degree nodes *)
+ spillWorklist: DLinkNode.t;
+ (* Nodes that have been coalesced *)
+ coalescedNodes: DLinkNode.t;
+ (* Moves that have been coalesced *)
+ coalescedMoves: DLinkMove.t;
+ (* Moves whose source and destination interfere *)
+ constrainedMoves: DLinkMove.t;
+ (* Moves that will no longer be considered for coalescing *)
+ frozenMoves: DLinkMove.t;
+ (* Moves enabled for possible coalescing *)
+ worklistMoves: DLinkMove.t;
+ (* Moves not yet ready for coalescing *)
+ activeMoves: DLinkMove.t
+}
+
+(* Register classes and reserved registers *)
+
+let num_register_classes = 2
+
+let class_of_type = function Tint -> 0 | Tfloat -> 1 | Tlong -> assert false
+
+let reserved_registers = ref ([]: mreg list)
+
+let rec remove_reserved = function
+ | [] -> []
+ | hd :: tl ->
+ if List.mem hd !reserved_registers
+ then remove_reserved tl
+ else hd :: remove_reserved tl
+
+(* Initialize and return an empty graph *)
+
+let init costs =
+ let int_caller_save = remove_reserved int_caller_save_regs
+ and float_caller_save = remove_reserved float_caller_save_regs
+ and int_callee_save = remove_reserved int_callee_save_regs
+ and float_callee_save = remove_reserved float_callee_save_regs in
+ {
+ caller_save_registers =
+ [| Array.of_list int_caller_save; Array.of_list float_caller_save |];
+ callee_save_registers =
+ [| Array.of_list int_callee_save; Array.of_list float_callee_save |];
+ num_available_registers =
+ [| List.length int_caller_save + List.length int_callee_save;
+ List.length float_caller_save + List.length float_callee_save |];
+ start_points =
+ [| 0; 0 |];
+ allocatable_registers =
+ int_caller_save @ int_callee_save @ float_caller_save @ float_callee_save;
+ stats_of_reg = costs;
+ varTable = Hashtbl.create 253;
+ nextIdent = 0;
+ adjSet = IntPairSet.empty;
+ simplifyWorklist = DLinkNode.make SimplifyWorklist;
+ freezeWorklist = DLinkNode.make FreezeWorklist;
+ spillWorklist = DLinkNode.make SpillWorklist;
+ coalescedNodes = DLinkNode.make CoalescedNodes;
+ coalescedMoves = DLinkMove.make CoalescedMoves;
+ constrainedMoves = DLinkMove.make ConstrainedMoves;
+ frozenMoves = DLinkMove.make FrozenMoves;
+ worklistMoves = DLinkMove.make WorklistMoves;
+ activeMoves = DLinkMove.make ActiveMoves
+ }
+
+(* Create nodes corresponding to XTL variables *)
+
+let weightedSpillCost st =
+ if st.cost < max_int
+ then float_of_int st.cost
+ else infinity
+
+let newNodeOfReg g r ty =
+ let st = g.stats_of_reg r in
+ g.nextIdent <- g.nextIdent + 1;
+ { ident = g.nextIdent; typ = ty;
+ var = V(r, ty); regclass = class_of_type ty;
+ accesses = st.usedefs;
+ spillcost = weightedSpillCost st;
+ adjlist = []; degree = 0; movelist = []; extra_adj = []; extra_pref = [];
+ alias = None;
+ color = None;
+ nstate = Initial;
+ nprev = DLinkNode.dummy; nnext = DLinkNode.dummy }
+
+let newNodeOfLoc g l =
+ let ty = Loc.coq_type l in
+ g.nextIdent <- g.nextIdent + 1;
+ { ident = g.nextIdent; typ = ty;
+ var = L l; regclass = class_of_type ty;
+ accesses = 0; spillcost = 0.0;
+ adjlist = []; degree = 0; movelist = []; extra_adj = []; extra_pref = [];
+ alias = None;
+ color = Some l;
+ nstate = Colored;
+ nprev = DLinkNode.dummy; nnext = DLinkNode.dummy }
+
+let nodeOfVar g v =
+ try
+ Hashtbl.find g.varTable v
+ with Not_found ->
+ let n =
+ match v with V(r, ty) -> newNodeOfReg g r ty | L l -> newNodeOfLoc g l in
+ Hashtbl.add g.varTable v n;
+ n
+
+(* Determine if two nodes interfere *)
+
+let interfere g n1 n2 =
+ let i1 = n1.ident and i2 = n2.ident in
+ let p = if i1 < i2 then (i1, i2) else (i2, i1) in
+ IntPairSet.mem p g.adjSet
+
+(* Add an edge to the graph. *)
+
+let recordInterf n1 n2 =
+ match n2.color with
+ | None | Some (R _) ->
+ if n1.regclass = n2.regclass then begin
+ n1.adjlist <- n2 :: n1.adjlist;
+ n1.degree <- 1 + n1.degree
+ end else begin
+ n1.extra_adj <- n2 :: n1.extra_adj
+ end
+ | Some (S _) ->
+ (*i printf "extra adj %s to %s\n" (name_of_node n1) (name_of_node n2); *)
+ n1.extra_adj <- n2 :: n1.extra_adj
+
+let addEdge g n1 n2 =
+ (*i printf "edge %s -- %s;\n" (name_of_node n1) (name_of_node n2);*)
+ assert (n1 != n2);
+ if not (interfere g n1 n2) then begin
+ let i1 = n1.ident and i2 = n2.ident in
+ let p = if i1 < i2 then (i1, i2) else (i2, i1) in
+ g.adjSet <- IntPairSet.add p g.adjSet;
+ if n1.nstate <> Colored then recordInterf n1 n2;
+ if n2.nstate <> Colored then recordInterf n2 n1
+ end
+
+(* Add a move preference. *)
+
+let recordMove g n1 n2 =
+ let m =
+ { src = n1; dst = n2; mstate = WorklistMoves;
+ mnext = DLinkMove.dummy; mprev = DLinkMove.dummy } in
+ n1.movelist <- m :: n1.movelist;
+ n2.movelist <- m :: n2.movelist;
+ DLinkMove.insert m g.worklistMoves
+
+let recordExtraPref n1 n2 =
+ let m =
+ { src = n1; dst = n2; mstate = FrozenMoves;
+ mnext = DLinkMove.dummy; mprev = DLinkMove.dummy } in
+ n1.extra_pref <- m :: n1.extra_pref
+
+let addMovePref g n1 n2 =
+ assert (n1.regclass = n2.regclass);
+ match n1.color, n2.color with
+ | None, None ->
+ recordMove g n1 n2
+ | Some (R mr1), None ->
+ if List.mem mr1 g.allocatable_registers then recordMove g n1 n2
+ | None, Some (R mr2) ->
+ if List.mem mr2 g.allocatable_registers then recordMove g n1 n2
+ | Some (S _), None ->
+ recordExtraPref n2 n1
+ | None, Some (S _) ->
+ recordExtraPref n1 n2
+ | _, _ ->
+ ()
+
+(* Apply the given function to the relevant adjacent nodes of a node *)
+
+let iterAdjacent f n =
+ List.iter
+ (fun n ->
+ match n.nstate with
+ | SelectStack | CoalescedNodes -> ()
+ | _ -> f n)
+ n.adjlist
+
+(* Determine the moves affecting a node *)
+
+let moveIsActiveOrWorklist m =
+ match m.mstate with
+ | ActiveMoves | WorklistMoves -> true
+ | _ -> false
+
+let nodeMoves n =
+ List.filter moveIsActiveOrWorklist n.movelist
+
+(* Determine whether a node is involved in a move *)
+
+let moveRelated n =
+ List.exists moveIsActiveOrWorklist n.movelist
+
+(* Initial partition of nodes into spill / freeze / simplify *)
+
+let initialNodePartition g =
+ let part_node v n =
+ match n.nstate with
+ | Initial ->
+ let k = g.num_available_registers.(n.regclass) in
+ if n.degree >= k then
+ DLinkNode.insert n g.spillWorklist
+ else if moveRelated n then
+ DLinkNode.insert n g.freezeWorklist
+ else
+ DLinkNode.insert n g.simplifyWorklist
+ | Colored -> ()
+ | _ -> assert false in
+ Hashtbl.iter part_node g.varTable
+
+
+(* Check invariants *)
+
+let degreeInvariant g n =
+ let c = ref 0 in
+ iterAdjacent (fun n -> incr c) n;
+ if !c <> n.degree then
+ failwith("degree invariant violated by " ^ name_of_node n)
+
+let simplifyWorklistInvariant g n =
+ if n.degree < g.num_available_registers.(n.regclass)
+ && not (moveRelated n)
+ then ()
+ else failwith("simplify worklist invariant violated by " ^ name_of_node n)
+
+let freezeWorklistInvariant g n =
+ if n.degree < g.num_available_registers.(n.regclass)
+ && moveRelated n
+ then ()
+ else failwith("freeze worklist invariant violated by " ^ name_of_node n)
+
+let spillWorklistInvariant g n =
+ if n.degree >= g.num_available_registers.(n.regclass)
+ then ()
+ else failwith("spill worklist invariant violated by " ^ name_of_node n)
+
+let checkInvariants g =
+ DLinkNode.iter
+ (fun n -> degreeInvariant g n; simplifyWorklistInvariant g n)
+ g.simplifyWorklist;
+ DLinkNode.iter
+ (fun n -> degreeInvariant g n; freezeWorklistInvariant g n)
+ g.freezeWorklist;
+ DLinkNode.iter
+ (fun n -> degreeInvariant g n; spillWorklistInvariant g n)
+ g.spillWorklist
+
+(* Enable moves that have become low-degree related *)
+
+let enableMoves g n =
+ List.iter
+ (fun m ->
+ if m.mstate = ActiveMoves
+ then DLinkMove.move m g.activeMoves g.worklistMoves)
+ (nodeMoves n)
+
+(* Simulate the removal of a node from the graph *)
+
+let decrementDegree g n =
+ let k = g.num_available_registers.(n.regclass) in
+ let d = n.degree in
+ n.degree <- d - 1;
+ if d = k then begin
+ enableMoves g n;
+ iterAdjacent (enableMoves g) n;
+ if moveRelated n
+ then DLinkNode.move n g.spillWorklist g.freezeWorklist
+ else DLinkNode.move n g.spillWorklist g.simplifyWorklist
+ end
+
+(* Simulate the effect of combining nodes [n1] and [n3] on [n2],
+ where [n2] is a node adjacent to [n3]. *)
+
+let combineEdge g n1 n2 =
+ assert (n1 != n2);
+ if interfere g n1 n2 then begin
+ (* The two edges n2--n3 and n2--n1 become one, so degree of n2 decreases *)
+ decrementDegree g n2
+ end else begin
+ (* Add new edge *)
+ let i1 = n1.ident and i2 = n2.ident in
+ let p = if i1 < i2 then (i1, i2) else (i2, i1) in
+ g.adjSet <- IntPairSet.add p g.adjSet;
+ if n1.nstate <> Colored then begin
+ n1.adjlist <- n2 :: n1.adjlist;
+ n1.degree <- 1 + n1.degree
+ end;
+ if n2.nstate <> Colored then begin
+ n2.adjlist <- n1 :: n2.adjlist;
+ (* n2's degree stays the same because the old edge n2--n3 disappears
+ and becomes the new edge n2--n1 *)
+ end
+ end
+
+(* Simplification of a low-degree node *)
+
+let simplify g =
+ let n = DLinkNode.pick g.simplifyWorklist in
+ (*i printf "Simplifying %s\n" (name_of_node n); *)
+ n.nstate <- SelectStack;
+ iterAdjacent (decrementDegree g) n;
+ n
+
+(* Briggs's conservative coalescing criterion. In the terminology of
+ Hailperin, "Comparing Conservative Coalescing Criteria",
+ TOPLAS 27(3) 2005, this is the full Briggs criterion, slightly
+ more powerful than the one in George and Appel's paper. *)
+
+let canCoalesceBriggs g u v =
+ let seen = ref IntSet.empty in
+ let k = g.num_available_registers.(u.regclass) in
+ let c = ref 0 in
+ let consider other n =
+ if not (IntSet.mem n.ident !seen) then begin
+ seen := IntSet.add n.ident !seen;
+ (* if n interferes with both u and v, its degree will decrease by one
+ after coalescing *)
+ let degree_after_coalescing =
+ if interfere g n other then n.degree - 1 else n.degree in
+ if degree_after_coalescing >= k || n.nstate = Colored then begin
+ incr c;
+ if !c >= k then raise Exit
+ end
+ end in
+ try
+ iterAdjacent (consider v) u;
+ iterAdjacent (consider u) v;
+ (*i printf " Briggs: OK for %s and %s\n" (name_of_node u) (name_of_node v); *)
+ true
+ with Exit ->
+ (*i printf " Briggs: no\n"; *)
+ false
+
+(* George's conservative coalescing criterion: all high-degree neighbors
+ of [v] are neighbors of [u]. *)
+
+let canCoalesceGeorge g u v =
+ let k = g.num_available_registers.(u.regclass) in
+ let isOK t =
+ if t.nstate = Colored then
+ if u.nstate = Colored || interfere g t u then () else raise Exit
+ else
+ if t.degree < k || interfere g t u then () else raise Exit
+ in
+ try
+ iterAdjacent isOK v;
+ (*i printf " George: OK for %s and %s\n" (name_of_node u) (name_of_node v); *)
+ true
+ with Exit ->
+ (*i printf " George: no\n"; *)
+ false
+
+(* The combined coalescing criterion. [u] can be precolored, but
+ [v] is not. According to George and Appel's paper:
+- If [u] is precolored, use George's criterion.
+- If [u] is not precolored, use Briggs's criterion.
+
+ As noted by Hailperin, for non-precolored nodes, George's criterion
+ is incomparable with Briggs's: there are cases where G says yes
+ and B says no. Typically, [u] is a long-lived variable with many
+ interferences, and [v] is a short-lived temporary copy of [u]
+ that has no more interferences than [u]. Coalescing [u] and [v]
+ is "weakly safe" in Hailperin's terminology: [u] is no harder to color,
+ [u]'s neighbors are no harder to color either, but if we end up
+ spilling [u], we'll spill [v] as well. So, we restrict this heuristic
+ to [v] having a small number of uses.
+*)
+
+let thresholdGeorge = 3
+
+let canCoalesce g u v =
+ (*i printf "canCoalesce %s[%.2f] %s[%.2f]\n"
+ (name_of_node u) u.spillcost (name_of_node v) v.spillcost; *)
+ if u.nstate = Colored
+ then canCoalesceGeorge g u v
+ else canCoalesceBriggs g u v
+ || (u.spillcost < infinity && v.spillcost < infinity &&
+ ((v.accesses <= thresholdGeorge && canCoalesceGeorge g u v)
+ || (u.accesses <= thresholdGeorge && canCoalesceGeorge g v u)))
+
+(* Update worklists after a move was processed *)
+
+let addWorkList g u =
+ if (not (u.nstate = Colored))
+ && u.degree < g.num_available_registers.(u.regclass)
+ && (not (moveRelated u))
+ then DLinkNode.move u g.freezeWorklist g.simplifyWorklist
+
+(* Return the canonical representative of a possibly coalesced node *)
+
+let rec getAlias n =
+ match n.alias with None -> n | Some n' -> getAlias n'
+
+(* Combine two nodes *)
+
+let combine g u v =
+ (*i printf "Combining %s and %s\n" (name_of_node u) (name_of_node v); *)
+ (*i if u.spillcost = infinity then
+ printf "Warning: combining unspillable %s\n" (name_of_node u);
+ if v.spillcost = infinity then
+ printf "Warning: combining unspillable %s\n" (name_of_node v);*)
+ if v.nstate = FreezeWorklist
+ then DLinkNode.move v g.freezeWorklist g.coalescedNodes
+ else DLinkNode.move v g.spillWorklist g.coalescedNodes;
+ v.alias <- Some u;
+ (* Precolored nodes often have big movelists, and if one of [u] and [v]
+ is precolored, it is []u. So, append [v.movelist] to [u.movelist]
+ instead of the other way around. *)
+ u.movelist <- List.rev_append v.movelist u.movelist;
+ u.spillcost <- u.spillcost +. v.spillcost;
+ iterAdjacent (combineEdge g u) v; (*r original code using [decrementDegree] is buggy *)
+ u.extra_adj <- u.extra_adj @ v.extra_adj;
+ u.extra_pref <- u.extra_pref @ v.extra_pref;
+ enableMoves g v; (*r added as per Appel's book erratum *)
+ if u.degree >= g.num_available_registers.(u.regclass)
+ && u.nstate = FreezeWorklist
+ then DLinkNode.move u g.freezeWorklist g.spillWorklist
+
+(* Attempt coalescing *)
+
+let coalesce g =
+ let m = DLinkMove.pick g.worklistMoves in
+ let x = getAlias m.src and y = getAlias m.dst in
+ let (u, v) = if y.nstate = Colored then (y, x) else (x, y) in
+ (*i printf "Attempt coalescing %s and %s\n" (name_of_node u) (name_of_node v);*)
+ if u == v then begin
+ DLinkMove.insert m g.coalescedMoves;
+ addWorkList g u
+ end else if v.nstate = Colored || interfere g u v then begin
+ DLinkMove.insert m g.constrainedMoves;
+ addWorkList g u;
+ addWorkList g v
+ end else if canCoalesce g u v then begin
+ DLinkMove.insert m g.coalescedMoves;
+ combine g u v;
+ addWorkList g u
+ end else begin
+ DLinkMove.insert m g.activeMoves
+ end
+
+(* Freeze moves associated with node [u] *)
+
+let freezeMoves g u =
+ let u' = getAlias u in
+ let freeze m =
+ let y = getAlias m.src in
+ let v = if y == u' then getAlias m.dst else y in
+ DLinkMove.move m g.activeMoves g.frozenMoves;
+ if not (moveRelated v)
+ && v.degree < g.num_available_registers.(v.regclass)
+ && v.nstate <> Colored
+ then DLinkNode.move v g.freezeWorklist g.simplifyWorklist in
+ List.iter freeze (nodeMoves u)
+
+(* Pick a move and freeze it *)
+
+let freeze g =
+ let u = DLinkNode.pick g.freezeWorklist in
+ (*i printf "Freezing %s\n" (name_of_node u); *)
+ DLinkNode.insert u g.simplifyWorklist;
+ freezeMoves g u
+
+(* This is the original spill cost function from Chaitin 1982 *)
+
+(*
+let spillCost n =
+(*i
+ printf "spillCost %s: cost = %.2f degree = %d rank = %.2f\n"
+ (name_of_node n) n.spillcost n.degree
+ (n.spillcost /. float n.degree);
+*)
+ n.spillcost /. float n.degree
+*)
+
+(* This is spill cost function h_0 from Bernstein et al 1989. It performs
+ slightly better than Chaitin's and than functions h_1 and h_2. *)
+
+let spillCost n =
+ let deg = float n.degree in n.spillcost /. (deg *. deg)
+
+(* Spill a node *)
+
+let selectSpill g =
+ (*i printf "Attempt spilling\n"; *)
+ (* Find a spillable node of minimal cost *)
+ let (n, cost) =
+ DLinkNode.fold
+ (fun n (best_node, best_cost as best) ->
+ let cost = spillCost n in
+ if cost <= best_cost then (n, cost) else best)
+ g.spillWorklist (DLinkNode.dummy, infinity) in
+ assert (n != DLinkNode.dummy);
+ if cost = infinity then begin
+ printf "Warning: spilling unspillable %s\n" (name_of_node n);
+ printf " spill queue is:";
+ DLinkNode.iter (fun n -> printf " %s" (name_of_node n)) g.spillWorklist;
+ printf "\n"
+ end;
+ DLinkNode.remove n g.spillWorklist;
+ (*i printf "Spilling %s\n" (name_of_node n); *)
+ freezeMoves g n;
+ n.nstate <- SelectStack;
+ iterAdjacent (decrementDegree g) n;
+ n
+
+(* Produce the order of nodes that we'll use for coloring *)
+
+let rec nodeOrder g stack =
+ (*i checkInvariants g; *)
+ if DLinkNode.notempty g.simplifyWorklist then
+ (let n = simplify g in nodeOrder g (n :: stack))
+ else if DLinkMove.notempty g.worklistMoves then
+ (coalesce g; nodeOrder g stack)
+ else if DLinkNode.notempty g.freezeWorklist then
+ (freeze g; nodeOrder g stack)
+ else if DLinkNode.notempty g.spillWorklist then
+ (let n = selectSpill g in nodeOrder g (n :: stack))
+ else
+ stack
+
+(* Assign a color (i.e. a hardware register or a stack location)
+ to a node. The color is chosen among the colors that are not
+ assigned to nodes with which this node interferes. The choice
+ is guided by the following heuristics: consider first caller-save
+ hardware register of the correct type; second, callee-save registers;
+ third, a stack location. Callee-save registers and stack locations
+ are ``expensive'' resources, so we try to minimize their number
+ by picking the smallest available callee-save register or stack location.
+ In contrast, caller-save registers are ``free'', so we pick an
+ available one pseudo-randomly. *)
+
+module Regset =
+ Set.Make(struct type t = mreg let compare = compare end)
+
+let find_reg g conflicts regclass =
+ let rec find avail curr last =
+ if curr >= last then None else begin
+ let r = avail.(curr) in
+ if Regset.mem r conflicts
+ then find avail (curr + 1) last
+ else Some (R r)
+ end in
+ let caller_save = g.caller_save_registers.(regclass)
+ and callee_save = g.callee_save_registers.(regclass)
+ and start = g.start_points.(regclass) in
+ match find caller_save start (Array.length caller_save) with
+ | Some _ as res ->
+ g.start_points.(regclass) <-
+ (if start + 1 < Array.length caller_save then start + 1 else 0);
+ res
+ | None ->
+ match find caller_save 0 start with
+ | Some _ as res ->
+ g.start_points.(regclass) <-
+ (if start + 1 < Array.length caller_save then start + 1 else 0);
+ res
+ | None ->
+ find callee_save 0 (Array.length callee_save)
+
+(* Aggressive coalescing of stack slots. When assigning a slot,
+ try first the slots assigned to the pseudoregs for which we
+ have a preference, provided no conflict occurs. *)
+
+let rec reuse_slot conflicts n mvlist =
+ match mvlist with
+ | [] -> None
+ | mv :: rem ->
+ let attempt_reuse n' =
+ match n'.color with
+ | Some(S(Local, _, _) as l)
+ when List.for_all (Loc.diff_dec l) conflicts -> Some l
+ | _ -> reuse_slot conflicts n rem in
+ let src = getAlias mv.src and dst = getAlias mv.dst in
+ if n == src then attempt_reuse dst
+ else if n == dst then attempt_reuse src
+ else reuse_slot conflicts n rem (* should not happen? *)
+
+(* If no reuse possible, assign lowest nonconflicting stack slot. *)
+
+let compare_slots s1 s2 =
+ match s1, s2 with
+ | S(_, ofs1, _), S(_, ofs2, _) -> Z.compare ofs1 ofs2
+ | _, _ -> assert false
+
+let find_slot conflicts typ =
+ let rec find curr = function
+ | [] ->
+ S(Local, curr, typ)
+ | S(Local, ofs, typ') :: l ->
+ if Z.le (Z.add curr (typesize typ)) ofs then
+ S(Local, curr, typ)
+ else begin
+ let ofs' = Z.add ofs (typesize typ') in
+ find (if Z.le ofs' curr then curr else ofs') l
+ end
+ | _ :: l ->
+ find curr l
+ in find Z.zero (List.stable_sort compare_slots conflicts)
+
+(* Record locations assigned to interfering nodes *)
+
+let record_reg_conflict cnf n =
+ match (getAlias n).color with
+ | Some (R r) -> Regset.add r cnf
+ | _ -> cnf
+
+let record_slot_conflict cnf n =
+ match (getAlias n).color with
+ | Some (S _ as l) -> l :: cnf
+ | _ -> cnf
+
+(* Assign a location, the best we can *)
+
+let assign_color g n =
+ let reg_conflicts =
+ List.fold_left record_reg_conflict Regset.empty n.adjlist in
+ (* First, try to assign a register *)
+ match find_reg g reg_conflicts n.regclass with
+ | Some loc ->
+ n.color <- Some loc
+ | None ->
+ (* Add extra conflicts for nonallocatable and preallocated stack slots *)
+ let slot_conflicts =
+ List.fold_left record_slot_conflict
+ (List.fold_left record_slot_conflict [] n.adjlist)
+ n.extra_adj in
+ (* Second, try to coalesce stack slots *)
+ match reuse_slot slot_conflicts n (n.extra_pref @ n.movelist) with
+ | Some loc ->
+ n.color <- Some loc
+ | None ->
+ (* Last, pick a Local stack slot *)
+ n.color <- Some (find_slot slot_conflicts n.typ)
+
+(* Extract the location of a variable *)
+
+let location_of_var g v =
+ match v with
+ | L l -> l
+ | V(r, ty) ->
+ try
+ let n = Hashtbl.find g.varTable v in
+ let n' = getAlias n in
+ match n'.color with
+ | None -> assert false
+ | Some l -> l
+ with Not_found ->
+ match ty with
+ | Tint -> R dummy_int_reg
+ | Tfloat -> R dummy_float_reg
+ | Tlong -> assert false
+
+(* The exported interface *)
+
+let add_interf g v1 v2 =
+ addEdge g (nodeOfVar g v1) (nodeOfVar g v2)
+
+let add_pref g v1 v2 =
+ addMovePref g (nodeOfVar g v1) (nodeOfVar g v2)
+
+let coloring g =
+ initialNodePartition g;
+ List.iter (assign_color g) (nodeOrder g []);
+ location_of_var g (* total function var -> location *)
diff --git a/backend/IRC.mli b/backend/IRC.mli
new file mode 100644
index 0000000..e81b6dc
--- /dev/null
+++ b/backend/IRC.mli
@@ -0,0 +1,44 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(* Iterated Register Coalescing: George and Appel's graph coloring algorithm *)
+
+open AST
+open Registers
+open Machregs
+open Locations
+open XTL
+
+(* The abstract type of interference and preference graphs. *)
+type graph
+
+(* Information associated to every variable. *)
+type var_stats = {
+ mutable cost: int; (* estimated cost of a spill *)
+ mutable usedefs: int (* number of uses and defs *)
+}
+
+(* Create an empty graph. The given function associates statistics to
+ every variable. *)
+val init: (reg -> var_stats) -> graph
+
+(* Add an interference between two variables. *)
+val add_interf: graph -> var -> var -> unit
+
+(* Add a preference between two variables. *)
+val add_pref: graph -> var -> var -> unit
+
+(* Color the graph. Return an assignment of locations to variables. *)
+val coloring: graph -> (var -> loc)
+
+(* Machine registers that are reserved and not available for allocation. *)
+val reserved_registers: mreg list ref
diff --git a/backend/Inlining.v b/backend/Inlining.v
index a9b1e7e..7b19f80 100644
--- a/backend/Inlining.v
+++ b/backend/Inlining.v
@@ -180,11 +180,17 @@ Next Obligation.
intros; constructor; simpl; xomega.
Qed.
-Fixpoint mlist_iter2 {A B: Type} (f: A -> B -> mon unit) (l: list (A*B)): mon unit :=
- match l with
- | nil => ret tt
- | (x,y) :: l' => do z <- f x y; mlist_iter2 f l'
- end.
+Program Definition ptree_mfold {A: Type} (f: positive -> A -> mon unit) (t: PTree.t A): mon unit :=
+ fun s =>
+ R tt
+ (PTree.fold (fun s1 k v => match f k v s1 return _ with R _ s2 _ => s2 end) t s)
+ _.
+Next Obligation.
+ apply PTree_Properties.fold_rec.
+ auto.
+ apply sincr_refl.
+ intros. destruct (f k v a). eapply sincr_trans; eauto.
+Qed.
(** ** Inlining contexts *)
@@ -422,7 +428,7 @@ Definition expand_instr (ctx: context) (pc: node) (i: instruction): mon unit :=
Definition expand_cfg_rec (ctx: context) (f: function): mon unit :=
do x <- request_stack (ctx.(dstk) + ctx.(mstk));
- mlist_iter2 (expand_instr ctx) (PTree.elements f.(fn_code)).
+ ptree_mfold (expand_instr ctx) f.(fn_code).
End EXPAND_CFG.
diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v
index 82ef9cf..e8dba67 100644
--- a/backend/Inliningspec.v
+++ b/backend/Inliningspec.v
@@ -180,6 +180,35 @@ Ltac monadInv H :=
((progress simpl in H) || unfold F in H); monadInv1 H
end.
+Fixpoint mlist_iter2 {A B: Type} (f: A -> B -> mon unit) (l: list (A*B)): mon unit :=
+ match l with
+ | nil => ret tt
+ | (x,y) :: l' => do z <- f x y; mlist_iter2 f l'
+ end.
+
+Remark mlist_iter2_fold:
+ forall (A B: Type) (f: A -> B -> mon unit) l s,
+ exists i,
+ mlist_iter2 f l s =
+ R tt (fold_left (fun a p => match f (fst p) (snd p) a with R _ s2 _ => s2 end) l s) i.
+Proof.
+ induction l; simpl; intros.
+ exists (sincr_refl s); auto.
+ destruct a as [x y]. unfold bind. simpl. destruct (f x y s) as [xx s1 i1].
+ destruct (IHl s1) as [i2 EQ]. rewrite EQ. econstructor; eauto.
+Qed.
+
+Lemma ptree_mfold_spec:
+ forall (A: Type) (f: positive -> A -> mon unit) t s x s' i,
+ ptree_mfold f t s = R x s' i ->
+ exists i', mlist_iter2 f (PTree.elements t) s = R tt s' i'.
+Proof.
+ intros.
+ destruct (mlist_iter2_fold _ _ f (PTree.elements t) s) as [i' EQ].
+ unfold ptree_mfold in H. inv H. rewrite PTree.fold_spec.
+ econstructor. eexact EQ.
+Qed.
+
(** ** Relational specification of the translation of moves *)
Inductive tr_moves (c: code) : node -> list reg -> list reg -> node -> Prop :=
@@ -416,6 +445,7 @@ Lemma expand_cfg_rec_unchanged:
Proof.
intros. unfold expand_cfg_rec in H. monadInv H. inversion EQ.
transitivity ((st_code s0)!pc).
+ exploit ptree_mfold_spec; eauto. intros [INCR' ITER].
eapply iter_expand_instr_unchanged; eauto.
subst s0; auto.
subst s0; simpl. xomega.
@@ -596,7 +626,8 @@ Proof.
intros. unfold expand_cfg_rec in H. monadInv H. inversion EQ.
constructor.
intros. rewrite H1. eapply max_def_function_params; eauto.
- intros. eapply iter_expand_instr_spec; eauto.
+ intros. exploit ptree_mfold_spec; eauto. intros [INCR' ITER].
+ eapply iter_expand_instr_spec; eauto.
apply PTree.elements_keys_norepet.
intros. rewrite H1. eapply max_def_function_instr; eauto.
eapply PTree.elements_complete; eauto.
diff --git a/backend/LTL.v b/backend/LTL.v
index 422b0e0..de10845 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -30,27 +30,33 @@ Require Import Conventions.
(** * Abstract syntax *)
-(** LTL is close to RTL, but uses locations instead of pseudo-registers. *)
+(** LTL is close to RTL, but uses machine registers and stack slots
+ instead of pseudo-registers. Also, the nodes of the control-flow
+ graph are basic blocks instead of single instructions. *)
Definition node := positive.
Inductive instruction: Type :=
- | Lnop: node -> instruction
- | Lop: operation -> list loc -> loc -> node -> instruction
- | Lload: memory_chunk -> addressing -> list loc -> loc -> node -> instruction
- | Lstore: memory_chunk -> addressing -> list loc -> loc -> node -> instruction
- | Lcall: signature -> loc + ident -> list loc -> loc -> node -> instruction
- | Ltailcall: signature -> loc + ident -> list loc -> instruction
- | Lbuiltin: external_function -> list loc -> loc -> node -> instruction
- | Lcond: condition -> list loc -> node -> node -> instruction
- | Ljumptable: loc -> list node -> instruction
- | Lreturn: option loc -> instruction.
-
-Definition code: Type := PTree.t instruction.
+ | Lop (op: operation) (args: list mreg) (res: mreg)
+ | Lload (chunk: memory_chunk) (addr: addressing) (args: list mreg) (dst: mreg)
+ | Lgetstack (sl: slot) (ofs: Z) (ty: typ) (dst: mreg)
+ | Lsetstack (src: mreg) (sl: slot) (ofs: Z) (ty: typ)
+ | Lstore (chunk: memory_chunk) (addr: addressing) (args: list mreg) (src: mreg)
+ | Lcall (sg: signature) (ros: mreg + ident)
+ | Ltailcall (sg: signature) (ros: mreg + ident)
+ | Lbuiltin (ef: external_function) (args: list mreg) (res: list mreg)
+ | Lannot (ef: external_function) (args: list loc)
+ | Lbranch (s: node)
+ | Lcond (cond: condition) (args: list mreg) (s1 s2: node)
+ | Ljumptable (arg: mreg) (tbl: list node)
+ | Lreturn.
+
+Definition bblock := list instruction.
+
+Definition code: Type := PTree.t bblock.
Record function: Type := mkfunction {
fn_sig: signature;
- fn_params: list loc;
fn_stacksize: Z;
fn_code: code;
fn_entrypoint: node
@@ -71,49 +77,56 @@ Definition funsig (fd: fundef) :=
Definition genv := Genv.t fundef unit.
Definition locset := Locmap.t.
-Definition locmap_optget (ol: option loc) (dfl: val) (ls: locset) : val :=
- match ol with
- | None => dfl
- | Some l => ls l
- end.
+(** Calling conventions are reflected at the level of location sets
+ (environments mapping locations to values) by the following two
+ functions.
-Fixpoint init_locs (vl: list val) (rl: list loc) {struct rl} : locset :=
- match rl, vl with
- | r1 :: rs, v1 :: vs => Locmap.set r1 v1 (init_locs vs rs)
- | _, _ => Locmap.init Vundef
- end.
+ [call_regs caller] returns the location set at function entry,
+ as a function of the location set [caller] of the calling function.
+- Machine registers have the same values as in the caller.
+- Incoming stack slots (used for parameter passing) have the same
+ values as the corresponding outgoing stack slots (used for argument
+ passing) in the caller.
+- Local and outgoing stack slots are initialized to undefined values.
+*)
-(** [postcall_locs ls] returns the location set [ls] after a function
- call. Caller-save registers and temporary registers
- are set to [undef], reflecting the fact that the called
- function can modify them freely. *)
+Definition call_regs (caller: locset) : locset :=
+ fun (l: loc) =>
+ match l with
+ | R r => caller (R r)
+ | S Local ofs ty => Vundef
+ | S Incoming ofs ty => caller (S Outgoing ofs ty)
+ | S Outgoing ofs ty => Vundef
+ end.
-Definition postcall_locs (ls: locset) : locset :=
+(** [return_regs caller callee] returns the location set after
+ a call instruction, as a function of the location set [caller]
+ of the caller before the call instruction and of the location
+ set [callee] of the callee at the return instruction.
+- Callee-save machine registers have the same values as in the caller
+ before the call.
+- Caller-save machine registers have the same values as in the callee.
+- Stack slots have the same values as in the caller.
+*)
+
+Definition return_regs (caller callee: locset) : locset :=
fun (l: loc) =>
match l with
| R r =>
- if In_dec Loc.eq (R r) temporaries then
- Vundef
- else if In_dec Loc.eq (R r) destroyed_at_call then
- Vundef
- else
- ls (R r)
- | S s => ls (S s)
+ if In_dec mreg_eq r destroyed_at_call
+ then callee (R r)
+ else caller (R r)
+ | S sl ofs ty => caller (S sl ofs ty)
end.
-(** Temporaries destroyed across instructions *)
-
-Definition undef_temps (ls: locset) := Locmap.undef temporaries ls.
-
(** LTL execution states. *)
Inductive stackframe : Type :=
| Stackframe:
- forall (res: loc) (**r where to store the result *)
- (f: function) (**r calling function *)
+ forall (f: function) (**r calling function *)
(sp: val) (**r stack pointer in calling function *)
(ls: locset) (**r location state in calling function *)
- (pc: node), (**r program point in calling function *)
+ (bb: bblock), (**r continuation in calling function *)
stackframe.
Inductive state : Type :=
@@ -125,15 +138,23 @@ Inductive state : Type :=
(ls: locset) (**r location state *)
(m: mem), (**r memory state *)
state
+ | Block:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: function) (**r function currently executing *)
+ (sp: val) (**r stack pointer *)
+ (bb: bblock) (**r current basic block *)
+ (ls: locset) (**r location state *)
+ (m: mem), (**r memory state *)
+ state
| Callstate:
forall (stack: list stackframe) (**r call stack *)
(f: fundef) (**r function to call *)
- (args: list val) (**r arguments to the call *)
+ (ls: locset) (**r location state of caller *)
(m: mem), (**r memory state *)
state
| Returnstate:
forall (stack: list stackframe) (**r call stack *)
- (v: val) (**r return value for the call *)
+ (ls: locset) (**r location state of callee *)
(m: mem), (**r memory state *)
state.
@@ -142,9 +163,24 @@ Section RELSEM.
Variable ge: genv.
-Definition find_function (los: loc + ident) (rs: locset) : option fundef :=
- match los with
- | inl l => Genv.find_funct ge (rs l)
+Definition reglist (rs: locset) (rl: list mreg) : list val :=
+ List.map (fun r => rs (R r)) rl.
+
+Fixpoint undef_regs (rl: list mreg) (rs: locset) : locset :=
+ match rl with
+ | nil => rs
+ | r1 :: rl => Locmap.set (R r1) Vundef (undef_regs rl rs)
+ end.
+
+Definition destroyed_by_getstack (s: slot) : list mreg :=
+ match s with
+ | Incoming => temp_for_parent_frame :: nil
+ | _ => nil
+ end.
+
+Definition find_function (ros: mreg + ident) (rs: locset) : option fundef :=
+ match ros with
+ | inl r => Genv.find_funct ge (rs (R r))
| inr symb =>
match Genv.find_symbol ge symb with
| None => None
@@ -152,95 +188,109 @@ Definition find_function (los: loc + ident) (rs: locset) : option fundef :=
end
end.
-(** The LTL transition relation is very similar to that of RTL,
- with locations being used in place of pseudo-registers.
- The main difference is for the [call] instruction: caller-save
- registers are set to [Vundef] across the call, using the [postcall_locs]
- function defined above. This forces the LTL producer to avoid
- storing values live across the call in a caller-save register. *)
+(** [parent_locset cs] returns the mapping of values for locations
+ of the caller function. *)
+
+Definition parent_locset (stack: list stackframe) : locset :=
+ match stack with
+ | nil => Locmap.init Vundef
+ | Stackframe f sp ls bb :: stack' => ls
+ end.
+
+(* REVISE
+(** [getslot sl ofs ty rs] looks up the value of location [S sl ofs ty] in [rs],
+ and normalizes it to the type [ty] of this location. *)
+
+Definition getslot (sl: slot) (ofs: Z) (ty: typ) (rs: locset) : val :=
+ Val.load_result
+ (match ty with Tint => Mint32 | Tfloat => Mfloat64 | Tlong => Mint64 end)
+ (rs (S sl ofs ty)).
+*)
Inductive step: state -> trace -> state -> Prop :=
- | exec_Lnop:
- forall s f sp pc rs m pc',
- (fn_code f)!pc = Some(Lnop pc') ->
+ | exec_start_block: forall s f sp pc rs m bb,
+ (fn_code f)!pc = Some bb ->
step (State s f sp pc rs m)
- E0 (State s f sp pc' rs m)
- | exec_Lop:
- forall s f sp pc rs m op args res pc' v,
- (fn_code f)!pc = Some(Lop op args res pc') ->
- eval_operation ge sp op (map rs args) m = Some v ->
- step (State s f sp pc rs m)
- E0 (State s f sp pc' (Locmap.set res v (undef_temps rs)) m)
- | exec_Lload:
- forall s f sp pc rs m chunk addr args dst pc' a v,
- (fn_code f)!pc = Some(Lload chunk addr args dst pc') ->
- eval_addressing ge sp addr (map rs args) = Some a ->
+ E0 (Block s f sp bb rs m)
+ | exec_Lop: forall s f sp op args res bb rs m v rs',
+ eval_operation ge sp op (reglist rs args) m = Some v ->
+ rs' = Locmap.set (R res) v (undef_regs (destroyed_by_op op) rs) ->
+ step (Block s f sp (Lop op args res :: bb) rs m)
+ E0 (Block s f sp bb rs' m)
+ | exec_Lload: forall s f sp chunk addr args dst bb rs m a v rs',
+ eval_addressing ge sp addr (reglist rs args) = Some a ->
Mem.loadv chunk m a = Some v ->
- step (State s f sp pc rs m)
- E0 (State s f sp pc' (Locmap.set dst v (undef_temps rs)) m)
- | exec_Lstore:
- forall s f sp pc rs m chunk addr args src pc' a m',
- (fn_code f)!pc = Some(Lstore chunk addr args src pc') ->
- eval_addressing ge sp addr (map rs args) = Some a ->
- Mem.storev chunk m a (rs src) = Some m' ->
- step (State s f sp pc rs m)
- E0 (State s f sp pc' (undef_temps rs) m')
- | exec_Lcall:
- forall s f sp pc rs m sig ros args res pc' f',
- (fn_code f)!pc = Some(Lcall sig ros args res pc') ->
- find_function ros rs = Some f' ->
- funsig f' = sig ->
- step (State s f sp pc rs m)
- E0 (Callstate (Stackframe res f sp (postcall_locs rs) pc' :: s)
- f' (List.map rs args) m)
- | exec_Ltailcall:
- forall s f stk pc rs m sig ros args f' m',
- (fn_code f)!pc = Some(Ltailcall sig ros args) ->
- find_function ros rs = Some f' ->
- funsig f' = sig ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) pc rs m)
- E0 (Callstate s f' (List.map rs args) m')
- | exec_Lbuiltin:
- forall s f sp pc rs m ef args res pc' t v m',
- (fn_code f)!pc = Some(Lbuiltin ef args res pc') ->
- external_call ef ge (map rs args) m t v m' ->
- step (State s f sp pc rs m)
- t (State s f sp pc' (Locmap.set res v rs) m')
- | exec_Lcond:
- forall s f sp pc rs m cond args ifso ifnot b pc',
- (fn_code f)!pc = Some(Lcond cond args ifso ifnot) ->
- eval_condition cond (map rs args) m = Some b ->
- pc' = (if b then ifso else ifnot) ->
- step (State s f sp pc rs m)
- E0 (State s f sp pc' (undef_temps rs) m)
- | exec_Ljumptable:
- forall s f sp pc rs m arg tbl n pc',
- (fn_code f)!pc = Some(Ljumptable arg tbl) ->
- rs arg = Vint n ->
- list_nth_z tbl (Int.unsigned n) = Some pc' ->
- step (State s f sp pc rs m)
- E0 (State s f sp pc' (undef_temps rs) m)
- | exec_Lreturn:
- forall s f stk pc rs m or m',
- (fn_code f)!pc = Some(Lreturn or) ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) pc rs m)
- E0 (Returnstate s (locmap_optget or Vundef rs) m')
- | exec_function_internal:
- forall s f args m m' stk,
- Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
- step (Callstate s (Internal f) args m)
- E0 (State s f (Vptr stk Int.zero) f.(fn_entrypoint) (init_locs args f.(fn_params)) m')
- | exec_function_external:
- forall s ef t args res m m',
- external_call ef ge args m t res m' ->
- step (Callstate s (External ef) args m)
- t (Returnstate s res m')
- | exec_return:
- forall res f sp rs pc s vres m,
- step (Returnstate (Stackframe res f sp rs pc :: s) vres m)
- E0 (State s f sp pc (Locmap.set res vres rs) m).
+ rs' = Locmap.set (R dst) v (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (Block s f sp (Lload chunk addr args dst :: bb) rs m)
+ E0 (Block s f sp bb rs' m)
+ | exec_Lgetstack: forall s f sp sl ofs ty dst bb rs m rs',
+ rs' = Locmap.set (R dst) (rs (S sl ofs ty)) (undef_regs (destroyed_by_getstack sl) rs) ->
+ step (Block s f sp (Lgetstack sl ofs ty dst :: bb) rs m)
+ E0 (Block s f sp bb rs' m)
+ | exec_Lsetstack: forall s f sp src sl ofs ty bb rs m rs',
+ rs' = Locmap.set (S sl ofs ty) (rs (R src)) (undef_regs (destroyed_by_op Omove) rs) ->
+ step (Block s f sp (Lsetstack src sl ofs ty :: bb) rs m)
+ E0 (Block s f sp bb rs' m)
+ | exec_Lstore: forall s f sp chunk addr args src bb rs m a rs' m',
+ eval_addressing ge sp addr (reglist rs args) = Some a ->
+ Mem.storev chunk m a (rs (R src)) = Some m' ->
+ rs' = undef_regs (destroyed_by_store chunk addr) rs ->
+ step (Block s f sp (Lstore chunk addr args src :: bb) rs m)
+ E0 (Block s f sp bb rs' m')
+ | exec_Lcall: forall s f sp sig ros bb rs m fd,
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ step (Block s f sp (Lcall sig ros :: bb) rs m)
+ E0 (Callstate (Stackframe f sp rs bb :: s) fd rs m)
+ | exec_Ltailcall: forall s f sp sig ros bb rs m fd rs' m',
+ rs' = return_regs (parent_locset s) rs ->
+ find_function ros rs' = Some fd ->
+ funsig fd = sig ->
+ Mem.free m sp 0 f.(fn_stacksize) = Some m' ->
+ step (Block s f (Vptr sp Int.zero) (Ltailcall sig ros :: bb) rs m)
+ E0 (Callstate s fd rs' m')
+ | exec_Lbuiltin: forall s f sp ef args res bb rs m t vl rs' m',
+ external_call' ef ge (reglist rs args) m t vl m' ->
+ rs' = Locmap.setlist (map R res) vl (undef_regs (destroyed_by_builtin ef) rs) ->
+ step (Block s f sp (Lbuiltin ef args res :: bb) rs m)
+ t (Block s f sp bb rs' m')
+ | exec_Lannot: forall s f sp ef args bb rs m t vl m',
+ external_call' ef ge (map rs args) m t vl m' ->
+ step (Block s f sp (Lannot ef args :: bb) rs m)
+ t (Block s f sp bb rs m')
+ | exec_Lbranch: forall s f sp pc bb rs m,
+ step (Block s f sp (Lbranch pc :: bb) rs m)
+ E0 (State s f sp pc rs m)
+ | exec_Lcond: forall s f sp cond args pc1 pc2 bb rs b pc rs' m,
+ eval_condition cond (reglist rs args) m = Some b ->
+ pc = (if b then pc1 else pc2) ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
+ step (Block s f sp (Lcond cond args pc1 pc2 :: bb) rs m)
+ E0 (State s f sp pc rs' m)
+ | exec_Ljumptable: forall s f sp arg tbl bb rs m n pc rs',
+ rs (R arg) = Vint n ->
+ list_nth_z tbl (Int.unsigned n) = Some pc ->
+ rs' = undef_regs (destroyed_by_jumptable) rs ->
+ step (Block s f sp (Ljumptable arg tbl :: bb) rs m)
+ E0 (State s f sp pc rs' m)
+ | exec_Lreturn: forall s f sp bb rs m m',
+ Mem.free m sp 0 f.(fn_stacksize) = Some m' ->
+ step (Block s f (Vptr sp Int.zero) (Lreturn :: bb) rs m)
+ E0 (Returnstate s (return_regs (parent_locset s) rs) m')
+ | exec_function_internal: forall s f rs m m' sp rs',
+ Mem.alloc m 0 f.(fn_stacksize) = (m', sp) ->
+ rs' = undef_regs destroyed_at_function_entry (call_regs rs) ->
+ step (Callstate s (Internal f) rs m)
+ E0 (State s f (Vptr sp Int.zero) f.(fn_entrypoint) rs' m')
+ | exec_function_external: forall s ef t args res rs m rs' m',
+ args = map rs (loc_arguments (ef_sig ef)) ->
+ external_call' ef ge args m t res m' ->
+ rs' = Locmap.setlist (map R (loc_result (ef_sig ef))) res rs ->
+ step (Callstate s (External ef) rs m)
+ t (Returnstate s rs' m')
+ | exec_return: forall f sp rs1 bb s rs m,
+ step (Returnstate (Stackframe f sp rs1 bb :: s) rs m)
+ E0 (Block s f sp bb rs m).
End RELSEM.
@@ -256,33 +306,32 @@ Inductive initial_state (p: program): state -> Prop :=
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
funsig f = mksignature nil (Some Tint) ->
- initial_state p (Callstate nil f nil m0).
+ initial_state p (Callstate nil f (Locmap.init Vundef) m0).
Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall n m,
- final_state (Returnstate nil (Vint n) m) n.
+ | final_state_intro: forall rs m r retcode,
+ loc_result (mksignature nil (Some Tint)) = r :: nil ->
+ rs (R r) = Vint retcode ->
+ final_state (Returnstate nil rs m) retcode.
Definition semantics (p: program) :=
Semantics step (initial_state p) final_state (Genv.globalenv p).
(** * Operations over LTL *)
-(** Computation of the possible successors of an instruction.
+(** Computation of the possible successors of a block.
This is used in particular for dataflow analyses. *)
-Definition successors_instr (i: instruction) : list node :=
- match i with
- | Lnop s => s :: nil
- | Lop op args res s => s :: nil
- | Lload chunk addr args dst s => s :: nil
- | Lstore chunk addr args src s => s :: nil
- | Lcall sig ros args res s => s :: nil
- | Ltailcall sig ros args => nil
- | Lbuiltin ef args res s => s :: nil
- | Lcond cond args ifso ifnot => ifso :: ifnot :: nil
- | Ljumptable arg tbl => tbl
- | Lreturn optarg => nil
+Fixpoint successors_block (b: bblock) : list node :=
+ match b with
+ | nil => nil (**r should never happen *)
+ | Ltailcall _ _ :: _ => nil
+ | Lbranch s :: _ => s :: nil
+ | Lcond _ _ s1 s2 :: _ => s1 :: s2 :: nil
+ | Ljumptable _ tbl :: _ => tbl
+ | Lreturn :: _ => nil
+ | instr :: b' => successors_block b'
end.
Definition successors (f: function): PTree.t (list node) :=
- PTree.map1 successors_instr f.(fn_code).
+ PTree.map1 successors_block f.(fn_code).
diff --git a/backend/LTLin.v b/backend/LTLin.v
deleted file mode 100644
index e0d5ca2..0000000
--- a/backend/LTLin.v
+++ /dev/null
@@ -1,268 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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 LTLin intermediate language: abstract syntax and semantcs *)
-
-(** The LTLin language is a variant of LTL where control-flow is not
- expressed as a graph of basic blocks, but as a linear list of
- instructions with explicit labels and ``goto'' instructions. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Values.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Op.
-Require Import Locations.
-Require Import LTL.
-
-(** * Abstract syntax *)
-
-Definition label := positive.
-
-(** LTLin instructions are similar to those of LTL.
- Except the last three, these instructions continue in sequence
- with the next instruction in the linear list of instructions.
- Unconditional branches [Lgoto] and conditional branches [Lcond]
- transfer control to the given code label. Labels are explicitly
- inserted in the instruction list as pseudo-instructions [Llabel]. *)
-
-Inductive instruction: Type :=
- | Lop: operation -> list loc -> loc -> instruction
- | Lload: memory_chunk -> addressing -> list loc -> loc -> instruction
- | Lstore: memory_chunk -> addressing -> list loc -> loc -> instruction
- | Lcall: signature -> loc + ident -> list loc -> loc -> instruction
- | Ltailcall: signature -> loc + ident -> list loc -> instruction
- | Lbuiltin: external_function -> list loc -> loc -> instruction
- | Llabel: label -> instruction
- | Lgoto: label -> instruction
- | Lcond: condition -> list loc -> label -> instruction
- | Ljumptable: loc -> list label -> instruction
- | Lreturn: option loc -> instruction.
-
-Definition code: Type := list instruction.
-
-Record function: Type := mkfunction {
- fn_sig: signature;
- fn_params: list loc;
- fn_stacksize: Z;
- fn_code: code
-}.
-
-Definition fundef := AST.fundef function.
-
-Definition program := AST.program fundef unit.
-
-Definition funsig (fd: fundef) :=
- match fd with
- | Internal f => fn_sig f
- | External ef => ef_sig ef
- end.
-
-Definition genv := Genv.t fundef unit.
-Definition locset := Locmap.t.
-
-(** * Operational semantics *)
-
-(** Looking up labels in the instruction list. *)
-
-Definition is_label (lbl: label) (instr: instruction) : bool :=
- match instr with
- | Llabel lbl' => if peq lbl lbl' then true else false
- | _ => false
- end.
-
-Lemma is_label_correct:
- forall lbl instr,
- if is_label lbl instr then instr = Llabel lbl else instr <> Llabel lbl.
-Proof.
- intros. destruct instr; simpl; try discriminate.
- case (peq lbl l); intro; congruence.
-Qed.
-
-(** [find_label lbl c] returns a list of instruction, suffix of the
- code [c], that immediately follows the [Llabel lbl] pseudo-instruction.
- If the label [lbl] is multiply-defined, the first occurrence is
- retained. If the label [lbl] is not defined, [None] is returned. *)
-
-Fixpoint find_label (lbl: label) (c: code) {struct c} : option code :=
- match c with
- | nil => None
- | i1 :: il => if is_label lbl i1 then Some il else find_label lbl il
- end.
-
-(** The states of the dynamic semantics are similar to those used
- in the LTL semantics (see module [LTL]). The only difference
- is that program points [pc] (nodes of the CFG in LTL) become
- code sequences [c] (suffixes of the code of the current function).
-*)
-
-Inductive stackframe : Type :=
- | Stackframe:
- forall (res: loc) (**r where to store the result *)
- (f: function) (**r calling function *)
- (sp: val) (**r stack pointer in calling function *)
- (ls: locset) (**r location state in calling function *)
- (c: code), (**r program point in calling function *)
- stackframe.
-
-Inductive state : Type :=
- | State:
- forall (stack: list stackframe) (**r call stack *)
- (f: function) (**r function currently executing *)
- (sp: val) (**r stack pointer *)
- (c: code) (**r current program point *)
- (ls: locset) (**r location state *)
- (m: mem), (**r memory state *)
- state
- | Callstate:
- forall (stack: list stackframe) (**r call stack *)
- (f: fundef) (**r function to call *)
- (args: list val) (**r arguments to the call *)
- (m: mem), (**r memory state *)
- state
- | Returnstate:
- forall (stack: list stackframe) (**r call stack *)
- (v: val) (**r return value for the call *)
- (m: mem), (**r memory state *)
- state.
-
-Section RELSEM.
-
-Variable ge: genv.
-
-Definition find_function (ros: loc + ident) (rs: locset) : option fundef :=
- match ros with
- | inl r => Genv.find_funct ge (rs r)
- | inr symb =>
- match Genv.find_symbol ge symb with
- | None => None
- | Some b => Genv.find_funct_ptr ge b
- end
- end.
-
-Inductive step: state -> trace -> state -> Prop :=
- | exec_Lop:
- forall s f sp op args res b rs m v,
- eval_operation ge sp op (map rs args) m = Some v ->
- step (State s f sp (Lop op args res :: b) rs m)
- E0 (State s f sp b (Locmap.set res v (undef_temps rs)) m)
- | exec_Lload:
- forall s f sp chunk addr args dst b rs m a v,
- eval_addressing ge sp addr (map rs args) = Some a ->
- Mem.loadv chunk m a = Some v ->
- step (State s f sp (Lload chunk addr args dst :: b) rs m)
- E0 (State s f sp b (Locmap.set dst v (undef_temps rs)) m)
- | exec_Lstore:
- forall s f sp chunk addr args src b rs m m' a,
- eval_addressing ge sp addr (map rs args) = Some a ->
- Mem.storev chunk m a (rs src) = Some m' ->
- step (State s f sp (Lstore chunk addr args src :: b) rs m)
- E0 (State s f sp b (undef_temps rs) m')
- | exec_Lcall:
- forall s f sp sig ros args res b rs m f',
- find_function ros rs = Some f' ->
- sig = funsig f' ->
- step (State s f sp (Lcall sig ros args res :: b) rs m)
- E0 (Callstate (Stackframe res f sp (postcall_locs rs) b :: s)
- f' (List.map rs args) m)
- | exec_Ltailcall:
- forall s f stk sig ros args b rs m f' m',
- find_function ros rs = Some f' ->
- sig = funsig f' ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) (Ltailcall sig ros args :: b) rs m)
- E0 (Callstate s f' (List.map rs args) m')
- | exec_Lbuiltin:
- forall s f sp rs m ef args res b t v m',
- external_call ef ge (map rs args) m t v m' ->
- step (State s f sp (Lbuiltin ef args res :: b) rs m)
- t (State s f sp b (Locmap.set res v rs) m')
- | exec_Llabel:
- forall s f sp lbl b rs m,
- step (State s f sp (Llabel lbl :: b) rs m)
- E0 (State s f sp b rs m)
- | exec_Lgoto:
- forall s f sp lbl b rs m b',
- find_label lbl f.(fn_code) = Some b' ->
- step (State s f sp (Lgoto lbl :: b) rs m)
- E0 (State s f sp b' rs m)
- | exec_Lcond_true:
- forall s f sp cond args lbl b rs m b',
- eval_condition cond (map rs args) m = Some true ->
- find_label lbl f.(fn_code) = Some b' ->
- step (State s f sp (Lcond cond args lbl :: b) rs m)
- E0 (State s f sp b' (undef_temps rs) m)
- | exec_Lcond_false:
- forall s f sp cond args lbl b rs m,
- eval_condition cond (map rs args) m = Some false ->
- step (State s f sp (Lcond cond args lbl :: b) rs m)
- E0 (State s f sp b (undef_temps rs) m)
- | exec_Ljumptable:
- forall s f sp arg tbl b rs m n lbl b',
- rs arg = Vint n ->
- list_nth_z tbl (Int.unsigned n) = Some lbl ->
- find_label lbl f.(fn_code) = Some b' ->
- step (State s f sp (Ljumptable arg tbl :: b) rs m)
- E0 (State s f sp b' (undef_temps rs) m)
- | exec_Lreturn:
- forall s f stk rs m or b m',
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step (State s f (Vptr stk Int.zero) (Lreturn or :: b) rs m)
- E0 (Returnstate s (locmap_optget or Vundef rs) m')
- | exec_function_internal:
- forall s f args m m' stk,
- Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
- step (Callstate s (Internal f) args m)
- E0 (State s f (Vptr stk Int.zero) f.(fn_code) (init_locs args f.(fn_params)) m')
- | exec_function_external:
- forall s ef args t res m m',
- external_call ef ge args m t res m' ->
- step (Callstate s (External ef) args m)
- t (Returnstate s res m')
- | exec_return:
- forall res f sp rs b s vres m,
- step (Returnstate (Stackframe res f sp rs b :: s) vres m)
- E0 (State s f sp b (Locmap.set res vres rs) m).
-
-End RELSEM.
-
-Inductive initial_state (p: program): state -> Prop :=
- | initial_state_intro: forall b f m0,
- let ge := Genv.globalenv p in
- Genv.init_mem p = Some m0 ->
- Genv.find_symbol ge p.(prog_main) = Some b ->
- Genv.find_funct_ptr ge b = Some f ->
- funsig f = mksignature nil (Some Tint) ->
- initial_state p (Callstate nil f nil m0).
-
-Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall n m,
- final_state (Returnstate nil (Vint n) m) n.
-
-Definition semantics (p: program) :=
- Semantics step (initial_state p) final_state (Genv.globalenv p).
-
-(** * Properties of the operational semantics *)
-
-Lemma find_label_is_tail:
- forall lbl c c', find_label lbl c = Some c' -> is_tail c' c.
-Proof.
- induction c; simpl; intros.
- discriminate.
- destruct (is_label lbl a). inv H. constructor. constructor.
- constructor. auto.
-Qed.
-
diff --git a/backend/LTLintyping.v b/backend/LTLintyping.v
deleted file mode 100644
index 0338667..0000000
--- a/backend/LTLintyping.v
+++ /dev/null
@@ -1,122 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Typing rules for LTLin. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Op.
-Require Import Locations.
-Require Import LTLin.
-Require LTLtyping.
-Require Import Conventions.
-
-(** The following predicates define a type system for LTLin similar to that
- of LTL. *)
-
-Section WT_INSTR.
-
-Variable funsig: signature.
-
-Inductive wt_instr : instruction -> Prop :=
- | wt_Lopmove:
- forall r1 r,
- Loc.type r1 = Loc.type r -> loc_acceptable r1 -> loc_acceptable r ->
- wt_instr (Lop Omove (r1 :: nil) r)
- | wt_Lop:
- forall op args res,
- op <> Omove ->
- (List.map Loc.type args, Loc.type res) = type_of_operation op ->
- locs_acceptable args -> loc_acceptable res ->
- wt_instr (Lop op args res)
- | wt_Lload:
- forall chunk addr args dst,
- List.map Loc.type args = type_of_addressing addr ->
- Loc.type dst = type_of_chunk chunk ->
- locs_acceptable args -> loc_acceptable dst ->
- wt_instr (Lload chunk addr args dst)
- | wt_Lstore:
- forall chunk addr args src,
- List.map Loc.type args = type_of_addressing addr ->
- Loc.type src = type_of_chunk chunk ->
- locs_acceptable args -> loc_acceptable src ->
- wt_instr (Lstore chunk addr args src)
- | wt_Lcall:
- forall sig ros args res,
- List.map Loc.type args = sig.(sig_args) ->
- Loc.type res = proj_sig_res sig ->
- LTLtyping.call_loc_acceptable sig ros ->
- locs_acceptable args -> loc_acceptable res ->
- wt_instr (Lcall sig ros args res)
- | wt_Ltailcall:
- forall sig ros args,
- List.map Loc.type args = sig.(sig_args) ->
- LTLtyping.call_loc_acceptable sig ros ->
- locs_acceptable args ->
- sig.(sig_res) = funsig.(sig_res) ->
- tailcall_possible sig ->
- wt_instr (Ltailcall sig ros args)
- | wt_Lbuiltin:
- forall ef args res,
- List.map Loc.type args = (ef_sig ef).(sig_args) ->
- Loc.type res = proj_sig_res (ef_sig ef) ->
- arity_ok (ef_sig ef).(sig_args) = true \/ ef_reloads ef = false ->
- locs_acceptable args -> loc_acceptable res ->
- wt_instr (Lbuiltin ef args res)
- | wt_Llabel: forall lbl,
- wt_instr (Llabel lbl)
- | wt_Lgoto: forall lbl,
- wt_instr (Lgoto lbl)
- | wt_Lcond:
- forall cond args lbl,
- List.map Loc.type args = type_of_condition cond ->
- locs_acceptable args ->
- wt_instr (Lcond cond args lbl)
- | wt_Ljumptable:
- forall arg tbl,
- Loc.type arg = Tint ->
- loc_acceptable arg ->
- list_length_z tbl * 4 <= Int.max_unsigned ->
- wt_instr (Ljumptable arg tbl)
- | wt_Lreturn:
- forall optres,
- option_map Loc.type optres = funsig.(sig_res) ->
- match optres with None => True | Some r => loc_acceptable r end ->
- wt_instr (Lreturn optres).
-
-Definition wt_code (c: code) : Prop :=
- forall i, In i c -> wt_instr i.
-
-End WT_INSTR.
-
-Record wt_function (f: function): Prop :=
- mk_wt_function {
- wt_params:
- List.map Loc.type f.(fn_params) = f.(fn_sig).(sig_args);
- wt_acceptable:
- locs_acceptable f.(fn_params);
- wt_norepet:
- Loc.norepet f.(fn_params);
- wt_instrs:
- wt_code f.(fn_sig) f.(fn_code)
-}.
-
-Inductive wt_fundef: fundef -> Prop :=
- | wt_fundef_external: forall ef,
- wt_fundef (External ef)
- | wt_function_internal: forall f,
- wt_function f ->
- wt_fundef (Internal f).
-
-Definition wt_program (p: program): Prop :=
- forall i f, In (i, Gfun f) (prog_defs p) -> wt_fundef f.
diff --git a/backend/LTLtyping.v b/backend/LTLtyping.v
deleted file mode 100644
index 0c90514..0000000
--- a/backend/LTLtyping.v
+++ /dev/null
@@ -1,143 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Typing rules for LTL. *)
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Op.
-Require Import Locations.
-Require Import LTL.
-Require Import Conventions.
-
-(** The following predicates define a type system for LTL similar to that
- of [RTL] (see file [RTLtyping]): it statically guarantees that operations
- and addressing modes are applied to the right number of arguments
- and that the arguments are of the correct types. Moreover, it also
- guarantees that the locations of arguments and results are "acceptable",
- i.e. either non-temporary registers or [Local] stack locations. *)
-
-Section WT_INSTR.
-
-Variable funct: function.
-
-Definition valid_successor (s: node) : Prop :=
- exists i, funct.(fn_code)!s = Some i.
-
-Definition call_loc_acceptable (sig: signature) (los: loc + ident) : Prop :=
- match los with
- | inl l => Loc.type l = Tint /\ loc_acceptable l /\ ~In l (loc_arguments sig)
- | inr s => True
- end.
-
-Inductive wt_instr : instruction -> Prop :=
- | wt_Lnop:
- forall s,
- valid_successor s ->
- wt_instr (Lnop s)
- | wt_Lopmove:
- forall r1 r s,
- Loc.type r1 = Loc.type r -> loc_acceptable r1 -> loc_acceptable r ->
- valid_successor s ->
- wt_instr (Lop Omove (r1 :: nil) r s)
- | wt_Lop:
- forall op args res s,
- op <> Omove ->
- (List.map Loc.type args, Loc.type res) = type_of_operation op ->
- locs_acceptable args -> loc_acceptable res ->
- valid_successor s ->
- wt_instr (Lop op args res s)
- | wt_Lload:
- forall chunk addr args dst s,
- List.map Loc.type args = type_of_addressing addr ->
- Loc.type dst = type_of_chunk chunk ->
- locs_acceptable args -> loc_acceptable dst ->
- valid_successor s ->
- wt_instr (Lload chunk addr args dst s)
- | wt_Lstore:
- forall chunk addr args src s,
- List.map Loc.type args = type_of_addressing addr ->
- Loc.type src = type_of_chunk chunk ->
- locs_acceptable args -> loc_acceptable src ->
- valid_successor s ->
- wt_instr (Lstore chunk addr args src s)
- | wt_Lcall:
- forall sig ros args res s,
- List.map Loc.type args = sig.(sig_args) ->
- Loc.type res = proj_sig_res sig ->
- call_loc_acceptable sig ros ->
- locs_acceptable args -> loc_acceptable res ->
- valid_successor s ->
- wt_instr (Lcall sig ros args res s)
- | wt_Ltailcall:
- forall sig ros args,
- List.map Loc.type args = sig.(sig_args) ->
- call_loc_acceptable sig ros ->
- locs_acceptable args ->
- sig.(sig_res) = funct.(fn_sig).(sig_res) ->
- tailcall_possible sig ->
- wt_instr (Ltailcall sig ros args)
- | wt_Lbuiltin:
- forall ef args res s,
- List.map Loc.type args = (ef_sig ef).(sig_args) ->
- Loc.type res = proj_sig_res (ef_sig ef) ->
- arity_ok (ef_sig ef).(sig_args) = true \/ ef_reloads ef = false ->
- locs_acceptable args -> loc_acceptable res ->
- valid_successor s ->
- wt_instr (Lbuiltin ef args res s)
- | wt_Lcond:
- forall cond args s1 s2,
- List.map Loc.type args = type_of_condition cond ->
- locs_acceptable args ->
- valid_successor s1 -> valid_successor s2 ->
- wt_instr (Lcond cond args s1 s2)
- | wt_Ljumptable:
- forall arg tbl,
- Loc.type arg = Tint ->
- loc_acceptable arg ->
- (forall lbl, In lbl tbl -> valid_successor lbl) ->
- list_length_z tbl * 4 <= Int.max_unsigned ->
- wt_instr (Ljumptable arg tbl)
- | wt_Lreturn:
- forall optres,
- option_map Loc.type optres = funct.(fn_sig).(sig_res) ->
- match optres with None => True | Some r => loc_acceptable r end ->
- wt_instr (Lreturn optres).
-
-End WT_INSTR.
-
-Record wt_function (f: function): Prop :=
- mk_wt_function {
- wt_params:
- List.map Loc.type f.(fn_params) = f.(fn_sig).(sig_args);
- wt_acceptable:
- locs_acceptable f.(fn_params);
- wt_norepet:
- Loc.norepet f.(fn_params);
- wt_instrs:
- forall pc instr,
- f.(fn_code)!pc = Some instr -> wt_instr f instr;
- wt_entrypoint:
- valid_successor f f.(fn_entrypoint)
-}.
-
-Inductive wt_fundef: fundef -> Prop :=
- | wt_fundef_external: forall ef,
- wt_fundef (External ef)
- | wt_function_internal: forall f,
- wt_function f ->
- wt_fundef (Internal f).
-
-Definition wt_program (p: program): Prop :=
- forall i f, In (i, Gfun f) (prog_defs p) -> wt_fundef f.
diff --git a/backend/Linear.v b/backend/Linear.v
index 52f5fd7..bdb08b4 100644
--- a/backend/Linear.v
+++ b/backend/Linear.v
@@ -12,10 +12,9 @@
(** The Linear intermediate language: abstract syntax and semantcs *)
-(** The Linear language is a variant of LTLin where arithmetic
- instructions operate on machine registers (type [mreg]) instead
- of arbitrary locations. Special instructions [Lgetstack] and
- [Lsetstack] are provided to access stack slots. *)
+(** The Linear language is a variant of LTL where control-flow is not
+ expressed as a graph of basic blocks, but as a linear list of
+ instructions with explicit labels and ``goto'' instructions. *)
Require Import Coqlib.
Require Import AST.
@@ -35,14 +34,14 @@ Require Import Conventions.
Definition label := positive.
Inductive instruction: Type :=
- | Lgetstack: slot -> mreg -> instruction
- | Lsetstack: mreg -> slot -> instruction
+ | Lgetstack: slot -> Z -> typ -> mreg -> instruction
+ | Lsetstack: mreg -> slot -> Z -> typ -> instruction
| Lop: operation -> list mreg -> mreg -> instruction
| Lload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lcall: signature -> mreg + ident -> instruction
| Ltailcall: signature -> mreg + ident -> instruction
- | Lbuiltin: external_function -> list mreg -> mreg -> instruction
+ | Lbuiltin: external_function -> list mreg -> list mreg -> instruction
| Lannot: external_function -> list loc -> instruction
| Llabel: label -> instruction
| Lgoto: label -> instruction
@@ -114,73 +113,6 @@ Definition find_function (ros: mreg + ident) (rs: locset) : option fundef :=
end
end.
-Definition reglist (rs: locset) (rl: list mreg) : list val :=
- List.map (fun r => rs (R r)) rl.
-
-(** Calling conventions are reflected at the level of location sets
- (environments mapping locations to values) by the following two
- functions.
-
- [call_regs caller] returns the location set at function entry,
- as a function of the location set [caller] of the calling function.
-- Temporary registers are undefined.
-- Other machine registers have the same values as in the caller.
-- Incoming stack slots (used for parameter passing) have the same
- values as the corresponding outgoing stack slots (used for argument
- passing) in the caller.
-- Local and outgoing stack slots are initialized to undefined values.
-*)
-
-Definition call_regs (caller: locset) : locset :=
- fun (l: loc) =>
- match l with
- | R r => if In_dec Loc.eq (R r) temporaries then Vundef else caller (R r)
- | S (Local ofs ty) => Vundef
- | S (Incoming ofs ty) => caller (S (Outgoing ofs ty))
- | S (Outgoing ofs ty) => Vundef
- end.
-
-(** [return_regs caller callee] returns the location set after
- a call instruction, as a function of the location set [caller]
- of the caller before the call instruction and of the location
- set [callee] of the callee at the return instruction.
-- Callee-save machine registers have the same values as in the caller
- before the call.
-- Caller-save machine registers have the same values
- as in the callee at the return point.
-- Stack slots have the same values as in the caller before the call.
-*)
-
-Definition return_regs (caller callee: locset) : locset :=
- fun (l: loc) =>
- match l with
- | R r =>
- if In_dec Loc.eq (R r) temporaries then
- callee (R r)
- else if In_dec Loc.eq (R r) destroyed_at_call then
- callee (R r)
- else
- caller (R r)
- | S s => caller (S s)
- end.
-
-(** Temporaries destroyed across operations *)
-
-Definition undef_op (op: operation) (rs: locset) :=
- match op with
- | Omove => Locmap.undef destroyed_at_move rs
- | _ => Locmap.undef temporaries rs
- end.
-
-Definition undef_getstack (s: slot) (rs: locset) :=
- match s with
- | Incoming _ _ => Locmap.set (R IT1) Vundef rs
- | _ => rs
- end.
-
-Definition undef_setstack (rs: locset) :=
- Locmap.undef destroyed_at_move rs.
-
(** Linear execution states. *)
Inductive stackframe: Type :=
@@ -214,73 +146,43 @@ Inductive state: Type :=
(** [parent_locset cs] returns the mapping of values for locations
of the caller function. *)
-
Definition parent_locset (stack: list stackframe) : locset :=
match stack with
| nil => Locmap.init Vundef
| Stackframe f sp ls c :: stack' => ls
end.
-(** The main difference between the Linear transition relation
- and the LTL transition relation is the handling of function calls.
- In LTL, arguments and results to calls are transmitted via
- [vargs] and [vres] components of [Callstate] and [Returnstate],
- respectively. The semantics takes care of transferring these values
- between the locations of the caller and of the callee.
-
- In Linear and lower-level languages (Mach, PPC), arguments and
- results are transmitted implicitly: the generated code for the
- caller arranges for arguments to be left in conventional registers
- and stack locations, as determined by the calling conventions, where
- the function being called will find them. Similarly, conventional
- registers will be used to pass the result value back to the caller.
- This is reflected in the definition of [Callstate] and [Returnstate]
- above, where a whole location state [rs] is passed instead of just
- the values of arguments or returns as in LTL.
-
- These location states passed across calls are treated in a way that
- reflects the callee-save/caller-save convention:
-- The [exec_Lcall] transition from [State] to [Callstate]
- saves the current location state [ls] in the call stack,
- and passes it to the callee.
-- The [exec_function_internal] transition from [Callstate] to [State]
- changes the view of stack slots ([Outgoing] slots slide to
- [Incoming] slots as per [call_regs]).
-- The [exec_Lreturn] transition from [State] to [Returnstate]
- restores the values of callee-save locations from
- the location state of the caller, using [return_regs].
-
-This protocol makes it much easier to later prove the correctness of
-the [Stacking] pass, which inserts actual code that performs the
-saving and restoring of callee-save registers described above.
-*)
-
Inductive step: state -> trace -> state -> Prop :=
| exec_Lgetstack:
- forall s f sp sl r b rs m,
- step (State s f sp (Lgetstack sl r :: b) rs m)
- E0 (State s f sp b (Locmap.set (R r) (rs (S sl)) (undef_getstack sl rs)) m)
+ forall s f sp sl ofs ty dst b rs m rs',
+ rs' = Locmap.set (R dst) (rs (S sl ofs ty)) (undef_regs (destroyed_by_getstack sl) rs) ->
+ step (State s f sp (Lgetstack sl ofs ty dst :: b) rs m)
+ E0 (State s f sp b rs' m)
| exec_Lsetstack:
- forall s f sp r sl b rs m,
- step (State s f sp (Lsetstack r sl :: b) rs m)
- E0 (State s f sp b (Locmap.set (S sl) (rs (R r)) (undef_setstack rs)) m)
+ forall s f sp src sl ofs ty b rs m rs',
+ rs' = Locmap.set (S sl ofs ty) (rs (R src)) (undef_regs (destroyed_by_op Omove) rs) ->
+ step (State s f sp (Lsetstack src sl ofs ty :: b) rs m)
+ E0 (State s f sp b rs' m)
| exec_Lop:
- forall s f sp op args res b rs m v,
+ forall s f sp op args res b rs m v rs',
eval_operation ge sp op (reglist rs args) m = Some v ->
+ rs' = Locmap.set (R res) v (undef_regs (destroyed_by_op op) rs) ->
step (State s f sp (Lop op args res :: b) rs m)
- E0 (State s f sp b (Locmap.set (R res) v (undef_op op rs)) m)
+ E0 (State s f sp b rs' m)
| exec_Lload:
- forall s f sp chunk addr args dst b rs m a v,
+ forall s f sp chunk addr args dst b rs m a v rs',
eval_addressing ge sp addr (reglist rs args) = Some a ->
Mem.loadv chunk m a = Some v ->
+ rs' = Locmap.set (R dst) v (undef_regs (destroyed_by_load chunk addr) rs) ->
step (State s f sp (Lload chunk addr args dst :: b) rs m)
- E0 (State s f sp b (Locmap.set (R dst) v (undef_temps rs)) m)
+ E0 (State s f sp b rs' m)
| exec_Lstore:
- forall s f sp chunk addr args src b rs m m' a,
+ forall s f sp chunk addr args src b rs m m' a rs',
eval_addressing ge sp addr (reglist rs args) = Some a ->
Mem.storev chunk m a (rs (R src)) = Some m' ->
+ rs' = undef_regs (destroyed_by_store chunk addr) rs ->
step (State s f sp (Lstore chunk addr args src :: b) rs m)
- E0 (State s f sp b (undef_temps rs) m')
+ E0 (State s f sp b rs' m')
| exec_Lcall:
forall s f sp sig ros b rs m f',
find_function ros rs = Some f' ->
@@ -288,20 +190,22 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp (Lcall sig ros :: b) rs m)
E0 (Callstate (Stackframe f sp rs b:: s) f' rs m)
| exec_Ltailcall:
- forall s f stk sig ros b rs m f' m',
- find_function ros rs = Some f' ->
+ forall s f stk sig ros b rs m rs' f' m',
+ rs' = return_regs (parent_locset s) rs ->
+ find_function ros rs' = Some f' ->
sig = funsig f' ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
step (State s f (Vptr stk Int.zero) (Ltailcall sig ros :: b) rs m)
- E0 (Callstate s f' (return_regs (parent_locset s) rs) m')
+ E0 (Callstate s f' rs' m')
| exec_Lbuiltin:
- forall s f sp rs m ef args res b t v m',
- external_call ef ge (reglist rs args) m t v m' ->
+ forall s f sp rs m ef args res b t vl rs' m',
+ external_call' ef ge (reglist rs args) m t vl m' ->
+ rs' = Locmap.setlist (map R res) vl (undef_regs (destroyed_by_builtin ef) rs) ->
step (State s f sp (Lbuiltin ef args res :: b) rs m)
- t (State s f sp b (Locmap.set (R res) v (undef_temps rs)) m')
+ t (State s f sp b rs' m')
| exec_Lannot:
forall s f sp rs m ef args b t v m',
- external_call ef ge (map rs args) m t v m' ->
+ external_call' ef ge (map rs args) m t v m' ->
step (State s f sp (Lannot ef args :: b) rs m)
t (State s f sp b rs m')
| exec_Llabel:
@@ -314,38 +218,42 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp (Lgoto lbl :: b) rs m)
E0 (State s f sp b' rs m)
| exec_Lcond_true:
- forall s f sp cond args lbl b rs m b',
+ forall s f sp cond args lbl b rs m rs' b',
eval_condition cond (reglist rs args) m = Some true ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
find_label lbl f.(fn_code) = Some b' ->
step (State s f sp (Lcond cond args lbl :: b) rs m)
- E0 (State s f sp b' (undef_temps rs) m)
+ E0 (State s f sp b' rs' m)
| exec_Lcond_false:
- forall s f sp cond args lbl b rs m,
+ forall s f sp cond args lbl b rs m rs',
eval_condition cond (reglist rs args) m = Some false ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
step (State s f sp (Lcond cond args lbl :: b) rs m)
- E0 (State s f sp b (undef_temps rs) m)
+ E0 (State s f sp b rs' m)
| exec_Ljumptable:
- forall s f sp arg tbl b rs m n lbl b',
+ forall s f sp arg tbl b rs m n lbl b' rs',
rs (R arg) = Vint n ->
list_nth_z tbl (Int.unsigned n) = Some lbl ->
find_label lbl f.(fn_code) = Some b' ->
+ rs' = undef_regs (destroyed_by_jumptable) rs ->
step (State s f sp (Ljumptable arg tbl :: b) rs m)
- E0 (State s f sp b' (undef_temps rs) m)
+ E0 (State s f sp b' rs' m)
| exec_Lreturn:
forall s f stk b rs m m',
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
step (State s f (Vptr stk Int.zero) (Lreturn :: b) rs m)
E0 (Returnstate s (return_regs (parent_locset s) rs) m')
| exec_function_internal:
- forall s f rs m m' stk,
+ forall s f rs m rs' m' stk,
Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
+ rs' = undef_regs destroyed_at_function_entry (call_regs rs) ->
step (Callstate s (Internal f) rs m)
- E0 (State s f (Vptr stk Int.zero) f.(fn_code) (call_regs rs) m')
+ E0 (State s f (Vptr stk Int.zero) f.(fn_code) rs' m')
| exec_function_external:
forall s ef args res rs1 rs2 m t m',
- external_call ef ge args m t res m' ->
- args = List.map rs1 (loc_arguments (ef_sig ef)) ->
- rs2 = Locmap.set (R (loc_result (ef_sig ef))) res rs1 ->
+ args = map rs1 (loc_arguments (ef_sig ef)) ->
+ external_call' ef ge args m t res m' ->
+ rs2 = Locmap.setlist (map R (loc_result (ef_sig ef))) res rs1 ->
step (Callstate s (External ef) rs1 m)
t (Returnstate s rs2 m')
| exec_return:
@@ -365,9 +273,10 @@ Inductive initial_state (p: program): state -> Prop :=
initial_state p (Callstate nil f (Locmap.init Vundef) m0).
Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall rs m r,
- rs (R (loc_result (mksignature nil (Some Tint)))) = Vint r ->
- final_state (Returnstate nil rs m) r.
+ | final_state_intro: forall rs m r retcode,
+ loc_result (mksignature nil (Some Tint)) = r :: nil ->
+ rs (R r) = Vint retcode ->
+ final_state (Returnstate nil rs m) retcode.
Definition semantics (p: program) :=
Semantics step (initial_state p) final_state (Genv.globalenv p).
diff --git a/backend/Linearize.v b/backend/Linearize.v
index 6fc8e48..388f459 100644
--- a/backend/Linearize.v
+++ b/backend/Linearize.v
@@ -10,8 +10,7 @@
(* *)
(* *********************************************************************)
-(** Linearization of the control-flow graph:
- translation from LTL to LTLin *)
+(** Linearization of the control-flow graph: translation from LTL to Linear *)
Require Import Coqlib.
Require Import Maps.
@@ -23,26 +22,26 @@ Require Import Errors.
Require Import Op.
Require Import Locations.
Require Import LTL.
-Require Import LTLin.
+Require Import Linear.
Require Import Kildall.
Require Import Lattice.
Open Scope error_monad_scope.
-(** To translate from LTL to LTLin, we must lay out the nodes
+(** To translate from LTL to Linear, we must lay out the nodes
of the LTL control-flow graph in some linear order, and insert
explicit branches and conditional branches to make sure that
each node jumps to its successors as prescribed by the
LTL control-flow graph. However, branches are not necessary
if the fall-through behaviour of LTLin instructions already
implements the desired flow of control. For instance,
- consider the two LTL instructions
+ consider the two LTL blocks
<<
- L1: Lop op args res L2
+ L1: Lop op args res; Lbranch L2
L2: ...
>>
If the instructions [L1] and [L2] are laid out consecutively in the LTLin
- code, we can generate the following LTLin code:
+ code, we can generate the following Linear code:
<<
L1: Lop op args res
L2: ...
@@ -65,7 +64,7 @@ Open Scope error_monad_scope.
- Choosing an order for the nodes. This returns an enumeration of CFG
nodes stating that they must be laid out in the order shown in the
list.
-- Generate LTLin code where each node branches explicitly to its
+- Generate Linear code where each node branches explicitly to its
successors, except if one of these successors is the immediately
following instruction.
@@ -102,7 +101,7 @@ Definition reachable (f: LTL.function) : PMap.t bool :=
| Some rs => rs
end.
-(** We then enumerate the nodes of reachable instructions.
+(** We then enumerate the nodes of reachable blocks.
This task is performed by external, untrusted Caml code. *)
Parameter enumerate_aux: LTL.function -> PMap.t bool -> list node.
@@ -126,7 +125,7 @@ Fixpoint nodeset_of_list (l: list node) (s: Nodeset.t)
Definition check_reachable_aux
(reach: PMap.t bool) (s: Nodeset.t)
- (ok: bool) (pc: node) (i: LTL.instruction) : bool :=
+ (ok: bool) (pc: node) (bb: LTL.bblock) : bool :=
if reach!!pc then ok && Nodeset.mem pc s else ok.
Definition check_reachable
@@ -141,10 +140,10 @@ Definition enumerate (f: LTL.function) : res (list node) :=
then OK enum
else Error (msg "Linearize: wrong enumeration").
-(** * Translation from LTL to LTLin *)
+(** * Translation from LTL to Linear *)
(** We now flatten the structure of the CFG graph, laying out
- LTL instructions consecutively in the order computed by [enumerate],
+ LTL blocks consecutively in the order computed by [enumerate],
and inserting branches to the labels of sucessors if necessary.
Whether to insert a branch or not is determined by
the [starts_with] function below.
@@ -169,31 +168,38 @@ Fixpoint starts_with (lbl: label) (k: code) {struct k} : bool :=
Definition add_branch (s: label) (k: code) : code :=
if starts_with s k then k else Lgoto s :: k.
-Definition linearize_instr (b: LTL.instruction) (k: code) : code :=
+Fixpoint linearize_block (b: LTL.bblock) (k: code) : code :=
match b with
- | LTL.Lnop s =>
+ | nil => k
+ | LTL.Lop op args res :: b' =>
+ Lop op args res :: linearize_block b' k
+ | LTL.Lload chunk addr args dst :: b' =>
+ Lload chunk addr args dst :: linearize_block b' k
+ | LTL.Lgetstack sl ofs ty dst :: b' =>
+ Lgetstack sl ofs ty dst :: linearize_block b' k
+ | LTL.Lsetstack src sl ofs ty :: b' =>
+ Lsetstack src sl ofs ty :: linearize_block b' k
+ | LTL.Lstore chunk addr args src :: b' =>
+ Lstore chunk addr args src :: linearize_block b' k
+ | LTL.Lcall sig ros :: b' =>
+ Lcall sig ros :: linearize_block b' k
+ | LTL.Ltailcall sig ros :: b' =>
+ Ltailcall sig ros :: k
+ | LTL.Lbuiltin ef args res :: b' =>
+ Lbuiltin ef args res :: linearize_block b' k
+ | LTL.Lannot ef args :: b' =>
+ Lannot ef args :: linearize_block b' k
+ | LTL.Lbranch s :: b' =>
add_branch s k
- | LTL.Lop op args res s =>
- Lop op args res :: add_branch s k
- | LTL.Lload chunk addr args dst s =>
- Lload chunk addr args dst :: add_branch s k
- | LTL.Lstore chunk addr args src s =>
- Lstore chunk addr args src :: add_branch s k
- | LTL.Lcall sig ros args res s =>
- Lcall sig ros args res :: add_branch s k
- | LTL.Ltailcall sig ros args =>
- Ltailcall sig ros args :: k
- | LTL.Lbuiltin ef args res s =>
- Lbuiltin ef args res :: add_branch s k
- | LTL.Lcond cond args s1 s2 =>
+ | LTL.Lcond cond args s1 s2 :: b' =>
if starts_with s1 k then
Lcond (negate_condition cond) args s2 :: add_branch s1 k
else
Lcond cond args s1 :: add_branch s2 k
- | LTL.Ljumptable arg tbl =>
+ | LTL.Ljumptable arg tbl :: b' =>
Ljumptable arg tbl :: k
- | LTL.Lreturn or =>
- Lreturn or :: k
+ | LTL.Lreturn :: b' =>
+ Lreturn :: k
end.
(** Linearize a function body according to an enumeration of its nodes. *)
@@ -201,7 +207,7 @@ Definition linearize_instr (b: LTL.instruction) (k: code) : code :=
Definition linearize_node (f: LTL.function) (pc: node) (k: code) : code :=
match f.(LTL.fn_code)!pc with
| None => k
- | Some b => Llabel pc :: linearize_instr b k
+ | Some b => Llabel pc :: linearize_block b k
end.
Definition linearize_body (f: LTL.function) (enum: list node) : code :=
@@ -209,16 +215,15 @@ Definition linearize_body (f: LTL.function) (enum: list node) : code :=
(** * Entry points for code linearization *)
-Definition transf_function (f: LTL.function) : res LTLin.function :=
+Definition transf_function (f: LTL.function) : res Linear.function :=
do enum <- enumerate f;
OK (mkfunction
(LTL.fn_sig f)
- (LTL.fn_params f)
(LTL.fn_stacksize f)
(add_branch (LTL.fn_entrypoint f) (linearize_body f enum))).
-Definition transf_fundef (f: LTL.fundef) : res LTLin.fundef :=
+Definition transf_fundef (f: LTL.fundef) : res Linear.fundef :=
AST.transf_partial_fundef transf_function f.
-Definition transf_program (p: LTL.program) : res LTLin.program :=
+Definition transf_program (p: LTL.program) : res Linear.program :=
transform_partial_program transf_fundef p.
diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml
index ac47ae8..5bb5838 100644
--- a/backend/Linearizeaux.ml
+++ b/backend/Linearizeaux.ml
@@ -47,8 +47,12 @@ let join_points f =
reached_twice := IntSet.add npc !reached_twice
end else begin
reached := IntSet.add npc !reached;
- List.iter traverse (Kildall.successors_list succs pc)
+ traverse_succs (Kildall.successors_list succs pc)
end
+ and traverse_succs = function
+ | [] -> ()
+ | [pc] -> traverse pc
+ | pc :: l -> traverse pc; traverse_succs l
in traverse f.fn_entrypoint; !reached_twice
(* Cut into reachable basic blocks, annotated with the min value of the PC *)
@@ -73,20 +77,18 @@ let basic_blocks f joins =
let minpc = min npc minpc in
match PTree.get pc f.fn_code with
| None -> assert false
- | Some i ->
- match i with
- | Lnop s -> next_in_block blk minpc s
- | Lop (op, args, res, s) -> next_in_block blk minpc s
- | Lload (chunk, addr, args, dst, s) -> next_in_block blk minpc s
- | Lstore (chunk, addr, args, src, s) -> next_in_block blk minpc s
- | Lcall (sig0, ros, args, res, s) -> next_in_block blk minpc s
- | Ltailcall (sig0, ros, args) -> end_block blk minpc
- | Lbuiltin (ef, args, res, s) -> next_in_block blk minpc s
- | Lcond (cond, args, ifso, ifnot) ->
+ | Some b ->
+ let rec do_instr_list = function
+ | [] -> assert false
+ | Lbranch s :: _ -> next_in_block blk minpc s
+ | Ltailcall (sig0, ros) :: _ -> end_block blk minpc
+ | Lcond (cond, args, ifso, ifnot) :: _ ->
end_block blk minpc; start_block ifso; start_block ifnot
- | Ljumptable(arg, tbl) ->
+ | Ljumptable(arg, tbl) :: _ ->
end_block blk minpc; List.iter start_block tbl
- | Lreturn optarg -> end_block blk minpc
+ | Lreturn :: _ -> end_block blk minpc
+ | instr :: b' -> do_instr_list b' in
+ do_instr_list b
(* next_in_block: check if join point and either extend block
or start block *)
and next_in_block blk minpc pc =
diff --git a/backend/Linearizeproof.v b/backend/Linearizeproof.v
index d368311..2548580 100644
--- a/backend/Linearizeproof.v
+++ b/backend/Linearizeproof.v
@@ -12,10 +12,11 @@
(** Correctness proof for code linearization *)
+Require Import FSets.
Require Import Coqlib.
Require Import Maps.
Require Import Ordered.
-Require Import FSets.
+Require Import Lattice.
Require Import AST.
Require Import Integers.
Require Import Values.
@@ -27,17 +28,15 @@ Require Import Smallstep.
Require Import Op.
Require Import Locations.
Require Import LTL.
-Require Import LTLtyping.
-Require Import LTLin.
+Require Import Linear.
Require Import Linearize.
-Require Import Lattice.
Module NodesetFacts := FSetFacts.Facts(Nodeset).
Section LINEARIZATION.
Variable prog: LTL.program.
-Variable tprog: LTLin.program.
+Variable tprog: Linear.program.
Hypothesis TRANSF: transf_program prog = OK tprog.
@@ -70,7 +69,7 @@ Proof (Genv.find_var_info_transf_partial transf_fundef _ TRANSF).
Lemma sig_preserved:
forall f tf,
transf_fundef f = OK tf ->
- LTLin.funsig tf = LTL.funsig f.
+ Linear.funsig tf = LTL.funsig f.
Proof.
unfold transf_fundef, transf_partial_fundef; intros.
destruct f. monadInv H. monadInv EQ. reflexivity.
@@ -80,7 +79,7 @@ Qed.
Lemma stacksize_preserved:
forall f tf,
transf_function f = OK tf ->
- LTLin.fn_stacksize tf = LTL.fn_stacksize f.
+ Linear.fn_stacksize tf = LTL.fn_stacksize f.
Proof.
intros. monadInv H. auto.
Qed.
@@ -117,8 +116,8 @@ Qed.
(** The successors of a reachable instruction are reachable. *)
Lemma reachable_successors:
- forall f pc pc' i,
- f.(LTL.fn_code)!pc = Some i -> In pc' (successors_instr i) ->
+ forall f pc pc' b,
+ f.(LTL.fn_code)!pc = Some b -> In pc' (successors_block b) ->
(reachable f)!!pc = true ->
(reachable f)!!pc' = true.
Proof.
@@ -222,7 +221,7 @@ Qed.
(** * Properties related to labels *)
-(** If labels are globally unique and the LTLin code [c] contains
+(** If labels are globally unique and the Linear code [c] contains
a subsequence [Llabel lbl :: c1], then [find_label lbl c] returns [c1].
*)
@@ -284,13 +283,13 @@ Proof.
intros. unfold add_branch. destruct (starts_with s k); auto.
Qed.
-Lemma find_label_lin_instr:
+Lemma find_label_lin_block:
forall lbl k b,
- find_label lbl (linearize_instr b k) = find_label lbl k.
+ find_label lbl (linearize_block b k) = find_label lbl k.
Proof.
intros lbl k. generalize (find_label_add_branch lbl k); intro.
- induction b; simpl; auto.
- case (starts_with n k); simpl; auto.
+ induction b; simpl; auto. destruct a; simpl; auto.
+ case (starts_with s1 k); simpl; auto.
Qed.
Remark linearize_body_cons:
@@ -298,7 +297,7 @@ Remark linearize_body_cons:
linearize_body f (pc :: enum) =
match f.(LTL.fn_code)!pc with
| None => linearize_body f enum
- | Some b => Llabel pc :: linearize_instr b (linearize_body f enum)
+ | Some b => Llabel pc :: linearize_block b (linearize_body f enum)
end.
Proof.
intros. unfold linearize_body. rewrite list_fold_right_eq.
@@ -309,7 +308,7 @@ Lemma find_label_lin_rec:
forall f enum pc b,
In pc enum ->
f.(LTL.fn_code)!pc = Some b ->
- exists k, find_label pc (linearize_body f enum) = Some (linearize_instr b k).
+ exists k, find_label pc (linearize_body f enum) = Some (linearize_block b k).
Proof.
induction enum; intros.
elim H.
@@ -320,7 +319,7 @@ Proof.
assert (In pc enum). simpl in H. tauto.
destruct (IHenum pc b H1 H0) as [k FIND].
exists k. destruct (LTL.fn_code f)!a.
- simpl. rewrite peq_false. rewrite find_label_lin_instr. auto. auto.
+ simpl. rewrite peq_false. rewrite find_label_lin_block. auto. auto.
auto.
Qed.
@@ -330,7 +329,7 @@ Lemma find_label_lin:
f.(LTL.fn_code)!pc = Some b ->
(reachable f)!!pc = true ->
exists k,
- find_label pc (fn_code tf) = Some (linearize_instr b k).
+ find_label pc (fn_code tf) = Some (linearize_block b k).
Proof.
intros. monadInv H. simpl.
rewrite find_label_add_branch. apply find_label_lin_rec.
@@ -343,25 +342,12 @@ Lemma find_label_lin_inv:
f.(LTL.fn_code)!pc = Some b ->
(reachable f)!!pc = true ->
find_label pc (fn_code tf) = Some k ->
- exists k', k = linearize_instr b k'.
+ exists k', k = linearize_block b k'.
Proof.
intros. exploit find_label_lin; eauto. intros [k' FIND].
exists k'. congruence.
Qed.
-Lemma find_label_lin_succ:
- forall f tf s,
- transf_function f = OK tf ->
- valid_successor f s ->
- (reachable f)!!s = true ->
- exists k,
- find_label s (fn_code tf) = Some k.
-Proof.
- intros. destruct H0 as [i AT].
- exploit find_label_lin; eauto. intros [k FIND].
- exists (linearize_instr i k); auto.
-Qed.
-
(** Unique label property for linearized code. *)
Lemma label_in_add_branch:
@@ -372,16 +358,16 @@ Proof.
destruct (starts_with s k); simpl; intuition congruence.
Qed.
-Lemma label_in_lin_instr:
+Lemma label_in_lin_block:
forall lbl k b,
- In (Llabel lbl) (linearize_instr b k) -> In (Llabel lbl) k.
+ In (Llabel lbl) (linearize_block b k) -> In (Llabel lbl) k.
Proof.
- induction b; simpl; intros;
- try (apply label_in_add_branch with n; intuition congruence);
- try (intuition congruence).
- destruct (starts_with n k); simpl in H.
- apply label_in_add_branch with n; intuition congruence.
- apply label_in_add_branch with n0; intuition congruence.
+ induction b; simpl; intros. auto.
+ destruct a; simpl in H; try (intuition congruence).
+ apply label_in_add_branch with s; intuition congruence.
+ destruct (starts_with s1 k); simpl in H.
+ apply label_in_add_branch with s1; intuition congruence.
+ apply label_in_add_branch with s2; intuition congruence.
Qed.
Lemma label_in_lin_rec:
@@ -392,7 +378,7 @@ Proof.
simpl; auto.
rewrite linearize_body_cons. destruct (LTL.fn_code f)!a.
simpl. intros [A|B]. left; congruence.
- right. apply IHenum. eapply label_in_lin_instr; eauto.
+ right. apply IHenum. eapply label_in_lin_block; eauto.
intro; right; auto.
Qed.
@@ -404,12 +390,13 @@ Proof.
destruct (starts_with lbl k); simpl; intuition.
Qed.
-Lemma unique_labels_lin_instr:
+Lemma unique_labels_lin_block:
forall k b,
- unique_labels k -> unique_labels (linearize_instr b k).
+ unique_labels k -> unique_labels (linearize_block b k).
Proof.
- induction b; intro; simpl; auto; try (apply unique_labels_add_branch; auto).
- case (starts_with n k); simpl; apply unique_labels_add_branch; auto.
+ induction b; intros; simpl. auto.
+ destruct a; auto; try (apply unique_labels_add_branch; auto).
+ case (starts_with s1 k); simpl; apply unique_labels_add_branch; auto.
Qed.
Lemma unique_labels_lin_rec:
@@ -423,8 +410,8 @@ Proof.
intro. destruct (LTL.fn_code f)!a.
simpl. split. red. intro. inversion H. elim H3.
apply label_in_lin_rec with f.
- apply label_in_lin_instr with i. auto.
- apply unique_labels_lin_instr. apply IHenum. inversion H; auto.
+ apply label_in_lin_block with b. auto.
+ apply unique_labels_lin_block. apply IHenum. inversion H; auto.
apply IHenum. inversion H; auto.
Qed.
@@ -458,6 +445,17 @@ Proof.
auto. eauto with coqlib.
Qed.
+Lemma is_tail_lin_block:
+ forall b c1 c2,
+ is_tail (linearize_block b c1) c2 -> is_tail c1 c2.
+Proof.
+ induction b; simpl; intros.
+ auto.
+ destruct a; eauto with coqlib.
+ eapply is_tail_add_branch; eauto.
+ destruct (starts_with s1 c1); eapply is_tail_add_branch; eauto with coqlib.
+Qed.
+
Lemma add_branch_correct:
forall lbl c k s f tf sp ls m,
transf_function f = OK tf ->
@@ -475,12 +473,11 @@ Qed.
(** * Correctness of linearization *)
-(** The proof of semantic preservation is a simulation argument
- based on diagrams of the following form:
+(** The proof of semantic preservation is a simulation argument of the "star" kind:
<<
st1 --------------- st2
| |
- t| +|t
+ t| t| + or ( 0 \/ |st1'| < |st1| )
| |
v v
st1'--------------- st2'
@@ -492,273 +489,260 @@ Qed.
control-flow graph, the transformed state is at a code
sequence [c] that starts with the label [pc]. *)
-Inductive match_stackframes: LTL.stackframe -> LTLin.stackframe -> Prop :=
+Inductive match_stackframes: LTL.stackframe -> Linear.stackframe -> Prop :=
| match_stackframe_intro:
- forall res f sp pc ls tf c,
+ forall f sp bb ls tf c,
transf_function f = OK tf ->
- (reachable f)!!pc = true ->
- valid_successor f pc ->
- is_tail c (fn_code tf) ->
- wt_function f ->
+ (forall pc, In pc (successors_block bb) -> (reachable f)!!pc = true) ->
+ is_tail c tf.(fn_code) ->
match_stackframes
- (LTL.Stackframe res f sp ls pc)
- (LTLin.Stackframe res tf sp ls (add_branch pc c)).
+ (LTL.Stackframe f sp ls bb)
+ (Linear.Stackframe tf sp ls (linearize_block bb c)).
-Inductive match_states: LTL.state -> LTLin.state -> Prop :=
- | match_states_intro:
+Inductive match_states: LTL.state -> Linear.state -> Prop :=
+ | match_states_add_branch:
forall s f sp pc ls m tf ts c
(STACKS: list_forall2 match_stackframes s ts)
(TRF: transf_function f = OK tf)
(REACH: (reachable f)!!pc = true)
- (AT: find_label pc (fn_code tf) = Some c)
- (WTF: wt_function f),
+ (TAIL: is_tail c tf.(fn_code)),
match_states (LTL.State s f sp pc ls m)
- (LTLin.State ts tf sp c ls m)
+ (Linear.State ts tf sp (add_branch pc c) ls m)
+ | match_states_cond_taken:
+ forall s f sp pc ls m tf ts cond args c
+ (STACKS: list_forall2 match_stackframes s ts)
+ (TRF: transf_function f = OK tf)
+ (REACH: (reachable f)!!pc = true)
+ (JUMP: eval_condition cond (reglist ls args) m = Some true),
+ match_states (LTL.State s f sp pc (undef_regs (destroyed_by_cond cond) ls) m)
+ (Linear.State ts tf sp (Lcond cond args pc :: c) ls m)
+ | match_states_jumptable:
+ forall s f sp pc ls m tf ts arg tbl c n
+ (STACKS: list_forall2 match_stackframes s ts)
+ (TRF: transf_function f = OK tf)
+ (REACH: (reachable f)!!pc = true)
+ (ARG: ls (R arg) = Vint n)
+ (JUMP: list_nth_z tbl (Int.unsigned n) = Some pc),
+ match_states (LTL.State s f sp pc (undef_regs destroyed_by_jumptable ls) m)
+ (Linear.State ts tf sp (Ljumptable arg tbl :: c) ls m)
+ | match_states_block:
+ forall s f sp bb ls m tf ts c
+ (STACKS: list_forall2 match_stackframes s ts)
+ (TRF: transf_function f = OK tf)
+ (REACH: forall pc, In pc (successors_block bb) -> (reachable f)!!pc = true)
+ (TAIL: is_tail c tf.(fn_code)),
+ match_states (LTL.Block s f sp bb ls m)
+ (Linear.State ts tf sp (linearize_block bb c) ls m)
| match_states_call:
forall s f ls m tf ts,
list_forall2 match_stackframes s ts ->
transf_fundef f = OK tf ->
- wt_fundef f ->
match_states (LTL.Callstate s f ls m)
- (LTLin.Callstate ts tf ls m)
+ (Linear.Callstate ts tf ls m)
| match_states_return:
forall s ls m ts,
list_forall2 match_stackframes s ts ->
match_states (LTL.Returnstate s ls m)
- (LTLin.Returnstate ts ls m).
+ (Linear.Returnstate ts ls m).
-Hypothesis wt_prog: wt_program prog.
+Definition measure (S: LTL.state) : nat :=
+ match S with
+ | LTL.State s f sp pc ls m => 0%nat
+ | LTL.Block s f sp bb ls m => 1%nat
+ | _ => 0%nat
+ end.
+
+Remark match_parent_locset:
+ forall s ts, list_forall2 match_stackframes s ts -> parent_locset ts = LTL.parent_locset s.
+Proof.
+ induction 1; simpl. auto. inv H; auto.
+Qed.
Theorem transf_step_correct:
forall s1 t s2, LTL.step ge s1 t s2 ->
forall s1' (MS: match_states s1 s1'),
- exists s2', plus LTLin.step tge s1' t s2' /\ match_states s2 s2'.
-Proof.
- induction 1; intros; try (inv MS);
- try (generalize (wt_instrs _ WTF _ _ H); intro WTI).
- (* Lnop *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; auto. intros [c'' AT'].
- econstructor; split.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_find_label. eauto.
+ (exists s2', plus Linear.step tge s1' t s2' /\ match_states s2 s2')
+ \/ (measure s2 < measure s1 /\ t = E0 /\ match_states s2 s1')%nat.
+Proof.
+ induction 1; intros; try (inv MS).
+
+ (* start of block, at an [add_branch] *)
+ exploit find_label_lin; eauto. intros [k F].
+ left; econstructor; split.
+ eapply add_branch_correct; eauto.
+ econstructor; eauto.
+ intros; eapply reachable_successors; eauto.
+ eapply is_tail_lin_block; eauto. eapply is_tail_find_label; eauto.
+
+ (* start of block, target of an [Lcond] *)
+ exploit find_label_lin; eauto. intros [k F].
+ left; econstructor; split.
+ apply plus_one. eapply exec_Lcond_true; eauto.
+ econstructor; eauto.
+ intros; eapply reachable_successors; eauto.
+ eapply is_tail_lin_block; eauto. eapply is_tail_find_label; eauto.
+
+ (* start of block, target of an [Ljumptable] *)
+ exploit find_label_lin; eauto. intros [k F].
+ left; econstructor; split.
+ apply plus_one. eapply exec_Ljumptable; eauto.
econstructor; eauto.
+ intros; eapply reachable_successors; eauto.
+ eapply is_tail_lin_block; eauto. eapply is_tail_find_label; eauto.
+
(* Lop *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; auto. intros [c'' AT'].
- econstructor; split.
- eapply plus_left'.
- eapply exec_Lop with (v := v); eauto.
- rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- traceEq.
- econstructor; eauto.
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor; eauto.
+ instantiate (1 := v); rewrite <- H; apply eval_operation_preserved.
+ exact symbols_preserved.
+ econstructor; eauto.
+
(* Lload *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; auto. intros [c'' AT'].
- econstructor; split.
- eapply plus_left'.
- apply exec_Lload with a.
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
- eauto.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- traceEq.
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor.
+ instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
+ exact symbols_preserved. eauto. eauto.
+ econstructor; eauto.
+
+ (* Lgetstack *)
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor; eauto.
+ econstructor; eauto.
+
+ (* Lsetstack *)
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor; eauto.
econstructor; eauto.
+
(* Lstore *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; auto. intros [c'' AT'].
- econstructor; split.
- eapply plus_left'.
- apply exec_Lstore with a.
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
- eauto.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- traceEq.
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor.
+ instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
+ exact symbols_preserved. eauto. eauto.
econstructor; eauto.
+
(* Lcall *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl; auto.
- assert (VALID: valid_successor f pc'). inv WTI; auto.
- exploit find_function_translated; eauto. intros [tf' [A B]].
- econstructor; split.
- apply plus_one. eapply exec_Lcall with (f' := tf'); eauto.
- symmetry; apply sig_preserved; auto.
- econstructor; eauto.
- constructor; auto. econstructor; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- destruct ros; simpl in H0.
- eapply Genv.find_funct_prop; eauto.
- destruct (Genv.find_symbol ge i); try discriminate.
- eapply Genv.find_funct_ptr_prop; eauto.
+ exploit find_function_translated; eauto. intros [tfd [A B]].
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor; eauto.
+ symmetry; eapply sig_preserved; eauto.
+ econstructor; eauto. constructor; auto. econstructor; eauto.
(* Ltailcall *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- exploit find_function_translated; eauto. intros [tf' [A B]].
- econstructor; split.
- apply plus_one. eapply exec_Ltailcall with (f' := tf'); eauto.
- symmetry; apply sig_preserved; auto.
- rewrite (stacksize_preserved _ _ TRF). eauto.
+ exploit find_function_translated; eauto. intros [tfd [A B]].
+ left; econstructor; split. simpl.
+ apply plus_one. econstructor; eauto.
+ rewrite (match_parent_locset _ _ STACKS). eauto.
+ symmetry; eapply sig_preserved; eauto.
+ rewrite (stacksize_preserved _ _ TRF); eauto.
+ rewrite (match_parent_locset _ _ STACKS).
econstructor; eauto.
- destruct ros; simpl in H0.
- eapply Genv.find_funct_prop; eauto.
- destruct (Genv.find_symbol ge i); try discriminate.
- eapply Genv.find_funct_ptr_prop; eauto.
(* Lbuiltin *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; auto. intros [c'' AT'].
- econstructor; split.
- eapply plus_left'.
- eapply exec_Lbuiltin.
- eapply external_call_symbols_preserved; eauto.
+ left; econstructor; split. simpl.
+ apply plus_one. eapply exec_Lbuiltin; eauto.
+ eapply external_call_symbols_preserved'; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ econstructor; eauto.
+
+ (* Lannot *)
+ left; econstructor; split. simpl.
+ apply plus_one. eapply exec_Lannot; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- traceEq.
econstructor; eauto.
+ (* Lbranch *)
+ assert ((reachable f)!!pc = true). apply REACH; simpl; auto.
+ right; split. simpl; omega. split. auto. simpl. econstructor; eauto.
+
(* Lcond *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
+ assert (REACH1: (reachable f)!!pc1 = true) by (apply REACH; simpl; auto).
+ assert (REACH2: (reachable f)!!pc2 = true) by (apply REACH; simpl; auto).
+ simpl linearize_block.
+ destruct (starts_with pc1 c).
+ (* branch if cond is false *)
+ assert (DC: destroyed_by_cond (negate_condition cond) = destroyed_by_cond cond).
+ destruct cond; reflexivity.
destruct b.
- (* true *)
- assert (REACH': (reachable f)!!ifso = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; eauto. intros [c'' AT'].
- destruct (starts_with ifso c').
- econstructor; split.
- eapply plus_left'.
- eapply exec_Lcond_false; eauto.
- rewrite eval_negate_condition; rewrite H0; auto.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- traceEq.
- econstructor; eauto.
- econstructor; split.
- apply plus_one. eapply exec_Lcond_true; eauto.
- econstructor; eauto.
- (* false *)
- assert (REACH': (reachable f)!!ifnot = true).
- eapply reachable_successors; eauto. simpl; auto.
- exploit find_label_lin_succ; eauto. inv WTI; auto. intros [c'' AT'].
- destruct (starts_with ifso c').
- econstructor; split.
- apply plus_one. eapply exec_Lcond_true; eauto.
- rewrite eval_negate_condition; rewrite H0; auto.
- econstructor; eauto.
- econstructor; split.
- eapply plus_left'.
- eapply exec_Lcond_false; eauto.
- eapply add_branch_correct; eauto.
- eapply is_tail_add_branch. eapply is_tail_cons_left.
- eapply is_tail_find_label. eauto.
- traceEq.
+ (* cond is true: no branch *)
+ left; econstructor; split.
+ apply plus_one. eapply exec_Lcond_false.
+ rewrite eval_negate_condition. rewrite H. auto. eauto.
+ rewrite DC. econstructor; eauto.
+ (* cond is false: branch is taken *)
+ right; split. simpl; omega. split. auto. rewrite <- DC. econstructor; eauto.
+ rewrite eval_negate_condition. rewrite H. auto.
+ (* branch if cond is true *)
+ destruct b.
+ (* cond is true: branch is taken *)
+ right; split. simpl; omega. split. auto. econstructor; eauto.
+ (* cond is false: no branch *)
+ left; econstructor; split.
+ apply plus_one. eapply exec_Lcond_false. eauto. eauto.
econstructor; eauto.
(* Ljumptable *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- assert (REACH': (reachable f)!!pc' = true).
- eapply reachable_successors; eauto. simpl. eapply list_nth_z_in; eauto.
- exploit find_label_lin_succ; eauto.
- inv WTI. apply H6. eapply list_nth_z_in; eauto.
- intros [c'' AT'].
- econstructor; split.
- apply plus_one. eapply exec_Ljumptable; eauto.
- econstructor; eauto.
+ assert (REACH': (reachable f)!!pc = true).
+ apply REACH. simpl. eapply list_nth_z_in; eauto.
+ right; split. simpl; omega. split. auto. econstructor; eauto.
(* Lreturn *)
- destruct (find_label_lin_inv _ _ _ _ _ TRF H REACH AT) as [c' EQ].
- simpl in EQ. subst c.
- econstructor; split.
- apply plus_one. eapply exec_Lreturn; eauto.
+ left; econstructor; split.
+ simpl. apply plus_one. econstructor; eauto.
rewrite (stacksize_preserved _ _ TRF). eauto.
- econstructor; eauto.
-
- (* internal function *)
+ rewrite (match_parent_locset _ _ STACKS). econstructor; eauto.
+
+ (* internal functions *)
assert (REACH: (reachable f)!!(LTL.fn_entrypoint f) = true).
apply reachable_entrypoint.
- inv H7. monadInv H6.
- exploit find_label_lin_succ; eauto. inv H1; auto. intros [c'' AT'].
- generalize EQ; intro. monadInv EQ0. econstructor; simpl; split.
- eapply plus_left'.
- eapply exec_function_internal; eauto.
- simpl. eapply add_branch_correct. eauto.
- simpl. eapply is_tail_add_branch. constructor. eauto.
- traceEq.
- econstructor; eauto.
+ monadInv H7.
+ left; econstructor; split.
+ apply plus_one. eapply exec_function_internal; eauto.
+ rewrite (stacksize_preserved _ _ EQ). eauto.
+ generalize EQ; intro EQ'; monadInv EQ'. simpl.
+ econstructor; eauto. simpl. eapply is_tail_add_branch. constructor.
(* external function *)
- monadInv H6. econstructor; split.
+ monadInv H8. left; econstructor; split.
apply plus_one. eapply exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto.
(* return *)
inv H3. inv H1.
- exploit find_label_lin_succ; eauto. intros [c' AT].
- econstructor; split.
- eapply plus_left'.
- eapply exec_return; eauto.
- eapply add_branch_correct; eauto. traceEq.
- econstructor; eauto.
+ left; econstructor; split.
+ apply plus_one. econstructor.
+ econstructor; eauto.
Qed.
Lemma transf_initial_states:
forall st1, LTL.initial_state prog st1 ->
- exists st2, LTLin.initial_state tprog st2 /\ match_states st1 st2.
+ exists st2, Linear.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.
+ exists (Callstate nil tf (Locmap.init Vundef) m0); split.
econstructor; eauto. eapply Genv.init_mem_transf_partial; eauto.
replace (prog_main tprog) with (prog_main prog).
rewrite symbols_preserved. eauto.
symmetry. apply (transform_partial_program_main transf_fundef _ TRANSF).
rewrite <- H3. apply sig_preserved. auto.
constructor. constructor. auto.
- eapply Genv.find_funct_ptr_prop; eauto.
Qed.
Lemma transf_final_states:
forall st1 st2 r,
- match_states st1 st2 -> LTL.final_state st1 r -> LTLin.final_state st2 r.
+ match_states st1 st2 -> LTL.final_state st1 r -> Linear.final_state st2 r.
Proof.
- intros. inv H0. inv H. inv H4. constructor.
+ intros. inv H0. inv H. inv H6. econstructor; eauto.
Qed.
Theorem transf_program_correct:
- forward_simulation (LTL.semantics prog) (LTLin.semantics tprog).
+ forward_simulation (LTL.semantics prog) (Linear.semantics tprog).
Proof.
- eapply forward_simulation_plus.
+ eapply forward_simulation_star.
eexact symbols_preserved.
eexact transf_initial_states.
eexact transf_final_states.
diff --git a/backend/Linearizetyping.v b/backend/Linearizetyping.v
deleted file mode 100644
index b4e25de..0000000
--- a/backend/Linearizetyping.v
+++ /dev/null
@@ -1,112 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Type preservation for the Linearize pass *)
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import Errors.
-Require Import AST.
-Require Import Op.
-Require Import Locations.
-Require Import LTL.
-Require Import LTLtyping.
-Require Import LTLin.
-Require Import Linearize.
-Require Import LTLintyping.
-Require Import Conventions.
-
-(** * Type preservation for the linearization pass *)
-
-Lemma wt_add_instr:
- forall f i k, wt_instr f i -> wt_code f k -> wt_code f (i :: k).
-Proof.
- intros; red; intros. elim H1; intro. subst i0; auto. auto.
-Qed.
-
-Lemma wt_add_branch:
- forall f s k, wt_code f k -> wt_code f (add_branch s k).
-Proof.
- intros; unfold add_branch. destruct (starts_with s k).
- auto. apply wt_add_instr; auto. constructor.
-Qed.
-
-Lemma wt_linearize_instr:
- forall f instr,
- LTLtyping.wt_instr f instr ->
- forall k,
- wt_code f.(LTL.fn_sig) k ->
- wt_code f.(LTL.fn_sig) (linearize_instr instr k).
-Proof.
- induction 1; simpl; intros.
- apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. auto.
- apply wt_add_instr. constructor; auto. apply wt_add_branch; auto.
- destruct (starts_with s1 k); apply wt_add_instr.
- constructor; auto. rewrite H. destruct cond; auto.
- apply wt_add_branch; auto.
- constructor; auto. apply wt_add_branch; auto.
- apply wt_add_instr. constructor; auto. auto.
- apply wt_add_instr. constructor; auto. auto.
-Qed.
-
-Lemma wt_linearize_body:
- forall f,
- (forall pc instr,
- f.(LTL.fn_code)!pc = Some instr -> LTLtyping.wt_instr f instr) ->
- forall enum,
- wt_code f.(LTL.fn_sig) (linearize_body f enum).
-Proof.
- unfold linearize_body; induction enum; rewrite list_fold_right_eq.
- red; simpl; intros; contradiction.
- unfold linearize_node. caseEq ((LTL.fn_code f)!a); intros.
- apply wt_add_instr. constructor. apply wt_linearize_instr; eauto with coqlib.
- auto.
-Qed.
-
-Lemma wt_transf_function:
- forall f tf,
- LTLtyping.wt_function f ->
- transf_function f = OK tf ->
- wt_function tf.
-Proof.
- intros. inv H. monadInv H0. constructor; auto.
- simpl. apply wt_add_branch.
- apply wt_linearize_body. auto.
-Qed.
-
-Lemma wt_transf_fundef:
- forall f tf,
- LTLtyping.wt_fundef f ->
- transf_fundef f = OK tf ->
- wt_fundef tf.
-Proof.
- induction 1; intros. monadInv H. constructor.
- monadInv H0. constructor; eapply wt_transf_function; eauto.
-Qed.
-
-Lemma program_typing_preserved:
- forall (p: LTL.program) (tp: LTLin.program),
- LTLtyping.wt_program p ->
- transf_program p = OK tp ->
- LTLintyping.wt_program tp.
-Proof.
- intros; red; intros.
- generalize (transform_partial_program_function transf_fundef _ _ _ H0 H1).
- intros [f0 [IN TR]].
- eapply wt_transf_fundef; eauto.
-Qed.
diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v
index 32d6045..c51db6f 100644
--- a/backend/Lineartyping.v
+++ b/backend/Lineartyping.v
@@ -16,124 +16,87 @@ Require Import Coqlib.
Require Import AST.
Require Import Integers.
Require Import Values.
+Require Import Events.
Require Import Op.
Require Import Locations.
+Require Import Conventions.
Require Import LTL.
Require Import Linear.
-Require Import Conventions.
-(** The typing rules for Linear are similar to those for LTLin: we check
- that instructions receive the right number of arguments,
- and that the types of the argument and result registers agree
- with what the instruction expects. Moreover, we also
- enforces some correctness conditions on the offsets of stack slots
- accessed through [Lgetstack] and [Lsetstack] Linear instructions. *)
+(** The typing rules for Linear enforce several properties useful for
+ the proof of the [Stacking] pass:
+- for each instruction, the type of the result register or slot
+ agrees with the type of values produced by the instruction;
+- correctness conditions on the offsets of stack slots
+ accessed through [Lgetstack] and [Lsetstack] Linear instructions.
+*)
+
+(** The rules are presented as boolean-valued functions so that we
+ get an executable type-checker for free. *)
Section WT_INSTR.
Variable funct: function.
-Definition slot_valid (s: slot) :=
- match s with
- | Local ofs ty => 0 <= ofs
- | Outgoing ofs ty => 0 <= ofs
- | Incoming ofs ty => In (S s) (loc_parameters funct.(fn_sig))
+Definition slot_valid (sl: slot) (ofs: Z) (ty: typ): bool :=
+ match sl with
+ | Local => zle 0 ofs
+ | Outgoing => zle 0 ofs
+ | Incoming => In_dec Loc.eq (S Incoming ofs ty) (loc_parameters funct.(fn_sig))
+ end &&
+ match ty with
+ | Tint | Tfloat => true
+ | Tlong => false
end.
-Definition slot_writable (s: slot) :=
- match s with
- | Local ofs ty => True
- | Outgoing ofs ty => True
- | Incoming ofs ty => False
+Definition slot_writable (sl: slot) : bool :=
+ match sl with
+ | Local => true
+ | Outgoing => true
+ | Incoming => false
end.
-Inductive wt_instr : instruction -> Prop :=
- | wt_Lgetstack:
- forall s r,
- slot_type s = mreg_type r ->
- slot_valid s ->
- wt_instr (Lgetstack s r)
- | wt_Lsetstack:
- forall s r,
- slot_type s = mreg_type r ->
- slot_valid s -> slot_writable s ->
- wt_instr (Lsetstack r s)
- | wt_Lopmove:
- forall r1 r,
- mreg_type r1 = mreg_type r ->
- wt_instr (Lop Omove (r1 :: nil) r)
- | wt_Lop:
- forall op args res,
- op <> Omove ->
- (List.map mreg_type args, mreg_type res) = type_of_operation op ->
- wt_instr (Lop op args res)
- | wt_Lload:
- forall chunk addr args dst,
- List.map mreg_type args = type_of_addressing addr ->
- mreg_type dst = type_of_chunk chunk ->
- wt_instr (Lload chunk addr args dst)
- | wt_Lstore:
- forall chunk addr args src,
- List.map mreg_type args = type_of_addressing addr ->
- mreg_type src = type_of_chunk chunk ->
- wt_instr (Lstore chunk addr args src)
- | wt_Lcall:
- forall sig ros,
- match ros with inl r => mreg_type r = Tint | _ => True end ->
- wt_instr (Lcall sig ros)
- | wt_Ltailcall:
- forall sig ros,
- tailcall_possible sig ->
- match ros with inl r => r = IT1 | _ => True end ->
- wt_instr (Ltailcall sig ros)
- | wt_Lbuiltin:
- forall ef args res,
- List.map mreg_type args = (ef_sig ef).(sig_args) ->
- mreg_type res = proj_sig_res (ef_sig ef) ->
- arity_ok (ef_sig ef).(sig_args) = true ->
- wt_instr (Lbuiltin ef args res)
- | wt_Lannot:
- forall ef args,
- List.map Loc.type args = (ef_sig ef).(sig_args) ->
- ef_reloads ef = false ->
- locs_acceptable args ->
- wt_instr (Lannot ef args)
- | wt_Llabel:
- forall lbl,
- wt_instr (Llabel lbl)
- | wt_Lgoto:
- forall lbl,
- wt_instr (Lgoto lbl)
- | wt_Lcond:
- forall cond args lbl,
- List.map mreg_type args = type_of_condition cond ->
- wt_instr (Lcond cond args lbl)
- | wt_Ljumptable:
- forall arg tbl,
- mreg_type arg = Tint ->
- list_length_z tbl * 4 <= Int.max_unsigned ->
- wt_instr (Ljumptable arg tbl)
- | wt_Lreturn:
- wt_instr (Lreturn).
+Definition loc_valid (l: loc) : bool :=
+ match l with
+ | R r => true
+ | S Local ofs ty => slot_valid Local ofs ty
+ | S _ _ _ => false
+ end.
+
+Definition wt_instr (i: instruction) : bool :=
+ match i with
+ | Lgetstack sl ofs ty r =>
+ typ_eq ty (mreg_type r) && slot_valid sl ofs ty
+ | Lsetstack r sl ofs ty =>
+ typ_eq ty (mreg_type r) && slot_valid sl ofs ty && slot_writable sl
+ | Lop op args res =>
+ match is_move_operation op args with
+ | Some arg =>
+ typ_eq (mreg_type arg) (mreg_type res)
+ | None =>
+ let (targs, tres) := type_of_operation op in
+ typ_eq (mreg_type res) tres
+ end
+ | Lload chunk addr args dst =>
+ typ_eq (mreg_type dst) (type_of_chunk chunk)
+ | Ltailcall sg ros =>
+ zeq (size_arguments sg) 0
+ | Lbuiltin ef args res =>
+ list_typ_eq (map mreg_type res) (proj_sig_res' (ef_sig ef))
+ | Lannot ef args =>
+ forallb loc_valid args
+ | _ =>
+ true
+ end.
End WT_INSTR.
-Definition wt_code (f: function) (c: code) : Prop :=
- forall instr, In instr c -> wt_instr f instr.
+Definition wt_code (f: function) (c: code) : bool :=
+ forallb (wt_instr f) c.
-Definition wt_function (f: function) : Prop :=
+Definition wt_function (f: function) : bool :=
wt_code f f.(fn_code).
-Inductive wt_fundef: fundef -> Prop :=
- | wt_fundef_external: forall ef,
- wt_fundef (External ef)
- | wt_function_internal: forall f,
- wt_function f ->
- wt_fundef (Internal f).
-
-Definition wt_program (p: program) : Prop :=
- forall i f, In (i, Gfun f) (prog_defs p) -> wt_fundef f.
-
(** Typing the run-time state. These definitions are used in [Stackingproof]. *)
Require Import Values.
@@ -148,48 +111,30 @@ Proof.
intros; red; intros.
unfold Locmap.set.
destruct (Loc.eq l l0). congruence.
- destruct (Loc.overlap l l0). red. auto.
- auto.
-Qed.
-
-Lemma wt_undef_locs:
- forall locs ls, wt_locset ls -> wt_locset (Locmap.undef locs ls).
-Proof.
- induction locs; simpl; intros. auto. apply IHlocs. apply wt_setloc; auto. red; auto.
+ destruct (Loc.diff_dec l l0). auto. red. auto.
Qed.
-Lemma wt_undef_temps:
- forall ls, wt_locset ls -> wt_locset (undef_temps ls).
+Lemma wt_setlocs:
+ forall ll vl ls,
+ Val.has_type_list vl (map Loc.type ll) -> wt_locset ls -> wt_locset (Locmap.setlist ll vl ls).
Proof.
- intros; unfold undef_temps. apply wt_undef_locs; auto.
+ induction ll; destruct vl; simpl; intuition.
+ apply IHll; auto. apply wt_setloc; auto.
Qed.
-Lemma wt_undef_op:
- forall op ls, wt_locset ls -> wt_locset (undef_op op ls).
+Lemma wt_undef_regs:
+ forall rs ls, wt_locset ls -> wt_locset (undef_regs rs ls).
Proof.
- intros. generalize (wt_undef_temps ls H); intro.
- unfold undef_op; destruct op; auto; apply wt_undef_locs; auto.
-Qed.
-
-Lemma wt_undef_getstack:
- forall s ls, wt_locset ls -> wt_locset (undef_getstack s ls).
-Proof.
- intros. unfold undef_getstack. destruct s; auto. apply wt_setloc; auto. red; auto.
-Qed.
-
-Lemma wt_undef_setstack:
- forall ls, wt_locset ls -> wt_locset (undef_setstack ls).
-Proof.
- intros. unfold undef_setstack. apply wt_undef_locs; auto.
+ induction rs; simpl; intros. auto. apply wt_setloc; auto. red; auto.
Qed.
Lemma wt_call_regs:
forall ls, wt_locset ls -> wt_locset (call_regs ls).
Proof.
- intros; red; intros. unfold call_regs. destruct l. auto.
- destruct (in_dec Loc.eq (R m) temporaries). red; auto. auto.
- destruct s. red; auto.
- change (Loc.type (S (Incoming z t))) with (Loc.type (S (Outgoing z t))). auto.
+ intros; red; intros. unfold call_regs. destruct l. auto.
+ destruct sl.
+ red; auto.
+ change (Loc.type (S Incoming pos ty)) with (Loc.type (S Outgoing pos ty)). auto.
red; auto.
Qed.
@@ -198,9 +143,8 @@ Lemma wt_return_regs:
wt_locset caller -> wt_locset callee -> wt_locset (return_regs caller callee).
Proof.
intros; red; intros.
- unfold return_regs. destruct l; auto.
- destruct (in_dec Loc.eq (R m) temporaries); auto.
- destruct (in_dec Loc.eq (R m) destroyed_at_call); auto.
+ unfold return_regs. destruct l; auto.
+ destruct (in_dec mreg_eq r destroyed_at_call); auto.
Qed.
Lemma wt_init:
@@ -208,3 +152,19 @@ Lemma wt_init:
Proof.
red; intros. unfold Locmap.init. red; auto.
Qed.
+
+Lemma wt_setlist_result:
+ forall sg v rs,
+ Val.has_type v (proj_sig_res sg) ->
+ wt_locset rs ->
+ wt_locset (Locmap.setlist (map R (loc_result sg)) (encode_long (sig_res sg) v) rs).
+Proof.
+ unfold loc_result, encode_long, proj_sig_res; intros.
+ destruct (sig_res sg) as [[] | ]; simpl.
+ apply wt_setloc; auto.
+ apply wt_setloc; auto.
+ destruct v; simpl in H; try contradiction;
+ simpl; apply wt_setloc; auto; apply wt_setloc; auto.
+ apply wt_setloc; auto.
+Qed.
+
diff --git a/backend/Locations.v b/backend/Locations.v
index 2381fea..2f2dae8 100644
--- a/backend/Locations.v
+++ b/backend/Locations.v
@@ -13,7 +13,10 @@
(** Locations are a refinement of RTL pseudo-registers, used to reflect
the results of register allocation (file [Allocation]). *)
+Require Import OrderedType.
Require Import Coqlib.
+Require Import Maps.
+Require Import Ordered.
Require Import AST.
Require Import Values.
Require Export Machregs.
@@ -42,9 +45,9 @@ Require Export Machregs.
calling conventions. *)
Inductive slot: Type :=
- | Local: Z -> typ -> slot
- | Incoming: Z -> typ -> slot
- | Outgoing: Z -> typ -> slot.
+ | Local
+ | Incoming
+ | Outgoing.
(** Morally, the [Incoming] slots of a function are the [Outgoing]
slots of its caller function.
@@ -56,46 +59,17 @@ as 32-bit integers/pointers (type [Tint]) or as 64-bit floats (type [Tfloat]).
The offset of a slot, combined with its type and its kind, identifies
uniquely the slot and will determine later where it resides within the
memory-allocated activation record. Offsets are always positive.
-
-Conceptually, slots will be mapped to four non-overlapping memory areas
-within activation records:
-- The area for [Local] slots of type [Tint]. The offset is interpreted
- as a 4-byte word index.
-- The area for [Local] slots of type [Tfloat]. The offset is interpreted
- as a 8-byte word index. Thus, two [Local] slots always refer either
- to the same memory chunk (if they have the same types and offsets)
- or to non-overlapping memory chunks (if the types or offsets differ).
-- The area for [Outgoing] slots. The offset is a 4-byte word index.
- Unlike [Local] slots, the PowerPC calling conventions demand that
- integer and float [Outgoing] slots reside in the same memory area.
- Therefore, [Outgoing Tint 0] and [Outgoing Tfloat 0] refer to
- overlapping memory chunks and cannot be used simultaneously: one will
- lose its value when the other is assigned. We will reflect this
- overlapping behaviour in the environments mapping locations to values
- defined later in this file.
-- The area for [Incoming] slots. Same structure as the [Outgoing] slots.
*)
-Definition slot_type (s: slot): typ :=
- match s with
- | Local ofs ty => ty
- | Incoming ofs ty => ty
- | Outgoing ofs ty => ty
- end.
-
Lemma slot_eq: forall (p q: slot), {p = q} + {p <> q}.
Proof.
- assert (typ_eq: forall (t1 t2: typ), {t1 = t2} + {t1 <> t2}).
- decide equality.
- generalize zeq; intro.
decide equality.
Defined.
-Global Opaque slot_eq.
Open Scope Z_scope.
Definition typesize (ty: typ) : Z :=
- match ty with Tint => 1 | Tfloat => 2 end.
+ match ty with Tint => 1 | Tlong => 2 | Tfloat => 2 end.
Lemma typesize_pos:
forall (ty: typ), typesize ty > 0.
@@ -109,20 +83,24 @@ Qed.
activation record slots. *)
Inductive loc : Type :=
- | R: mreg -> loc
- | S: slot -> loc.
+ | R (r: mreg)
+ | S (sl: slot) (pos: Z) (ty: typ).
Module Loc.
Definition type (l: loc) : typ :=
match l with
| R r => mreg_type r
- | S s => slot_type s
+ | S sl pos ty => ty
end.
Lemma eq: forall (p q: loc), {p = q} + {p <> q}.
Proof.
- decide equality. apply mreg_eq. apply slot_eq.
+ decide equality.
+ apply mreg_eq.
+ apply typ_eq.
+ apply zeq.
+ apply slot_eq.
Defined.
(** As mentioned previously, two locations can be different (in the sense
@@ -138,13 +116,10 @@ Module Loc.
*)
Definition diff (l1 l2: loc) : Prop :=
match l1, l2 with
- | R r1, R r2 => r1 <> r2
- | S (Local d1 t1), S (Local d2 t2) =>
- d1 <> d2 \/ t1 <> t2
- | S (Incoming d1 t1), S (Incoming d2 t2) =>
- d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1
- | S (Outgoing d1 t1), S (Outgoing d2 t2) =>
- d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1
+ | R r1, R r2 =>
+ r1 <> r2
+ | S s1 d1 t1, S s2 d2 t2 =>
+ s1 <> s2 \/ d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1
| _, _ =>
True
end.
@@ -152,11 +127,8 @@ Module Loc.
Lemma same_not_diff:
forall l, ~(diff l l).
Proof.
- destruct l; unfold diff; try tauto.
- destruct s.
- tauto.
- generalize (typesize_pos t); omega.
- generalize (typesize_pos t); omega.
+ destruct l; unfold diff; auto.
+ red; intros. destruct H; auto. generalize (typesize_pos ty); omega.
Qed.
Lemma diff_not_eq:
@@ -169,124 +141,22 @@ Module Loc.
forall l1 l2, diff l1 l2 -> diff l2 l1.
Proof.
destruct l1; destruct l2; unfold diff; auto.
- destruct s; auto.
- destruct s; destruct s0; intuition auto.
- Qed.
-
- Lemma diff_reg_right:
- forall l r, l <> R r -> diff (R r) l.
- Proof.
- intros. simpl. destruct l. congruence. auto.
- Qed.
-
- Lemma diff_reg_left:
- forall l r, l <> R r -> diff l (R r).
- Proof.
- intros. apply diff_sym. apply diff_reg_right. auto.
- Qed.
-
-(** [Loc.overlap l1 l2] returns [false] if [l1] and [l2] are different and
- non-overlapping, and [true] otherwise: either [l1 = l2] or they partially
- overlap. *)
-
- Definition overlap_aux (t1: typ) (d1 d2: Z) : bool :=
- if zeq d1 d2 then true else
- match t1 with
- | Tint => false
- | Tfloat => if zeq (d1 + 1) d2 then true else false
- end.
-
- Definition overlap (l1 l2: loc) : bool :=
- match l1, l2 with
- | S (Incoming d1 t1), S (Incoming d2 t2) =>
- overlap_aux t1 d1 d2 || overlap_aux t2 d2 d1
- | S (Outgoing d1 t1), S (Outgoing d2 t2) =>
- overlap_aux t1 d1 d2 || overlap_aux t2 d2 d1
- | _, _ => false
- end.
-
- Lemma overlap_aux_true_1:
- forall d1 t1 d2 t2,
- overlap_aux t1 d1 d2 = true ->
- ~(d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1).
- Proof.
- intros until t2.
- generalize (typesize_pos t1); intro.
- generalize (typesize_pos t2); intro.
- unfold overlap_aux.
- case (zeq d1 d2).
- intros. omega.
- case t1. intros; discriminate.
- case (zeq (d1 + 1) d2); intros.
- subst d2. simpl. omega.
- discriminate.
- Qed.
-
- Lemma overlap_aux_true_2:
- forall d1 t1 d2 t2,
- overlap_aux t2 d2 d1 = true ->
- ~(d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1).
- Proof.
- intros. generalize (overlap_aux_true_1 d2 t2 d1 t1 H).
- tauto.
+ intuition.
Qed.
- Lemma overlap_not_diff:
- forall l1 l2, overlap l1 l2 = true -> ~(diff l1 l2).
+ Definition diff_dec (l1 l2: loc) : { Loc.diff l1 l2 } + { ~Loc.diff l1 l2 }.
Proof.
- unfold overlap, diff; destruct l1; destruct l2; intros; try discriminate.
- destruct s; discriminate.
- destruct s; destruct s0; try discriminate.
- elim (orb_true_elim _ _ H); intro.
- apply overlap_aux_true_1; auto.
- apply overlap_aux_true_2; auto.
- elim (orb_true_elim _ _ H); intro.
- apply overlap_aux_true_1; auto.
- apply overlap_aux_true_2; auto.
- Qed.
-
- Lemma overlap_aux_false_1:
- forall t1 d1 t2 d2,
- overlap_aux t1 d1 d2 || overlap_aux t2 d2 d1 = false ->
- d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1.
- Proof.
- intros until d2. intro OV.
- generalize (orb_false_elim _ _ OV). intro OV'. elim OV'.
- unfold overlap_aux.
- case (zeq d1 d2); intro.
- intros; discriminate.
- case (zeq d2 d1); intro.
- intros; discriminate.
- case t1; case t2; simpl.
- intros; omega.
- case (zeq (d2 + 1) d1); intros. discriminate. omega.
- case (zeq (d1 + 1) d2); intros. discriminate. omega.
- case (zeq (d1 + 1) d2); intros H1 H2. discriminate.
- case (zeq (d2 + 1) d1); intros. discriminate. omega.
- Qed.
-
- Lemma non_overlap_diff:
- forall l1 l2, l1 <> l2 -> overlap l1 l2 = false -> diff l1 l2.
- Proof.
- intros. unfold diff; destruct l1; destruct l2.
- congruence.
- auto.
- destruct s; auto.
- destruct s; destruct s0; auto.
- case (zeq z z0); intro.
- compare t t0; intro.
- congruence. tauto. tauto.
- apply overlap_aux_false_1. exact H0.
- apply overlap_aux_false_1. exact H0.
- Qed.
-
- Definition diff_dec (l1 l2: loc) : { Loc.diff l1 l2 } + {~Loc.diff l1 l2}.
- Proof.
- intros. case (eq l1 l2); intros.
- right. rewrite e. apply same_not_diff.
- case_eq (overlap l1 l2); intros.
- right. apply overlap_not_diff; auto.
- left. apply non_overlap_diff; auto.
+ intros. destruct l1; destruct l2; simpl.
+ - destruct (mreg_eq r r0). right; tauto. left; auto.
+ - left; auto.
+ - left; auto.
+ - destruct (slot_eq sl sl0).
+ destruct (zle (pos + typesize ty) pos0).
+ left; auto.
+ destruct (zle (pos0 + typesize ty0) pos).
+ left; auto.
+ right; red; intros [P | [P | P]]. congruence. omega. omega.
+ left; auto.
Defined.
(** We now redefine some standard notions over lists, using the [Loc.diff]
@@ -316,12 +186,16 @@ Module Loc.
elim (diff_not_eq l l); auto.
Qed.
- Lemma reg_notin:
- forall r ll, ~(In (R r) ll) -> notin (R r) ll.
+ Lemma notin_dec (l: loc) (ll: list loc) : {notin l ll} + {~notin l ll}.
Proof.
- intros. rewrite notin_iff; intros.
- destruct l'; simpl. congruence. auto.
- Qed.
+ induction ll; simpl.
+ left; auto.
+ destruct (diff_dec l a).
+ destruct IHll.
+ left; auto.
+ right; tauto.
+ right; tauto.
+ Defined.
(** [Loc.disjoint l1 l2] is true if the locations in list [l1]
are different from all locations in list [l2]. *)
@@ -376,6 +250,17 @@ Module Loc.
| norepet_cons:
forall hd tl, notin hd tl -> norepet tl -> norepet (hd :: tl).
+ Lemma norepet_dec (ll: list loc) : {norepet ll} + {~norepet ll}.
+ Proof.
+ induction ll.
+ left; constructor.
+ destruct (notin_dec a ll).
+ destruct IHll.
+ left; constructor; auto.
+ right; red; intros P; inv P; contradiction.
+ right; red; intros P; inv P; contradiction.
+ Defined.
+
(** [Loc.no_overlap l1 l2] holds if elements of [l1] never overlap partially
with elements of [l2]. *)
@@ -384,8 +269,6 @@ Module Loc.
End Loc.
-Global Opaque Loc.eq Loc.diff_dec.
-
(** * Mappings from locations to values *)
(** The [Locmap] module defines mappings from locations to values,
@@ -413,20 +296,20 @@ Module Locmap.
Definition set (l: loc) (v: val) (m: t) : t :=
fun (p: loc) =>
- if Loc.eq l p then v else if Loc.overlap l p then Vundef else m p.
+ if Loc.eq l p then v else if Loc.diff_dec l p then m p else Vundef.
Lemma gss: forall l v m, (set l v m) l = v.
Proof.
- intros. unfold set. case (Loc.eq l l); tauto.
+ intros. unfold set. rewrite dec_eq_true. auto.
Qed.
Lemma gso: forall l v m p, Loc.diff l p -> (set l v m) p = m p.
Proof.
- intros. unfold set. case (Loc.eq l p); intro.
- subst p. elim (Loc.same_not_diff _ H).
- caseEq (Loc.overlap l p); intro.
- elim (Loc.overlap_not_diff _ _ H0 H).
+ intros. unfold set. destruct (Loc.eq l p).
+ subst p. elim (Loc.same_not_diff _ H).
+ destruct (Loc.diff_dec l p).
auto.
+ contradiction.
Qed.
Fixpoint undef (ll: list loc) (m: t) {struct ll} : t :=
@@ -446,10 +329,172 @@ Module Locmap.
assert (P: forall ll l m, m l = Vundef -> (undef ll m) l = Vundef).
induction ll; simpl; intros. auto. apply IHll.
unfold set. destruct (Loc.eq a l); auto.
- destruct (Loc.overlap a l); auto.
+ destruct (Loc.diff_dec a l); auto.
induction ll; simpl; intros. contradiction.
destruct H. apply P. subst a. apply gss.
auto.
Qed.
+ Fixpoint setlist (ll: list loc) (vl: list val) (m: t) {struct ll} : t :=
+ match ll, vl with
+ | l1 :: ls, v1 :: vs => setlist ls vs (set l1 v1 m)
+ | _, _ => m
+ end.
+
+ Lemma gsetlisto: forall l ll vl m, Loc.notin l ll -> (setlist ll vl m) l = m l.
+ Proof.
+ induction ll; simpl; intros.
+ auto.
+ destruct vl; auto. destruct H. rewrite IHll; auto. apply gso; auto. apply Loc.diff_sym; auto.
+ Qed.
+
End Locmap.
+
+(** * Total ordering over locations *)
+
+Module IndexedTyp <: INDEXED_TYPE.
+ Definition t := typ.
+ Definition index (x: t) :=
+ match x with Tint => 1%positive | Tfloat => 2%positive | Tlong => 3%positive end.
+ Lemma index_inj: forall x y, index x = index y -> x = y.
+ Proof. destruct x; destruct y; simpl; congruence. Qed.
+ Definition eq := typ_eq.
+End IndexedTyp.
+
+Module OrderedTyp := OrderedIndexed(IndexedTyp).
+
+Module IndexedSlot <: INDEXED_TYPE.
+ Definition t := slot.
+ Definition index (x: t) :=
+ match x with Local => 1%positive | Incoming => 2%positive | Outgoing => 3%positive end.
+ Lemma index_inj: forall x y, index x = index y -> x = y.
+ Proof. destruct x; destruct y; simpl; congruence. Qed.
+ Definition eq := slot_eq.
+End IndexedSlot.
+
+Module OrderedSlot := OrderedIndexed(IndexedSlot).
+
+Module OrderedLoc <: OrderedType.
+ Definition t := loc.
+ Definition eq (x y: t) := x = y.
+ Definition lt (x y: t) :=
+ match x, y with
+ | R r1, R r2 => Plt (IndexedMreg.index r1) (IndexedMreg.index r2)
+ | R _, S _ _ _ => True
+ | S _ _ _, R _ => False
+ | S sl1 ofs1 ty1, S sl2 ofs2 ty2 =>
+ OrderedSlot.lt sl1 sl2 \/ (sl1 = sl2 /\
+ (ofs1 < ofs2 \/ (ofs1 = ofs2 /\ OrderedTyp.lt ty1 ty2)))
+ end.
+ Lemma eq_refl : forall x : t, eq x x.
+ Proof (@refl_equal t).
+ Lemma eq_sym : forall x y : t, eq x y -> eq y x.
+ Proof (@sym_equal t).
+ Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
+ Proof (@trans_equal t).
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof.
+ unfold lt; intros.
+ destruct x; destruct y; destruct z; try tauto.
+ eapply Plt_trans; eauto.
+ destruct H.
+ destruct H0. left; eapply OrderedSlot.lt_trans; eauto.
+ destruct H0. subst sl0. auto.
+ destruct H. subst sl.
+ destruct H0. auto.
+ destruct H.
+ right. split. auto.
+ intuition.
+ right; split. congruence. eapply OrderedTyp.lt_trans; eauto.
+ Qed.
+ Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Proof.
+ unfold lt, eq; intros; red; intros. subst y.
+ destruct x.
+ eelim Plt_strict; eauto.
+ destruct H. eelim OrderedSlot.lt_not_eq; eauto. red; auto.
+ destruct H. destruct H0. omega.
+ destruct H0. eelim OrderedTyp.lt_not_eq; eauto. red; auto.
+ Qed.
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ intros. destruct x; destruct y.
+ - destruct (OrderedPositive.compare (IndexedMreg.index r) (IndexedMreg.index r0)).
+ + apply LT. red. auto.
+ + apply EQ. red. f_equal. apply IndexedMreg.index_inj. auto.
+ + apply GT. red. auto.
+ - apply LT. red; auto.
+ - apply GT. red; auto.
+ - destruct (OrderedSlot.compare sl sl0).
+ + apply LT. red; auto.
+ + destruct (OrderedZ.compare pos pos0).
+ * apply LT. red. auto.
+ * destruct (OrderedTyp.compare ty ty0).
+ apply LT. red; auto.
+ apply EQ. red; red in e; red in e0; red in e1. congruence.
+ apply GT. red; auto.
+ * apply GT. red. auto.
+ + apply GT. red; auto.
+ Defined.
+ Definition eq_dec := Loc.eq.
+
+(** Connection between the ordering defined here and the [Loc.diff] predicate. *)
+
+ Definition diff_low_bound (l: loc) : loc :=
+ match l with
+ | R mr => l
+ | S sl ofs ty => S sl (ofs - 1) Tfloat
+ end.
+
+ Definition diff_high_bound (l: loc) : loc :=
+ match l with
+ | R mr => l
+ | S sl ofs ty => S sl (ofs + typesize ty - 1) Tlong
+ end.
+
+ Lemma outside_interval_diff:
+ forall l l', lt l' (diff_low_bound l) \/ lt (diff_high_bound l) l' -> Loc.diff l l'.
+ Proof.
+ intros.
+ destruct l as [mr | sl ofs ty]; destruct l' as [mr' | sl' ofs' ty']; simpl in *; auto.
+ - assert (IndexedMreg.index mr <> IndexedMreg.index mr').
+ { destruct H. apply sym_not_equal. apply Plt_ne; auto. apply Plt_ne; auto. }
+ congruence.
+ - assert (RANGE: forall ty, 1 <= typesize ty <= 2).
+ { intros; unfold typesize. destruct ty0; omega. }
+ destruct H.
+ + destruct H. left. apply sym_not_equal. apply OrderedSlot.lt_not_eq; auto.
+ destruct H. right.
+ destruct H0. right. generalize (RANGE ty'); omega.
+ destruct H0.
+ assert (ty' = Tint).
+ { unfold OrderedTyp.lt in H1. destruct ty'; compute in H1; congruence. }
+ subst ty'. right. simpl typesize. omega.
+ + destruct H. left. apply OrderedSlot.lt_not_eq; auto.
+ destruct H. right.
+ destruct H0. left; omega.
+ destruct H0. exfalso. destruct ty'; compute in H1; congruence.
+ Qed.
+
+ Lemma diff_outside_interval:
+ forall l l', Loc.diff l l' -> lt l' (diff_low_bound l) \/ lt (diff_high_bound l) l'.
+ Proof.
+ intros.
+ destruct l as [mr | sl ofs ty]; destruct l' as [mr' | sl' ofs' ty']; simpl in *; auto.
+ - unfold Plt, Pos.lt. destruct (Pos.compare (IndexedMreg.index mr) (IndexedMreg.index mr')) eqn:C.
+ elim H. apply IndexedMreg.index_inj. apply Pos.compare_eq_iff. auto.
+ auto.
+ rewrite Pos.compare_antisym. rewrite C. auto.
+ - destruct (OrderedSlot.compare sl sl'); auto.
+ destruct H. contradiction.
+ destruct H.
+ right; right; split; auto. left; omega.
+ left; right; split; auto. destruct ty'; simpl in *.
+ destruct (zlt ofs' (ofs - 1)). left; auto.
+ right; split. omega. compute. auto.
+ left; omega.
+ left; omega.
+ Qed.
+
+End OrderedLoc.
+
diff --git a/backend/Mach.v b/backend/Mach.v
index 0728c4d..5030de1 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -60,7 +60,7 @@ Inductive instruction: Type :=
| Mstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mcall: signature -> mreg + ident -> instruction
| Mtailcall: signature -> mreg + ident -> instruction
- | Mbuiltin: external_function -> list mreg -> mreg -> instruction
+ | Mbuiltin: external_function -> list mreg -> list mreg -> instruction
| Mannot: external_function -> list annot_param -> instruction
| Mlabel: label -> instruction
| Mgoto: label -> instruction
@@ -124,7 +124,7 @@ store in the reserved location.
*)
Definition chunk_of_type (ty: typ) :=
- match ty with Tint => Mint32 | Tfloat => Mfloat64al32 end.
+ match ty with Tint => Mint32 | Tfloat => Mfloat64al32 | Tlong => Mint64 end.
Definition load_stack (m: mem) (sp: val) (ty: typ) (ofs: int) :=
Mem.loadv (chunk_of_type ty) m (Val.add sp (Vint ofs)).
@@ -147,38 +147,29 @@ Notation "a # b <- c" := (Regmap.set b c a) (at level 1, b at next level).
Fixpoint undef_regs (rl: list mreg) (rs: regset) {struct rl} : regset :=
match rl with
| nil => rs
- | r1 :: rl' => undef_regs rl' (Regmap.set r1 Vundef rs)
+ | r1 :: rl' => Regmap.set r1 Vundef (undef_regs rl' rs)
end.
Lemma undef_regs_other:
forall r rl rs, ~In r rl -> undef_regs rl rs r = rs r.
Proof.
- induction rl; simpl; intros. auto. rewrite IHrl. apply Regmap.gso. intuition. intuition.
+ induction rl; simpl; intros. auto. rewrite Regmap.gso. apply IHrl. intuition. intuition.
Qed.
Lemma undef_regs_same:
- forall r rl rs, In r rl \/ rs r = Vundef -> undef_regs rl rs r = Vundef.
+ forall r rl rs, In r rl -> undef_regs rl rs r = Vundef.
Proof.
induction rl; simpl; intros. tauto.
- destruct H. destruct H. apply IHrl. right. subst; apply Regmap.gss.
- auto.
- apply IHrl. right. unfold Regmap.set. destruct (RegEq.eq r a); auto.
+ destruct H. subst a. apply Regmap.gss.
+ unfold Regmap.set. destruct (RegEq.eq r a); auto.
Qed.
-Definition undef_temps (rs: regset) :=
- undef_regs temporary_regs rs.
-
-Definition undef_move (rs: regset) :=
- undef_regs destroyed_at_move_regs rs.
-
-Definition undef_op (op: operation) (rs: regset) :=
- match op with
- | Omove => undef_move rs
- | _ => undef_temps rs
+Fixpoint set_regs (rl: list mreg) (vl: list val) (rs: regset) : regset :=
+ match rl, vl with
+ | r1 :: rl', v1 :: vl' => set_regs rl' vl' (Regmap.set r1 v1 rs)
+ | _, _ => rs
end.
-Definition undef_setstack (rs: regset) := undef_move rs.
-
Definition is_label (lbl: label) (instr: instruction) : bool :=
match instr with
| Mlabel lbl' => if peq lbl lbl' then true else false
@@ -231,7 +222,7 @@ Inductive extcall_arg: regset -> mem -> val -> loc -> val -> Prop :=
extcall_arg rs m sp (R r) (rs r)
| extcall_arg_stack: forall rs m sp ofs ty v,
load_stack m sp ty (Int.repr (Stacklayout.fe_ofs_arg + 4 * ofs)) = Some v ->
- extcall_arg rs m sp (S (Outgoing ofs ty)) v.
+ extcall_arg rs m sp (S Outgoing ofs ty) v.
Definition extcall_arguments
(rs: regset) (m: mem) (sp: val) (sg: signature) (args: list val) : Prop :=
@@ -306,34 +297,39 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp (Mgetstack ofs ty dst :: c) rs m)
E0 (State s f sp c (rs#dst <- v) m)
| exec_Msetstack:
- forall s f sp src ofs ty c rs m m',
+ forall s f sp src ofs ty c rs m m' rs',
store_stack m sp ty ofs (rs src) = Some m' ->
+ rs' = undef_regs (destroyed_by_op Omove) rs ->
step (State s f sp (Msetstack src ofs ty :: c) rs m)
- E0 (State s f sp c (undef_setstack rs) m')
+ E0 (State s f sp c rs' m')
| exec_Mgetparam:
- forall s fb f sp ofs ty dst c rs m v,
+ forall s fb f sp ofs ty dst c rs m v rs',
Genv.find_funct_ptr ge fb = Some (Internal f) ->
load_stack m sp Tint f.(fn_link_ofs) = Some (parent_sp s) ->
load_stack m (parent_sp s) ty ofs = Some v ->
+ rs' = (rs # temp_for_parent_frame <- Vundef # dst <- v) ->
step (State s fb sp (Mgetparam ofs ty dst :: c) rs m)
- E0 (State s fb sp c (rs # IT1 <- Vundef # dst <- v) m)
+ E0 (State s fb sp c rs' m)
| exec_Mop:
- forall s f sp op args res c rs m v,
+ forall s f sp op args res c rs m v rs',
eval_operation ge sp op rs##args m = Some v ->
+ rs' = ((undef_regs (destroyed_by_op op) rs)#res <- v) ->
step (State s f sp (Mop op args res :: c) rs m)
- E0 (State s f sp c ((undef_op op rs)#res <- v) m)
+ E0 (State s f sp c rs' m)
| exec_Mload:
- forall s f sp chunk addr args dst c rs m a v,
+ forall s f sp chunk addr args dst c rs m a v rs',
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
+ rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- v) ->
step (State s f sp (Mload chunk addr args dst :: c) rs m)
- E0 (State s f sp c ((undef_temps rs)#dst <- v) m)
+ E0 (State s f sp c rs' m)
| exec_Mstore:
- forall s f sp chunk addr args src c rs m m' a,
+ forall s f sp chunk addr args src c rs m m' a rs',
eval_addressing ge sp addr rs##args = Some a ->
Mem.storev chunk m a (rs src) = Some m' ->
+ rs' = undef_regs (destroyed_by_store chunk addr) rs ->
step (State s f sp (Mstore chunk addr args src :: c) rs m)
- E0 (State s f sp c (undef_temps rs) m')
+ E0 (State s f sp c rs' m')
| exec_Mcall:
forall s fb sp sig ros c rs m f f' ra,
find_function_ptr ge ros rs = Some f' ->
@@ -352,14 +348,15 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s fb (Vptr stk soff) (Mtailcall sig ros :: c) rs m)
E0 (Callstate s f' rs m')
| exec_Mbuiltin:
- forall s f sp rs m ef args res b t v m',
- external_call ef ge rs##args m t v m' ->
+ forall s f sp rs m ef args res b t vl rs' m',
+ external_call' ef ge rs##args m t vl m' ->
+ rs' = set_regs res vl (undef_regs (destroyed_by_builtin ef) rs) ->
step (State s f sp (Mbuiltin ef args res :: b) rs m)
- t (State s f sp b ((undef_temps rs)#res <- v) m')
+ t (State s f sp b rs' m')
| exec_Mannot:
forall s f sp rs m ef args b vargs t v m',
annot_arguments rs m sp args vargs ->
- external_call ef ge vargs m t v m' ->
+ external_call' ef ge vargs m t v m' ->
step (State s f sp (Mannot ef args :: b) rs m)
t (State s f sp b rs m')
| exec_Mgoto:
@@ -369,25 +366,28 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s fb sp (Mgoto lbl :: c) rs m)
E0 (State s fb sp c' rs m)
| exec_Mcond_true:
- forall s fb f sp cond args lbl c rs m c',
+ forall s fb f sp cond args lbl c rs m c' rs',
eval_condition cond rs##args m = Some true ->
Genv.find_funct_ptr ge fb = Some (Internal f) ->
find_label lbl f.(fn_code) = Some c' ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
step (State s fb sp (Mcond cond args lbl :: c) rs m)
- E0 (State s fb sp c' (undef_temps rs) m)
+ E0 (State s fb sp c' rs' m)
| exec_Mcond_false:
- forall s f sp cond args lbl c rs m,
+ forall s f sp cond args lbl c rs m rs',
eval_condition cond rs##args m = Some false ->
+ rs' = undef_regs (destroyed_by_cond cond) rs ->
step (State s f sp (Mcond cond args lbl :: c) rs m)
- E0 (State s f sp c (undef_temps rs) m)
+ E0 (State s f sp c rs' m)
| exec_Mjumptable:
- forall s fb f sp arg tbl c rs m n lbl c',
+ forall s fb f sp arg tbl c rs m n lbl c' rs',
rs arg = Vint n ->
list_nth_z tbl (Int.unsigned n) = Some lbl ->
Genv.find_funct_ptr ge fb = Some (Internal f) ->
find_label lbl f.(fn_code) = Some c' ->
+ rs' = undef_regs destroyed_by_jumptable rs ->
step (State s fb sp (Mjumptable arg tbl :: c) rs m)
- E0 (State s fb sp c' (undef_temps rs) m)
+ E0 (State s fb sp c' rs' m)
| exec_Mreturn:
forall s fb stk soff c rs m f m',
Genv.find_funct_ptr ge fb = Some (Internal f) ->
@@ -397,20 +397,21 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s fb (Vptr stk soff) (Mreturn :: c) rs m)
E0 (Returnstate s rs m')
| exec_function_internal:
- forall s fb rs m f m1 m2 m3 stk,
+ forall s fb rs m f m1 m2 m3 stk rs',
Genv.find_funct_ptr ge fb = Some (Internal f) ->
Mem.alloc m 0 f.(fn_stacksize) = (m1, stk) ->
let sp := Vptr stk Int.zero in
store_stack m1 sp Tint f.(fn_link_ofs) (parent_sp s) = Some m2 ->
store_stack m2 sp Tint f.(fn_retaddr_ofs) (parent_ra s) = Some m3 ->
+ rs' = undef_regs destroyed_at_function_entry rs ->
step (Callstate s fb rs m)
- E0 (State s fb sp f.(fn_code) (undef_temps rs) m3)
+ E0 (State s fb sp f.(fn_code) rs' m3)
| exec_function_external:
forall s fb rs m t rs' ef args res m',
Genv.find_funct_ptr ge fb = Some (External ef) ->
- external_call ef ge args m t res m' ->
extcall_arguments rs m (parent_sp s) (ef_sig ef) args ->
- rs' = (rs#(loc_result (ef_sig ef)) <- res) ->
+ external_call' ef ge args m t res m' ->
+ rs' = set_regs (loc_result (ef_sig ef)) res rs ->
step (Callstate s fb rs m)
t (Returnstate s rs' m')
| exec_return:
@@ -428,9 +429,10 @@ Inductive initial_state (p: program): state -> Prop :=
initial_state p (Callstate nil fb (Regmap.init Vundef) m0).
Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall rs m r,
- rs (loc_result (mksignature nil (Some Tint))) = Vint r ->
- final_state (Returnstate nil rs m) r.
+ | final_state_intro: forall rs m r retcode,
+ loc_result (mksignature nil (Some Tint)) = r :: nil ->
+ rs r = Vint retcode ->
+ final_state (Returnstate nil rs m) retcode.
Definition semantics (rao: function -> code -> int -> Prop) (p: program) :=
Semantics (step rao) (initial_state p) final_state (Genv.globalenv p).
diff --git a/backend/Machtyping.v b/backend/Machtyping.v
deleted file mode 100644
index 2dc19be..0000000
--- a/backend/Machtyping.v
+++ /dev/null
@@ -1,108 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Type system for the Mach intermediate language. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Op.
-Require Import Locations.
-Require Import Conventions.
-Require Import Mach.
-
-(** * Typing rules *)
-
-Inductive wt_instr : instruction -> Prop :=
- | wt_Mlabel:
- forall lbl,
- wt_instr (Mlabel lbl)
- | wt_Mgetstack:
- forall ofs ty r,
- mreg_type r = ty ->
- wt_instr (Mgetstack ofs ty r)
- | wt_Msetstack:
- forall ofs ty r,
- mreg_type r = ty ->
- wt_instr (Msetstack r ofs ty)
- | wt_Mgetparam:
- forall ofs ty r,
- mreg_type r = ty ->
- wt_instr (Mgetparam ofs ty r)
- | wt_Mopmove:
- forall r1 r,
- mreg_type r1 = mreg_type r ->
- wt_instr (Mop Omove (r1 :: nil) r)
- | wt_Mop:
- forall op args res,
- op <> Omove ->
- (List.map mreg_type args, mreg_type res) = type_of_operation op ->
- wt_instr (Mop op args res)
- | wt_Mload:
- forall chunk addr args dst,
- List.map mreg_type args = type_of_addressing addr ->
- mreg_type dst = type_of_chunk chunk ->
- wt_instr (Mload chunk addr args dst)
- | wt_Mstore:
- forall chunk addr args src,
- List.map mreg_type args = type_of_addressing addr ->
- mreg_type src = type_of_chunk chunk ->
- wt_instr (Mstore chunk addr args src)
- | wt_Mcall:
- forall sig ros,
- match ros with inl r => mreg_type r = Tint | inr s => True end ->
- wt_instr (Mcall sig ros)
- | wt_Mtailcall:
- forall sig ros,
- tailcall_possible sig ->
- match ros with inl r => mreg_type r = Tint | inr s => True end ->
- wt_instr (Mtailcall sig ros)
- | wt_Mbuiltin:
- forall ef args res,
- List.map mreg_type args = (ef_sig ef).(sig_args) ->
- mreg_type res = proj_sig_res (ef_sig ef) ->
- wt_instr (Mbuiltin ef args res)
- | wt_Mannot:
- forall ef args,
- ef_reloads ef = false ->
- wt_instr (Mannot ef args)
- | wt_Mgoto:
- forall lbl,
- wt_instr (Mgoto lbl)
- | wt_Mcond:
- forall cond args lbl,
- List.map mreg_type args = type_of_condition cond ->
- wt_instr (Mcond cond args lbl)
- | wt_Mjumptable:
- forall arg tbl,
- mreg_type arg = Tint ->
- list_length_z tbl * 4 <= Int.max_unsigned ->
- wt_instr (Mjumptable arg tbl)
- | wt_Mreturn:
- wt_instr Mreturn.
-
-Record wt_function (f: function) : Prop := mk_wt_function {
- wt_function_instrs:
- forall instr, In instr f.(fn_code) -> wt_instr instr;
- wt_function_stacksize:
- 0 <= f.(fn_stacksize) <= Int.max_unsigned
-}.
-
-Inductive wt_fundef: fundef -> Prop :=
- | wt_fundef_external: forall ef,
- wt_fundef (External ef)
- | wt_function_internal: forall f,
- wt_function f ->
- wt_fundef (Internal f).
-
-Definition wt_program (p: program) : Prop :=
- forall i f, In (i, Gfun f) (prog_defs p) -> wt_fundef f.
diff --git a/backend/PrintCminor.ml b/backend/PrintCminor.ml
index 01ea1e6..8612b73 100644
--- a/backend/PrintCminor.ml
+++ b/backend/PrintCminor.ml
@@ -31,13 +31,13 @@ let rec precedence = function
| Evar _ -> (16, NA)
| Econst _ -> (16, NA)
| Eunop _ -> (15, RtoL)
- | Ebinop((Omul|Odiv|Odivu|Omod|Omodu|Omulf|Odivf), _, _) -> (13, LtoR)
- | Ebinop((Oadd|Osub|Oaddf|Osubf), _, _) -> (12, LtoR)
- | Ebinop((Oshl|Oshr|Oshru), _, _) -> (11, LtoR)
- | Ebinop((Ocmp _|Ocmpu _|Ocmpf _), _, _) -> (10, LtoR)
- | Ebinop(Oand, _, _) -> (8, LtoR)
- | Ebinop(Oxor, _, _) -> (7, LtoR)
- | Ebinop(Oor, _, _) -> (6, LtoR)
+ | Ebinop((Omul|Odiv|Odivu|Omod|Omodu|Omulf|Odivf|Omull|Odivl|Odivlu|Omodl|Omodlu), _, _) -> (13, LtoR)
+ | Ebinop((Oadd|Osub|Oaddf|Osubf|Oaddl|Osubl), _, _) -> (12, LtoR)
+ | Ebinop((Oshl|Oshr|Oshru|Oshll|Oshrl|Oshrlu), _, _) -> (11, LtoR)
+ | Ebinop((Ocmp _|Ocmpu _|Ocmpf _|Ocmpl _|Ocmplu _), _, _) -> (10, LtoR)
+ | Ebinop((Oand|Oandl), _, _) -> (8, LtoR)
+ | Ebinop((Oxor|Oxorl), _, _) -> (7, LtoR)
+ | Ebinop((Oor|Oorl), _, _) -> (6, LtoR)
| Eload _ -> (15, RtoL)
(* Naming idents. *)
@@ -60,6 +60,15 @@ let name_of_unop = function
| Ointuoffloat -> "intuoffloat"
| Ofloatofint -> "floatofint"
| Ofloatofintu -> "floatofintu"
+ | Onegl -> "-l"
+ | Onotl -> "~l"
+ | Ointoflong -> "intofflong"
+ | Olongofint -> "longofint"
+ | Olongofintu -> "longofintu"
+ | Olongoffloat -> "longoffloat"
+ | Olonguoffloat -> "longuoffloat"
+ | Ofloatoflong -> "floatoflong"
+ | Ofloatoflongu -> "floatoflongu"
let comparison_name = function
| Ceq -> "=="
@@ -87,9 +96,24 @@ let name_of_binop = function
| Osubf -> "-f"
| Omulf -> "*f"
| Odivf -> "/f"
+ | Oaddl -> "+l"
+ | Osubl -> "-l"
+ | Omull -> "*l"
+ | Odivl -> "/l"
+ | Odivlu -> "/lu"
+ | Omodl -> "%l"
+ | Omodlu -> "%lu"
+ | Oandl -> "&l"
+ | Oorl -> "|l"
+ | Oxorl -> "^l"
+ | Oshll -> "<<l"
+ | Oshrl -> ">>l"
+ | Oshrlu -> ">>lu"
| Ocmp c -> comparison_name c
| Ocmpu c -> comparison_name c ^ "u"
| Ocmpf c -> comparison_name c ^ "f"
+ | Ocmpl c -> comparison_name c ^ "l"
+ | Ocmplu c -> comparison_name c ^ "lu"
(* Expressions *)
@@ -109,6 +133,8 @@ let rec expr p (prec, e) =
fprintf p "%ld" (camlint_of_coqint n)
| Econst(Ofloatconst f) ->
fprintf p "%F" (camlfloat_of_coqfloat f)
+ | Econst(Olongconst n) ->
+ fprintf p "%LdLL" (camlint64_of_coqint n)
| Econst(Oaddrsymbol(id, ofs)) ->
let ofs = camlint_of_coqint ofs in
if ofs = 0l
@@ -141,6 +167,7 @@ let rec print_expr_list p (first, rl) =
let name_of_type = function
| Tint -> "int"
| Tfloat -> "float"
+ | Tlong -> "long"
let rec print_sig p = function
| {sig_args = []; sig_res = None} -> fprintf p "void"
@@ -262,6 +289,7 @@ let print_init_data p = function
| Init_int8 i -> fprintf p "int8 %ld" (camlint_of_coqint i)
| Init_int16 i -> fprintf p "int16 %ld" (camlint_of_coqint i)
| Init_int32 i -> fprintf p "%ld" (camlint_of_coqint i)
+ | Init_int64 i -> fprintf p "%Ld" (camlint64_of_coqint i)
| Init_float32 f -> fprintf p "float32 %F" (camlfloat_of_coqfloat f)
| Init_float64 f -> fprintf p "%F" (camlfloat_of_coqfloat f)
| Init_space i -> fprintf p "[%ld]" (camlint_of_coqint i)
diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml
index 94f5af0..1149dd0 100644
--- a/backend/PrintLTL.ml
+++ b/backend/PrintLTL.ml
@@ -22,21 +22,30 @@ open Locations
open LTL
open PrintAST
open PrintOp
-open PrintRTL
+
+let mreg pp r =
+ match Machregsaux.name_of_register r with
+ | Some s -> fprintf pp "%s" s
+ | None -> fprintf pp "<unknown machreg>"
+
+let rec mregs pp = function
+ | [] -> ()
+ | [r] -> mreg pp r
+ | r1::rl -> fprintf pp "%a, %a" mreg r1 mregs rl
+
+let slot pp (sl, ofs, ty) =
+ match sl with
+ | Local ->
+ fprintf pp "local(%ld,%s)" (camlint_of_coqint ofs) (name_of_type ty)
+ | Incoming ->
+ fprintf pp "incoming(%ld,%s)" (camlint_of_coqint ofs) (name_of_type ty)
+ | Outgoing ->
+ fprintf pp "outgoing(%ld,%s)" (camlint_of_coqint ofs) (name_of_type ty)
let loc pp l =
match l with
- | R r ->
- begin match Machregsaux.name_of_register r with
- | Some s -> fprintf pp "%s" s
- | None -> fprintf pp "<unknown machreg>"
- end
- | S(Local(ofs, ty)) ->
- fprintf pp "local(%ld,%s)" (camlint_of_coqint ofs) (name_of_typ ty)
- | S(Incoming(ofs, ty)) ->
- fprintf pp "incoming(%ld,%s)" (camlint_of_coqint ofs) (name_of_typ ty)
- | S(Outgoing(ofs, ty)) ->
- fprintf pp "outgoing(%ld,%s)" (camlint_of_coqint ofs) (name_of_typ ty)
+ | R r -> mreg pp r
+ | S(sl, ofs, ty) -> slot pp (sl, ofs, ty)
let rec locs pp = function
| [] -> ()
@@ -44,77 +53,93 @@ let rec locs pp = function
| r1::rl -> fprintf pp "%a, %a" loc r1 locs rl
let ros pp = function
- | Coq_inl r -> loc pp r
+ | Coq_inl r -> mreg pp r
| Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s)
let print_succ pp s dfl =
- let s = camlint_of_positive s in
- if s <> dfl then fprintf pp " goto %ld@ " s
-
-let print_instruction pp (pc, i) =
- fprintf pp "%5ld: " pc;
- match i with
- | Lnop s ->
- let s = camlint_of_positive s in
- if s = Int32.pred pc
- then fprintf pp "nop@ "
- else fprintf pp "goto %ld@ " s
- | Lop(op, args, res, s) ->
- fprintf pp "%a = %a@ " loc res (print_operator loc) (op, args);
- print_succ pp s (Int32.pred pc)
- | Lload(chunk, addr, args, dst, s) ->
- fprintf pp "%a = %s[%a]@ "
- loc dst (name_of_chunk chunk) (print_addressing loc) (addr, args);
- print_succ pp s (Int32.pred pc)
- | Lstore(chunk, addr, args, src, s) ->
- fprintf pp "%s[%a] = %a@ "
- (name_of_chunk chunk) (print_addressing loc) (addr, args) loc src;
- print_succ pp s (Int32.pred pc)
- | Lcall(sg, fn, args, res, s) ->
- fprintf pp "%a = %a(%a)@ "
- loc res ros fn locs args;
- print_succ pp s (Int32.pred pc)
- | Ltailcall(sg, fn, args) ->
- fprintf pp "tailcall %a(%a)@ "
- ros fn locs args
- | Lbuiltin(ef, args, res, s) ->
- fprintf pp "%a = builtin %s(%a)@ "
- loc res (name_of_external ef) locs args;
- print_succ pp s (Int32.pred pc)
+ let s = P.to_int32 s in
+ if s <> dfl then fprintf pp "goto %ld" s
+
+let print_instruction pp succ = function
+ | Lop(op, args, res) ->
+ fprintf pp "%a = %a;" mreg res (print_operation mreg) (op, args)
+ | Lload(chunk, addr, args, dst) ->
+ fprintf pp "%a = %s[%a];"
+ mreg dst (name_of_chunk chunk) (print_addressing mreg) (addr, args)
+ | Lgetstack(sl, ofs, ty, dst) ->
+ fprintf pp "%a = %a;" mreg dst slot (sl, ofs, ty)
+ | Lsetstack(src, sl, ofs, ty) ->
+ fprintf pp "%a = %a;" slot (sl, ofs, ty) mreg src
+ | Lstore(chunk, addr, args, src) ->
+ fprintf pp "%s[%a] = %a;"
+ (name_of_chunk chunk) (print_addressing mreg) (addr, args) mreg src
+ | Lcall(sg, fn) ->
+ fprintf pp "call %a;" ros fn
+ | Ltailcall(sg, fn) ->
+ fprintf pp "tailcall %a;" ros fn
+ | Lbuiltin(ef, args, res) ->
+ fprintf pp "%a = builtin %s(%a);"
+ mregs res (name_of_external ef) mregs args
+ | Lannot(ef, args) ->
+ fprintf pp "builtin %s(%a);"
+ (name_of_external ef) locs args
+ | Lbranch s ->
+ print_succ pp s succ
| Lcond(cond, args, s1, s2) ->
- fprintf pp "if (%a) goto %ld else goto %ld@ "
- (print_condition loc) (cond, args)
- (camlint_of_positive s1) (camlint_of_positive s2)
+ fprintf pp "if (%a) goto %ld else goto %ld"
+ (print_condition mreg) (cond, args)
+ (P.to_int32 s1) (P.to_int32 s2)
| Ljumptable(arg, tbl) ->
let tbl = Array.of_list tbl in
- fprintf pp "@[<v 2>jumptable (%a)" loc arg;
+ fprintf pp "@[<v 2>jumptable (%a)" mreg arg;
for i = 0 to Array.length tbl - 1 do
- fprintf pp "@ case %d: goto %ld" i (camlint_of_positive tbl.(i))
+ fprintf pp "@ case %d: goto %ld" i (P.to_int32 tbl.(i))
done;
- fprintf pp "@]@ "
- | Lreturn None ->
- fprintf pp "return@ "
- | Lreturn (Some arg) ->
- fprintf pp "return %a@ " loc arg
-
-let print_function pp f =
- fprintf pp "@[<v 2>f(%a) {@ " locs f.fn_params;
+ fprintf pp "@]"
+ | Lreturn ->
+ fprintf pp "return"
+
+let rec print_instructions pp succ = function
+ | [] -> ()
+ | [i] -> print_instruction pp succ i
+ | i :: il ->
+ print_instruction pp succ i;
+ fprintf pp "@ ";
+ print_instructions pp succ il
+
+let print_block pp (pc, blk) =
+ fprintf pp "%5ld: @[<hov 0>" pc;
+ print_instructions pp (Int32.pred pc) blk;
+ fprintf pp "@]@ "
+
+let print_function pp id f =
+ fprintf pp "@[<v 2>%s() {@ " (extern_atom id);
let instrs =
List.sort
(fun (pc1, _) (pc2, _) -> Pervasives.compare pc2 pc1)
(List.map
- (fun (pc, i) -> (camlint_of_positive pc, i))
+ (fun (pc, i) -> (P.to_int32 pc, i))
(PTree.elements f.fn_code)) in
print_succ pp f.fn_entrypoint
(match instrs with (pc1, _) :: _ -> pc1 | [] -> -1l);
- List.iter (print_instruction pp) instrs;
+ List.iter (print_block pp) instrs;
fprintf pp "@;<0 -2>}@]@."
-let print_fundef fd =
- begin match fd with
- | Internal f -> print_function std_formatter f
- | External _ -> ()
- end;
- fd
+let print_globdef pp (id, gd) =
+ match gd with
+ | Gfun(Internal f) -> print_function pp id f
+ | _ -> ()
+
+let print_program pp (prog: LTL.program) =
+ List.iter (print_globdef pp) prog.prog_defs
+let destination : string option ref = ref None
+let print_if prog =
+ match !destination with
+ | None -> ()
+ | Some f ->
+ let oc = open_out f in
+ let pp = formatter_of_out_channel oc in
+ print_program pp prog;
+ close_out oc
diff --git a/backend/PrintMach.ml b/backend/PrintMach.ml
index 7e6c343..e7cb947 100644
--- a/backend/PrintMach.ml
+++ b/backend/PrintMach.ml
@@ -76,7 +76,7 @@ let print_instruction pp i =
fprintf pp "tailcall %a@ " ros fn
| Mbuiltin(ef, args, res) ->
fprintf pp "%a = builtin %s(%a)@ "
- reg res (name_of_external ef) regs args
+ regs res (name_of_external ef) regs args
| Mannot(ef, args) ->
fprintf pp "%s(%a)@ " (name_of_external ef) annot_params args
| Mlabel lbl ->
diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml
new file mode 100644
index 0000000..756fc58
--- /dev/null
+++ b/backend/PrintXTL.ml
@@ -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. *)
+(* *)
+(* *********************************************************************)
+
+(** Pretty-printer for XTL *)
+
+open Format
+open Camlcoq
+open Datatypes
+open Maps
+open AST
+open Registers
+open Op
+open Locations
+open PrintAST
+open PrintOp
+open XTL
+
+let mreg pp r =
+ match Machregsaux.name_of_register r with
+ | Some s -> fprintf pp "%s" s
+ | None -> fprintf pp "<unknown machreg>"
+
+let short_name_of_type = function Tint -> 'i' | Tfloat -> 'f' | Tlong -> 'l'
+
+let loc pp = function
+ | R r -> mreg pp r
+ | S(Local, ofs, ty) ->
+ fprintf pp "L%c%ld" (short_name_of_type ty) (camlint_of_coqint ofs)
+ | S(Incoming, ofs, ty) ->
+ fprintf pp "I%c%ld" (short_name_of_type ty) (camlint_of_coqint ofs)
+ | S(Outgoing, ofs, ty) ->
+ fprintf pp "O%c%ld" (short_name_of_type ty) (camlint_of_coqint ofs)
+
+let current_alloc = ref (None: (var -> loc) option)
+let current_liveness = ref (None: VSet.t PMap.t option)
+
+let reg pp r ty =
+ match !current_alloc with
+ | None -> fprintf pp "x%ld" (P.to_int32 r)
+ | Some alloc -> fprintf pp "x%ld{%a}" (P.to_int32 r) loc (alloc (V(r, ty)))
+
+let var pp = function
+ | V(r, ty) -> reg pp r ty
+ | L l -> loc pp l
+
+let rec vars pp = function
+ | [] -> ()
+ | [r] -> var pp r
+ | r1::rl -> fprintf pp "%a, %a" var r1 vars rl
+
+let ros pp = function
+ | Coq_inl r -> var pp r
+ | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s)
+
+let liveset pp lv =
+ fprintf pp "@[<hov 2>{";
+ VSet.iter (function V(r, ty) -> fprintf pp "@ x%ld" (P.to_int32 r)
+ | L l -> ())
+ lv;
+ fprintf pp " }@]"
+
+let print_succ pp s dfl =
+ let s = P.to_int32 s in
+ if s <> dfl then fprintf pp "goto %ld" s
+
+let print_instruction pp succ = function
+ | Xmove(src, dst) ->
+ fprintf pp "%a = %a;" var dst var src
+ | Xreload(src, dst) ->
+ fprintf pp "%a =r %a;" var dst var src
+ | Xspill(src, dst) ->
+ fprintf pp "%a =s %a;" var dst var src
+ | Xparmove(srcs, dsts, t1, t2) ->
+ fprintf pp "(%a) = (%a) using %a, %a;" vars dsts vars srcs var t1 var t2
+ | Xop(op, args, res) ->
+ fprintf pp "%a = %a;" var res (print_operation var) (op, args)
+ | Xload(chunk, addr, args, dst) ->
+ fprintf pp "%a = %s[%a];"
+ var dst (name_of_chunk chunk) (print_addressing var) (addr, args)
+ | Xstore(chunk, addr, args, src) ->
+ fprintf pp "%s[%a] = %a;"
+ (name_of_chunk chunk) (print_addressing var) (addr, args) var src
+ | Xcall(sg, fn, args, res) ->
+ fprintf pp "%a = call %a(%a);" vars res ros fn vars args
+ | Xtailcall(sg, fn, args) ->
+ fprintf pp "tailcall %a(%a);" ros fn vars args
+ | Xbuiltin(ef, args, res) ->
+ fprintf pp "%a = builtin %s(%a);"
+ vars res (name_of_external ef) vars args
+ | Xbranch s ->
+ print_succ pp s succ
+ | Xcond(cond, args, s1, s2) ->
+ fprintf pp "if (%a) goto %ld else goto %ld"
+ (print_condition var) (cond, args)
+ (P.to_int32 s1) (P.to_int32 s2)
+ | Xjumptable(arg, tbl) ->
+ let tbl = Array.of_list tbl in
+ fprintf pp "@[<v 2>jumptable (%a)" var arg;
+ for i = 0 to Array.length tbl - 1 do
+ fprintf pp "@ case %d: goto %ld" i (P.to_int32 tbl.(i))
+ done;
+ fprintf pp "@]"
+ | Xreturn args ->
+ fprintf pp "return %a" vars args
+
+let rec print_instructions pp succ = function
+ | [] -> ()
+ | [i] -> print_instruction pp succ i
+ | i :: il ->
+ print_instruction pp succ i;
+ fprintf pp "@ ";
+ print_instructions pp succ il
+
+let print_block pp (pc, blk) =
+ fprintf pp "%5ld: @[<hov 0>" pc;
+ print_instructions pp (Int32.pred pc) blk;
+ fprintf pp "@]@ ";
+ match !current_liveness with
+ | None -> ()
+ | Some liveness -> fprintf pp "%a@ " liveset (PMap.get (P.of_int32 pc) liveness)
+
+let print_function pp ?alloc ?live f =
+ current_alloc := alloc;
+ current_liveness := live;
+ fprintf pp "@[<v 2>f() {@ ";
+ let instrs =
+ List.sort
+ (fun (pc1, _) (pc2, _) -> Pervasives.compare pc2 pc1)
+ (List.map
+ (fun (pc, i) -> (P.to_int32 pc, i))
+ (PTree.elements f.fn_code)) in
+ print_succ pp f.fn_entrypoint
+ (match instrs with (pc1, _) :: _ -> pc1 | [] -> -1l);
+ List.iter (print_block pp) instrs;
+ fprintf pp "@;<0 -2>}@]@.";
+ current_alloc := None;
+ current_liveness := None
+
diff --git a/backend/RRE.v b/backend/RRE.v
deleted file mode 100644
index bee57f6..0000000
--- a/backend/RRE.v
+++ /dev/null
@@ -1,173 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Redundant Reloads Elimination *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Op.
-Require Import Locations.
-Require Import Conventions.
-Require Import Linear.
-
-(** * Equations between slots and machine registers *)
-
-(** The RRE pass keeps track of which register holds the value of which
- stack slot, using sets of equations like the following. *)
-
-Record equation := mkeq { e_reg: mreg; e_slot: slot }.
-
-Definition equations : Type := list equation.
-
-Fixpoint find_reg_containing (s: slot) (eqs: equations) : option mreg :=
- match eqs with
- | nil => None
- | e :: eqs' => if slot_eq (e_slot e) s then Some (e_reg e) else find_reg_containing s eqs'
- end.
-
-Definition eq_equation (eq1 eq2: equation) : {eq1=eq2} + {eq1<>eq2}.
-Proof.
- generalize slot_eq mreg_eq. decide equality.
-Defined.
-
-Definition contains_equation (s: slot) (r: mreg) (eqs: equations) : bool :=
- In_dec eq_equation (mkeq r s) eqs.
-
-(** Remove equations that are invalidated by an assignment to location [l]. *)
-
-Fixpoint kill_loc (l: loc) (eqs: equations) : equations :=
- match eqs with
- | nil => nil
- | e :: eqs' =>
- if Loc.diff_dec (S (e_slot e)) l && Loc.diff_dec (R (e_reg e)) l
- then e :: kill_loc l eqs'
- else kill_loc l eqs'
- end.
-
-(** Same, for a list of locations [ll]. *)
-
-Definition kill_locs (ll: list loc) (eqs: equations) : equations :=
- List.fold_left (fun eqs l => kill_loc l eqs) ll eqs.
-
-Definition kill_temps (eqs: equations) : equations :=
- kill_locs temporaries eqs.
-
-Definition kill_at_move (eqs: equations) : equations :=
- kill_locs destroyed_at_move eqs.
-
-Definition kill_op (op: operation) (eqs: equations) : equations :=
- match op with Omove => kill_at_move eqs | _ => kill_temps eqs end.
-
-(** * Safety criterion *)
-
-Definition is_incoming (s: slot) : bool :=
- match s with
- | Incoming _ _ => true
- | _ => false
- end.
-
-(** Turning a [Lgetstack] into a register-to-register move is not always
- safe: at least on x86, the move destroys some registers
- (those from [destroyed_at_move] list) while the [Lgetstack] does not.
- Therefore, we perform this transformation only if the destroyed
- registers are not used before being destroyed by a later
- [Lop], [Lload], [Lstore], [Lbuiltin], [Lcond] or [Ljumptable] operation. *)
-
-Fixpoint safe_move_insertion (c: code) : bool :=
- match c with
- | Lgetstack s r :: c' =>
- negb(In_dec mreg_eq r destroyed_at_move_regs) && safe_move_insertion c'
- | Lsetstack r s :: c' =>
- negb(In_dec mreg_eq r destroyed_at_move_regs)
- | Lop op args res :: c' =>
- list_disjoint_dec mreg_eq args destroyed_at_move_regs
- | Lload chunk addr args res :: c' =>
- list_disjoint_dec mreg_eq args destroyed_at_move_regs
- | Lstore chunk addr args src :: c' =>
- list_disjoint_dec mreg_eq (src :: args) destroyed_at_move_regs
- | Lbuiltin ef args res :: c' =>
- list_disjoint_dec mreg_eq args destroyed_at_move_regs
- | Lcond cond args lbl :: c' =>
- list_disjoint_dec mreg_eq args destroyed_at_move_regs
- | Ljumptable arg tbl :: c' =>
- negb(In_dec mreg_eq arg destroyed_at_move_regs)
- | _ => false
- end.
-
-(** * Code transformation *)
-
-(** The following function eliminates [Lgetstack] instructions or turn them
- into register-to-register move whenever possible. Simultaneously,
- it propagates valid (register, slot) equations across basic blocks. *)
-
-(** [transf_code] is written in accumulator-passing style so that it runs
- in constant stack space. The [k] parameter accumulates the instructions
- to be generated, in reverse order, and is then reversed at the end *)
-
-Fixpoint transf_code (eqs: equations) (c: code) (k: code) : code :=
- match c with
- | nil => List.rev' k
- | Lgetstack s r :: c =>
- if is_incoming s then
- transf_code (kill_loc (R r) (kill_loc (R IT1) eqs)) c (Lgetstack s r :: k)
- else if contains_equation s r eqs then
- transf_code eqs c k
- else
- match find_reg_containing s eqs with
- | Some r' =>
- if safe_move_insertion c then
- transf_code (kill_at_move (mkeq r s :: kill_loc (R r) eqs)) c (Lop Omove (r' :: nil) r :: k)
- else
- transf_code (mkeq r s :: kill_loc (R r) eqs) c (Lgetstack s r :: k)
- | None =>
- transf_code (mkeq r s :: kill_loc (R r) eqs) c (Lgetstack s r :: k)
- end
- | Lsetstack r s :: c =>
- transf_code (kill_at_move (mkeq r s :: kill_loc (S s) eqs)) c (Lsetstack r s :: k)
- | Lop op args res :: c =>
- transf_code (kill_loc (R res) (kill_op op eqs)) c (Lop op args res :: k)
- | Lload chunk addr args res :: c =>
- transf_code (kill_loc (R res) (kill_temps eqs)) c (Lload chunk addr args res :: k)
- | Lstore chunk addr args src :: c =>
- transf_code (kill_temps eqs) c (Lstore chunk addr args src :: k)
- | Lcall sg ros :: c =>
- transf_code nil c (Lcall sg ros :: k)
- | Ltailcall sg ros :: c =>
- transf_code nil c (Ltailcall sg ros :: k)
- | Lbuiltin ef args res :: c =>
- transf_code (kill_loc (R res) (kill_temps eqs)) c (Lbuiltin ef args res :: k)
- | Lannot ef args :: c =>
- transf_code eqs c (Lannot ef args :: k)
- | Llabel lbl :: c =>
- transf_code nil c (Llabel lbl :: k)
- | Lgoto lbl :: c =>
- transf_code nil c (Lgoto lbl :: k)
- | Lcond cond args lbl :: c =>
- transf_code (kill_temps eqs) c (Lcond cond args lbl :: k)
- | Ljumptable arg lbls :: c =>
- transf_code nil c (Ljumptable arg lbls :: k)
- | Lreturn :: c =>
- transf_code nil c (Lreturn :: k)
- end.
-
-Definition transf_function (f: function) : function :=
- mkfunction
- (fn_sig f)
- (fn_stacksize f)
- (transf_code nil (fn_code f) nil).
-
-Definition transf_fundef (fd: fundef) : fundef :=
- transf_fundef transf_function fd.
-
-Definition transf_program (p: program) : program :=
- transform_program transf_fundef p.
-
diff --git a/backend/RREproof.v b/backend/RREproof.v
deleted file mode 100644
index 98de7a8..0000000
--- a/backend/RREproof.v
+++ /dev/null
@@ -1,658 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Correctness proof for the [RRE] pass. *)
-
-Require Import Axioms.
-Require Import Coqlib.
-Require Import AST.
-Require Import Values.
-Require Import Globalenvs.
-Require Import Events.
-Require Import Smallstep.
-Require Import Op.
-Require Import Locations.
-Require Import Conventions.
-Require Import Linear.
-Require Import RRE.
-
-(** * Operations over equations *)
-
-Lemma find_reg_containing_sound:
- forall s r eqs, find_reg_containing s eqs = Some r -> In (mkeq r s) eqs.
-Proof.
- induction eqs; simpl; intros.
- congruence.
- destruct (slot_eq (e_slot a) s). inv H. left; destruct a; auto. right; eauto.
-Qed.
-
-Definition equations_hold (ls: locset) (eqs: equations) : Prop :=
- forall e, In e eqs -> ls (S (e_slot e)) = ls (R (e_reg e)).
-
-Lemma nil_hold:
- forall ls, equations_hold ls nil.
-Proof.
- red; intros; contradiction.
-Qed.
-
-Lemma In_kill_loc:
- forall e l eqs,
- In e (kill_loc l eqs) ->
- In e eqs /\ Loc.diff (S (e_slot e)) l /\ Loc.diff (R (e_reg e)) l.
-Proof.
- induction eqs; simpl kill_loc; simpl In; intros.
- tauto.
- destruct (Loc.diff_dec (S (e_slot a)) l).
- destruct (Loc.diff_dec (R (e_reg a)) l).
- simpl in H. intuition congruence.
- simpl in H. intuition.
- simpl in H. intuition.
-Qed.
-
-Lemma kill_loc_hold:
- forall ls eqs l v,
- equations_hold ls eqs ->
- equations_hold (Locmap.set l v ls) (kill_loc l eqs).
-Proof.
- intros; red; intros.
- exploit In_kill_loc; eauto. intros [A [B C]].
- repeat rewrite Locmap.gso; auto; apply Loc.diff_sym; auto.
-Qed.
-
-Lemma In_kill_locs:
- forall e ll eqs,
- In e (kill_locs ll eqs) ->
- In e eqs /\ Loc.notin (S (e_slot e)) ll /\ Loc.notin (R (e_reg e)) ll.
-Proof.
-Opaque Loc.diff.
- induction ll; simpl; intros.
- tauto.
- exploit IHll; eauto. intros [A [B C]]. exploit In_kill_loc; eauto. intros [D [E F]].
- tauto.
-Qed.
-
-Lemma kill_locs_hold:
- forall ll ls eqs,
- equations_hold ls eqs ->
- equations_hold (Locmap.undef ll ls) (kill_locs ll eqs).
-Proof.
- intros; red; intros. exploit In_kill_locs; eauto. intros [A [B C]].
- repeat rewrite Locmap.guo; auto.
-Qed.
-
-Lemma kill_temps_hold:
- forall ls eqs,
- equations_hold ls eqs ->
- equations_hold (LTL.undef_temps ls) (kill_temps eqs).
-Proof.
- exact (kill_locs_hold temporaries).
-Qed.
-
-Lemma kill_at_move_hold:
- forall ls eqs,
- equations_hold ls eqs ->
- equations_hold (undef_setstack ls) (kill_at_move eqs).
-Proof.
- exact (kill_locs_hold destroyed_at_move).
-Qed.
-
-Lemma kill_at_op_hold:
- forall op ls eqs,
- equations_hold ls eqs ->
- equations_hold (undef_op op ls) (kill_op op eqs).
-Proof.
- intros op.
- destruct op; exact kill_temps_hold || exact kill_at_move_hold.
-Qed.
-
-Lemma eqs_getstack_hold:
- forall rs r s eqs,
- equations_hold rs eqs ->
- equations_hold (Locmap.set (R r) (rs (S s)) rs)
- (mkeq r s :: kill_loc (R r) eqs).
-Proof.
-Transparent Loc.diff.
- intros; red; intros. simpl in H0; destruct H0.
- subst e. simpl. rewrite Locmap.gss; rewrite Locmap.gso; auto. red; auto.
- exploit In_kill_loc; eauto. intros [D [E F]].
- repeat rewrite Locmap.gso. auto.
- apply Loc.diff_sym; auto. apply Loc.diff_sym; auto.
-Qed.
-
-Lemma eqs_movestack_hold:
- forall rs r s eqs,
- equations_hold rs eqs ->
- equations_hold (Locmap.set (R r) (rs (S s)) (undef_setstack rs))
- (kill_at_move (mkeq r s :: kill_loc (R r) eqs)).
-Proof.
- unfold undef_setstack, kill_at_move; intros; red; intros.
- exploit In_kill_locs; eauto. intros [A [B C]].
- simpl in A; destruct A.
- subst e. rewrite Locmap.gss. rewrite Locmap.gso. apply Locmap.guo. auto.
- simpl; auto.
- exploit In_kill_loc; eauto. intros [D [E F]].
- repeat rewrite Locmap.gso. repeat rewrite Locmap.guo; auto.
- apply Loc.diff_sym; auto. apply Loc.diff_sym; auto.
-Qed.
-
-Lemma eqs_setstack_hold:
- forall rs r s eqs,
- equations_hold rs eqs ->
- equations_hold (Locmap.set (S s) (rs (R r)) (undef_setstack rs))
- (kill_at_move (mkeq r s :: kill_loc (S s) eqs)).
-Proof.
- unfold undef_setstack, kill_at_move; intros; red; intros.
- exploit In_kill_locs; eauto. intros [A [B C]].
- simpl in A; destruct A.
- subst e. rewrite Locmap.gss. rewrite Locmap.gso. rewrite Locmap.guo. auto.
- auto. simpl. destruct s; auto.
- exploit In_kill_loc; eauto. intros [D [E F]].
- repeat rewrite Locmap.gso. repeat rewrite Locmap.guo; auto.
- apply Loc.diff_sym; auto. apply Loc.diff_sym; auto.
-Qed.
-
-Lemma locmap_set_reg_same:
- forall rs r,
- Locmap.set (R r) (rs (R r)) rs = rs.
-Proof.
- intros. apply extensionality; intros.
- destruct (Loc.eq x (R r)).
- subst x. apply Locmap.gss.
- apply Locmap.gso. apply Loc.diff_reg_right; auto.
-Qed.
-
-(** * Agreement between values of locations *)
-
-(** Values of locations may differ between the original and transformed
- program: after a [Lgetstack] is optimized to a [Lop Omove],
- the values of [destroyed_at_move] temporaries differ. This
- can only happen in parts of the code where the [safe_move_insertion]
- function returns [true]. *)
-
-Definition agree (sm: bool) (rs rs': locset) : Prop :=
- forall l, sm = false \/ Loc.notin l destroyed_at_move -> rs' l = rs l.
-
-Lemma agree_false:
- forall rs rs',
- agree false rs rs' <-> rs' = rs.
-Proof.
- intros; split; intros.
- apply extensionality; intros. auto.
- subst rs'. red; auto.
-Qed.
-
-Lemma agree_slot:
- forall sm rs rs' s,
- agree sm rs rs' -> rs' (S s) = rs (S s).
-Proof.
-Transparent Loc.diff.
- intros. apply H. right. simpl; destruct s; tauto.
-Qed.
-
-Lemma agree_reg:
- forall sm rs rs' r,
- agree sm rs rs' ->
- sm = false \/ ~In r destroyed_at_move_regs -> rs' (R r) = rs (R r).
-Proof.
- intros. apply H. destruct H0; auto. right.
- simpl in H0; simpl; intuition congruence.
-Qed.
-
-Lemma agree_regs:
- forall sm rs rs' rl,
- agree sm rs rs' ->
- sm = false \/ list_disjoint rl destroyed_at_move_regs -> reglist rs' rl = reglist rs rl.
-Proof.
- induction rl; intros; simpl.
- auto.
- decEq. apply agree_reg with sm. auto.
- destruct H0. auto. right. eapply list_disjoint_notin; eauto with coqlib.
- apply IHrl; auto. destruct H0; auto. right. eapply list_disjoint_cons_left; eauto.
-Qed.
-
-Lemma agree_set:
- forall sm rs rs' l v,
- agree sm rs rs' ->
- agree sm (Locmap.set l v rs) (Locmap.set l v rs').
-Proof.
- intros; red; intros.
- unfold Locmap.set.
- destruct (Loc.eq l l0). auto.
- destruct (Loc.overlap l l0). auto.
- apply H; auto.
-Qed.
-
-Lemma agree_undef_move_1:
- forall sm rs rs',
- agree sm rs rs' ->
- agree true rs (undef_setstack rs').
-Proof.
- intros. unfold undef_setstack. red; intros.
- destruct H0. congruence. rewrite Locmap.guo; auto.
-Qed.
-
-Remark locmap_undef_equal:
- forall x ll rs rs',
- (forall l, Loc.notin l ll -> rs' l = rs l) ->
- Locmap.undef ll rs' x = Locmap.undef ll rs x.
-Proof.
- induction ll; intros; simpl.
- apply H. simpl. auto.
- apply IHll. intros. unfold Locmap.set.
- destruct (Loc.eq a l). auto. destruct (Loc.overlap a l) eqn:?. auto.
- apply H. simpl. split; auto. apply Loc.diff_sym. apply Loc.non_overlap_diff; auto.
-Qed.
-
-Lemma agree_undef_move_2:
- forall sm rs rs',
- agree sm rs rs' ->
- agree false (undef_setstack rs) (undef_setstack rs').
-Proof.
- intros. rewrite agree_false.
- apply extensionality; intros. unfold undef_setstack. apply locmap_undef_equal. auto.
-Qed.
-
-Lemma agree_undef_temps:
- forall sm rs rs',
- agree sm rs rs' ->
- agree false (LTL.undef_temps rs) (LTL.undef_temps rs').
-Proof.
- intros. rewrite agree_false.
- apply extensionality; intros. unfold LTL.undef_temps. apply locmap_undef_equal.
- intros. apply H. right. simpl in H0; simpl; tauto.
-Qed.
-
-Lemma agree_undef_op:
- forall op sm rs rs',
- agree sm rs rs' ->
- agree false (undef_op op rs) (undef_op op rs').
-Proof.
- intros op.
- destruct op; exact agree_undef_temps || exact agree_undef_move_2.
-Qed.
-
-Lemma transf_code_cont:
- forall c eqs k1 k2,
- transf_code eqs c (k1 ++ k2) = List.rev k2 ++ transf_code eqs c k1.
-Proof.
- induction c; simpl; intros.
- unfold rev'; rewrite <- ! rev_alt; apply rev_app_distr.
- destruct a; try (rewrite <- IHc; reflexivity).
- destruct (is_incoming s).
- rewrite <- IHc; reflexivity.
- destruct (contains_equation s m eqs).
- auto.
- destruct (find_reg_containing s eqs).
- destruct (safe_move_insertion c).
- rewrite <- IHc; reflexivity.
- rewrite <- IHc; reflexivity.
- rewrite <- IHc; reflexivity.
-Qed.
-
-Corollary transf_code_eq:
- forall eqs c i, transf_code eqs c (i :: nil) = i :: transf_code eqs c nil.
-Proof.
- intros. change (i :: nil) with (nil ++ (i :: nil)).
- rewrite transf_code_cont. auto.
-Qed.
-
-Lemma transl_find_label:
- forall lbl c eqs,
- find_label lbl (transf_code eqs c nil) =
- option_map (fun c => transf_code nil c nil) (find_label lbl c).
-Proof.
- induction c; intros.
- auto.
- destruct a; simpl; try (rewrite transf_code_eq; simpl; auto).
- destruct (is_incoming s); simpl; auto.
- destruct (contains_equation s m eqs). auto.
- destruct (find_reg_containing s eqs); rewrite !transf_code_eq.
- destruct (safe_move_insertion c); simpl; auto.
- simpl; auto.
- destruct (peq lbl l); simpl; auto.
-Qed.
-
-(** * Semantic preservation *)
-
-Section PRESERVATION.
-
-Variable prog: program.
-Let tprog := transf_program prog.
-
-Let ge := Genv.globalenv prog.
-Let tge := Genv.globalenv tprog.
-
-Lemma functions_translated:
- forall v f,
- 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 function_ptr_translated:
- forall v f,
- Genv.find_funct_ptr ge v = Some f ->
- Genv.find_funct_ptr tge v = Some (transf_fundef f).
-Proof (@Genv.find_funct_ptr_transf _ _ _ transf_fundef prog).
-
-Lemma symbols_preserved:
- forall id,
- Genv.find_symbol tge id = Genv.find_symbol ge id.
-Proof (@Genv.find_symbol_transf _ _ _ transf_fundef prog).
-
-Lemma varinfo_preserved:
- forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
-Proof (@Genv.find_var_info_transf _ _ _ transf_fundef prog).
-
-Lemma sig_preserved:
- forall f, funsig (transf_fundef f) = funsig f.
-Proof.
- destruct f; reflexivity.
-Qed.
-
-Lemma find_function_translated:
- forall ros rs fd,
- find_function ge ros rs = Some fd ->
- find_function tge ros rs = Some (transf_fundef fd).
-Proof.
- intros. destruct ros; simpl in *.
- apply functions_translated; auto.
- rewrite symbols_preserved. destruct (Genv.find_symbol ge i).
- apply function_ptr_translated; auto.
- congruence.
-Qed.
-
-Inductive match_frames: stackframe -> stackframe -> Prop :=
- | match_frames_intro:
- forall f sp rs c,
- match_frames (Stackframe f sp rs c)
- (Stackframe (transf_function f) sp rs (transf_code nil c nil)).
-
-Inductive match_states: state -> state -> Prop :=
- | match_states_regular:
- forall sm stk f sp c rs m stk' rs' eqs
- (STK: list_forall2 match_frames stk stk')
- (EQH: equations_hold rs' eqs)
- (AG: agree sm rs rs')
- (SAFE: sm = false \/ safe_move_insertion c = true),
- match_states (State stk f sp c rs m)
- (State stk' (transf_function f) sp (transf_code eqs c nil) rs' m)
- | match_states_call:
- forall stk f rs m stk'
- (STK: list_forall2 match_frames stk stk'),
- match_states (Callstate stk f rs m)
- (Callstate stk' (transf_fundef f) rs m)
- | match_states_return:
- forall stk rs m stk'
- (STK: list_forall2 match_frames stk stk'),
- match_states (Returnstate stk rs m)
- (Returnstate stk' rs m).
-
-Definition measure (S: state) : nat :=
- match S with
- | State s f sp c rs m => List.length c
- | _ => 0%nat
- end.
-
-Remark match_parent_locset:
- forall stk stk',
- list_forall2 match_frames stk stk' ->
- return_regs (parent_locset stk') = return_regs (parent_locset stk).
-Proof.
- intros. inv H; auto. inv H0; auto.
-Qed.
-
-Theorem 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')
- \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat.
-Proof.
-Opaque destroyed_at_move_regs.
- induction 1; intros; inv MS; simpl.
-(** getstack *)
- simpl in SAFE.
- assert (SAFE': sm = false \/ ~In r destroyed_at_move_regs /\ safe_move_insertion b = true).
- destruct (in_dec mreg_eq r destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- destruct (is_incoming sl) eqn:?.
- (* incoming, stays as getstack *)
- assert (UGS: forall rs, undef_getstack sl rs = Locmap.set (R IT1) Vundef rs).
- destruct sl; simpl in Heqb0; discriminate || auto.
- left; econstructor; split.
- rewrite transf_code_eq; constructor.
- repeat rewrite UGS.
- apply match_states_regular with sm. auto.
- apply kill_loc_hold. apply kill_loc_hold; auto.
- rewrite (agree_slot _ _ _ sl AG). apply agree_set. apply agree_set. auto.
- tauto.
- (* not incoming *)
- assert (UGS: forall rs, undef_getstack sl rs = rs).
- destruct sl; simpl in Heqb0; discriminate || auto.
- unfold contains_equation.
- destruct (in_dec eq_equation (mkeq r sl) eqs); simpl.
- (* eliminated *)
- right. split. omega. split. auto. rewrite UGS.
- exploit EQH; eauto. simpl. intro EQ.
- assert (EQ1: rs' (S sl) = rs (S sl)) by (eapply agree_slot; eauto).
- assert (EQ2: rs' (R r) = rs (R r)) by (eapply agree_reg; eauto; tauto).
- rewrite <- EQ1; rewrite EQ; rewrite EQ2. rewrite locmap_set_reg_same.
- apply match_states_regular with sm; auto; tauto.
- (* found an equation *)
- destruct (find_reg_containing sl eqs) as [r'|] eqn:?.
- exploit EQH. eapply find_reg_containing_sound; eauto.
- simpl; intro EQ.
- (* turned into a move *)
- destruct (safe_move_insertion b) eqn:?.
- left; econstructor; split.
- rewrite transf_code_eq. constructor. simpl; eauto.
- rewrite UGS. rewrite <- EQ.
- apply match_states_regular with true; auto.
- apply eqs_movestack_hold; auto.
- rewrite (agree_slot _ _ _ sl AG). apply agree_set. eapply agree_undef_move_1; eauto.
- (* left as a getstack *)
- left; econstructor; split.
- rewrite transf_code_eq. constructor.
- repeat rewrite UGS.
- apply match_states_regular with sm. auto.
- apply eqs_getstack_hold; auto.
- rewrite (agree_slot _ _ _ sl AG). apply agree_set. auto.
- intuition congruence.
- (* no equation, left as a getstack *)
- left; econstructor; split.
- rewrite transf_code_eq; constructor.
- repeat rewrite UGS.
- apply match_states_regular with sm. auto.
- apply eqs_getstack_hold; auto.
- rewrite (agree_slot _ _ _ sl AG). apply agree_set. auto.
- tauto.
-
-(* setstack *)
- left; econstructor; split. rewrite transf_code_eq; constructor.
- apply match_states_regular with false; auto.
- apply eqs_setstack_hold; auto.
- rewrite (agree_reg _ _ _ r AG). apply agree_set. eapply agree_undef_move_2; eauto.
- simpl in SAFE. destruct (in_dec mreg_eq r destroyed_at_move_regs); simpl in SAFE; intuition congruence.
-
-(* op *)
- left; econstructor; split. rewrite transf_code_eq; constructor.
- instantiate (1 := v). rewrite <- H.
- rewrite (agree_regs _ _ _ args AG).
- apply eval_operation_preserved. exact symbols_preserved.
- simpl in SAFE. destruct (list_disjoint_dec mreg_eq args destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- apply match_states_regular with false; auto.
- apply kill_loc_hold; apply kill_at_op_hold; auto.
- apply agree_set. eapply agree_undef_op; eauto.
-
-(* load *)
- left; econstructor; split.
- rewrite transf_code_eq; econstructor. instantiate (1 := a). rewrite <- H.
- rewrite (agree_regs _ _ _ args AG).
- apply eval_addressing_preserved. exact symbols_preserved.
- simpl in SAFE. destruct (list_disjoint_dec mreg_eq args destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- eauto.
- apply match_states_regular with false; auto.
- apply kill_loc_hold; apply kill_temps_hold; auto.
- apply agree_set. eapply agree_undef_temps; eauto.
-
-(* store *)
-Opaque list_disjoint_dec.
- simpl in SAFE.
- assert (sm = false \/ ~In src destroyed_at_move_regs /\ list_disjoint args destroyed_at_move_regs).
- destruct SAFE. auto. right.
- destruct (list_disjoint_dec mreg_eq (src :: args) destroyed_at_move_regs); try discriminate.
- split. eapply list_disjoint_notin; eauto with coqlib. eapply list_disjoint_cons_left; eauto.
- left; econstructor; split.
- rewrite transf_code_eq; econstructor. instantiate (1 := a). rewrite <- H.
- rewrite (agree_regs _ _ _ args AG).
- apply eval_addressing_preserved. exact symbols_preserved.
- tauto.
- rewrite (agree_reg _ _ _ src AG).
- eauto.
- tauto.
- apply match_states_regular with false; auto.
- apply kill_temps_hold; auto.
- eapply agree_undef_temps; eauto.
-
-(* call *)
- simpl in SAFE. assert (sm = false) by intuition congruence.
- subst sm. rewrite agree_false in AG. subst rs'.
- left; econstructor; split.
- rewrite transf_code_eq; constructor. eapply find_function_translated; eauto.
- symmetry; apply sig_preserved.
- constructor. constructor; auto. constructor.
-
-(* tailcall *)
- simpl in SAFE. assert (sm = false) by intuition congruence.
- subst sm. rewrite agree_false in AG. subst rs'.
- left; econstructor; split.
- rewrite transf_code_eq; constructor. eapply find_function_translated; eauto.
- symmetry; apply sig_preserved. eauto.
- rewrite (match_parent_locset _ _ STK). constructor; auto.
-
-(* builtin *)
- left; econstructor; split.
- rewrite transf_code_eq; constructor.
- rewrite (agree_regs _ _ _ args AG).
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact varinfo_preserved.
- simpl in SAFE. destruct (list_disjoint_dec mreg_eq args destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- apply match_states_regular with false; auto.
- apply kill_loc_hold; apply kill_temps_hold; auto.
- apply agree_set. eapply agree_undef_temps; eauto.
-
-(* annot *)
- simpl in SAFE. assert (sm = false) by intuition congruence.
- subst sm. rewrite agree_false in AG. subst rs'.
- left; econstructor; split.
- rewrite transf_code_eq; econstructor. eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact varinfo_preserved.
- apply match_states_regular with false; auto.
- rewrite agree_false; auto.
-
-(* label *)
- left; econstructor; split. rewrite transf_code_eq; constructor.
- apply match_states_regular with false; auto.
- apply nil_hold.
- simpl in SAFE. destruct SAFE. subst sm. auto. congruence.
-
-(* goto *)
- generalize (transl_find_label lbl (fn_code f) nil). rewrite H. simpl. intros.
- left; econstructor; split. rewrite transf_code_eq; constructor; eauto.
- apply match_states_regular with false; auto.
- apply nil_hold.
- simpl in SAFE. destruct SAFE. subst sm. auto. congruence.
-
-(* cond true *)
- generalize (transl_find_label lbl (fn_code f) nil). rewrite H0. simpl. intros.
- left; econstructor; split.
- rewrite transf_code_eq; apply exec_Lcond_true; auto.
- rewrite (agree_regs _ _ _ args AG). auto.
- simpl in SAFE. destruct (list_disjoint_dec mreg_eq args destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- eauto.
- apply match_states_regular with false; auto.
- apply nil_hold.
- eapply agree_undef_temps; eauto.
-
-(* cond false *)
- left; econstructor; split. rewrite transf_code_eq; apply exec_Lcond_false; auto.
- rewrite (agree_regs _ _ _ args AG). auto.
- simpl in SAFE. destruct (list_disjoint_dec mreg_eq args destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- apply match_states_regular with false; auto.
- apply kill_temps_hold; auto.
- eapply agree_undef_temps; eauto.
-
-(* jumptable *)
- generalize (transl_find_label lbl (fn_code f) nil). rewrite H1. simpl. intros.
- left; econstructor; split. rewrite transf_code_eq; econstructor; eauto.
- rewrite (agree_reg _ _ _ arg AG). auto.
- simpl in SAFE. destruct (in_dec mreg_eq arg destroyed_at_move_regs); simpl in SAFE; intuition congruence.
- apply match_states_regular with false; auto.
- apply nil_hold.
- eapply agree_undef_temps; eauto.
-
-(* return *)
- simpl in SAFE. destruct SAFE; try discriminate. subst sm. rewrite agree_false in AG. subst rs'.
- left; econstructor; split.
- rewrite transf_code_eq; constructor. simpl. eauto.
- rewrite (match_parent_locset _ _ STK).
- constructor; auto.
-
-(* internal *)
- left; econstructor; split.
- constructor. simpl; eauto.
- simpl. apply match_states_regular with false; auto. apply nil_hold. rewrite agree_false; auto.
-
-(* external *)
- left; econstructor; split.
- econstructor. eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact varinfo_preserved.
- auto. eauto.
- constructor; auto.
-
-(* return *)
- inv STK. inv H1. left; econstructor; split. constructor.
- apply match_states_regular with false; auto.
- apply nil_hold.
- rewrite agree_false; 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.
- econstructor; split.
- econstructor.
- apply Genv.init_mem_transf; eauto.
- rewrite symbols_preserved. eauto.
- apply function_ptr_translated; eauto.
- rewrite sig_preserved. auto.
- econstructor; eauto. constructor.
-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 STK. econstructor. auto.
-Qed.
-
-Theorem transf_program_correct:
- forward_simulation (Linear.semantics prog) (Linear.semantics tprog).
-Proof.
- eapply forward_simulation_opt.
- eexact symbols_preserved.
- eexact transf_initial_states.
- eexact transf_final_states.
- eexact transf_step_correct.
-Qed.
-
-End PRESERVATION.
diff --git a/backend/RREtyping.v b/backend/RREtyping.v
deleted file mode 100644
index 170d8ad..0000000
--- a/backend/RREtyping.v
+++ /dev/null
@@ -1,110 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Proof of type preservation for the [RRE] pass. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Locations.
-Require Import Linear.
-Require Import Lineartyping.
-Require Import Conventions.
-Require Import RRE.
-Require Import RREproof.
-
-Remark wt_cons:
- forall f c i, wt_instr f i -> wt_code f c -> wt_code f (i::c).
-Proof.
- intros; red; intros. simpl in H1; destruct H1. congruence. auto.
-Qed.
-
-Hint Constructors wt_instr : linearty.
-Hint Resolve wt_cons: linearty.
-
-Definition wt_eqs (eqs: equations) :=
- forall e, In e eqs -> slot_type (e_slot e) = mreg_type (e_reg e).
-
-Lemma wt_eqs_nil:
- wt_eqs nil.
-Proof. red; simpl; tauto. Qed.
-
-Lemma wt_eqs_cons:
- forall r s eqs,
- slot_type s = mreg_type r -> wt_eqs eqs -> wt_eqs (mkeq r s :: eqs).
-Proof.
- intros; red; simpl; intros. destruct H1. subst; simpl; auto. auto.
-Qed.
-
-Lemma wt_kill_loc:
- forall l eqs,
- wt_eqs eqs -> wt_eqs (kill_loc l eqs).
-Proof.
- intros; red; intros. exploit In_kill_loc; eauto. intros [A B]. auto.
-Qed.
-
-Lemma wt_kill_locs:
- forall ll eqs,
- wt_eqs eqs -> wt_eqs (kill_locs ll eqs).
-Proof.
- intros; red; intros. exploit In_kill_locs; eauto. intros [A B]. auto.
-Qed.
-
-Lemma wt_kill_temps:
- forall eqs, wt_eqs eqs -> wt_eqs (kill_temps eqs).
-Proof.
- exact (wt_kill_locs temporaries).
-Qed.
-
-Lemma wt_kill_at_move:
- forall eqs, wt_eqs eqs -> wt_eqs (kill_at_move eqs).
-Proof.
- exact (wt_kill_locs destroyed_at_move).
-Qed.
-
-Hint Resolve wt_eqs_nil wt_eqs_cons wt_kill_loc wt_kill_locs
- wt_kill_temps wt_kill_at_move: linearty.
-
-Lemma wt_kill_op:
- forall op eqs, wt_eqs eqs -> wt_eqs (kill_op op eqs).
-Proof.
- intros; destruct op; simpl; apply wt_kill_locs; auto.
-Qed.
-
-Hint Resolve wt_kill_op: linearty.
-
-Lemma wt_transf_code:
- forall f c eqs, wt_code f c -> wt_eqs eqs ->
- wt_code (transf_function f) (transf_code eqs c nil).
-Proof.
- induction c; intros; simpl.
- red; simpl; tauto.
- assert (WI: wt_instr f a) by auto with coqlib.
- assert (WC: wt_code f c) by (red; auto with coqlib).
- clear H.
- inv WI; rewrite ? transf_code_eq; auto 10 with linearty.
- destruct (is_incoming s) eqn:?. auto with linearty.
- destruct (contains_equation s r eqs). auto with linearty.
- destruct (find_reg_containing s eqs) as [r'|] eqn:?; auto with linearty.
- assert (mreg_type r' = mreg_type r).
- exploit H0. eapply find_reg_containing_sound; eauto. simpl. congruence.
- rewrite ! transf_code_eq.
- destruct (safe_move_insertion c); auto 10 with linearty.
-Qed.
-
-Lemma program_typing_preserved:
- forall p, wt_program p -> wt_program (transf_program p).
-Proof.
- intros. red; intros. exploit transform_program_function; eauto.
- intros [f0 [A B]]. subst f. exploit H; eauto. intros WTFD.
- inv WTFD; simpl; constructor. red; simpl.
- apply wt_transf_code; auto with linearty.
-Qed.
diff --git a/backend/RTLgen.v b/backend/RTLgen.v
index 007191a..f5e34e4 100644
--- a/backend/RTLgen.v
+++ b/backend/RTLgen.v
@@ -362,18 +362,6 @@ Fixpoint alloc_regs (map: mapping) (al: exprlist)
ret (r :: rl)
end.
-(** A variant of [alloc_regs] for two-address instructions:
- reuse the result register as destination for the first argument. *)
-
-Definition alloc_regs_2addr (map: mapping) (al: exprlist) (rd: reg)
- : mon (list reg) :=
- match al with
- | Enil =>
- ret nil
- | Econs a bl =>
- do rl <- alloc_regs map bl; ret (rd :: rl)
- end.
-
(** [alloc_optreg] is used for function calls. If a destination is
specified for the call, it is returned. Otherwise, a fresh
register is returned. *)
@@ -406,9 +394,7 @@ Fixpoint transl_expr (map: mapping) (a: expr) (rd: reg) (nd: node)
| Evar v =>
do r <- find_var map v; add_move r rd nd
| Eop op al =>
- do rl <- if two_address_op op
- then alloc_regs_2addr map al rd
- else alloc_regs map al;
+ do rl <- alloc_regs map al;
do no <- add_instr (Iop op rl rd nd);
transl_exprlist map al rl no
| Eload chunk addr al =>
@@ -427,6 +413,14 @@ Fixpoint transl_expr (map: mapping) (a: expr) (rd: reg) (nd: node)
transl_expr map b r nc
| Eletvar n =>
do r <- find_letvar map n; add_move r rd nd
+ | Ebuiltin ef al =>
+ do rl <- alloc_regs map al;
+ do no <- add_instr (Ibuiltin ef rl rd nd);
+ transl_exprlist map al rl no
+ | Eexternal id sg al =>
+ do rl <- alloc_regs map al;
+ do no <- add_instr (Icall sg (inr id) rl rd nd);
+ transl_exprlist map al rl no
end
(** Translation of a list of expressions. The expressions are evaluated
@@ -495,16 +489,6 @@ Fixpoint transl_switch (r: reg) (nexits: list node) (t: comptree)
add_instr (Iop op (r :: nil) rt n3)
end.
-(** Detect a two-address operator at the top of an expression. *)
-
-Fixpoint expr_is_2addr_op (e: expr) : bool :=
- match e with
- | Eop op _ => two_address_op op
- | Econdition cond el e2 e3 => expr_is_2addr_op e2 || expr_is_2addr_op e3
- | Elet e1 e2 => expr_is_2addr_op e2
- | _ => false
- end.
-
(** Translation of statements. [transl_stmt map s nd nexits nret rret]
enriches the current CFG with the RTL instructions necessary to
execute the CminorSel statement [s], and returns the node of the first
@@ -524,12 +508,7 @@ Fixpoint transl_stmt (map: mapping) (s: stmt) (nd: node)
ret nd
| Sassign v b =>
do r <- find_var map v;
- if expr_is_2addr_op b then
- do rd <- new_reg;
- do n1 <- add_move rd r nd;
- transl_expr map b rd n1
- else
- transl_expr map b r nd
+ transl_expr map b r nd
| Sstore chunk addr al b =>
do rl <- alloc_regs map al;
do r <- alloc_reg map b;
diff --git a/backend/RTLgenaux.ml b/backend/RTLgenaux.ml
index 363bc2b..150de5a 100644
--- a/backend/RTLgenaux.ml
+++ b/backend/RTLgenaux.ml
@@ -34,6 +34,8 @@ let rec size_expr = function
1 + size_exprs args + max (size_expr e1) (size_expr e2)
| Elet(e1, e2) -> size_expr e1 + size_expr e2
| Eletvar n -> 0
+ | Ebuiltin(ef, el) -> 2 + size_exprs el
+ | Eexternal(id, sg, el) -> 5 + size_exprs el
and size_exprs = function
| Enil -> 0
diff --git a/backend/RTLgenproof.v b/backend/RTLgenproof.v
index 690611c..c3cae28 100644
--- a/backend/RTLgenproof.v
+++ b/backend/RTLgenproof.v
@@ -485,19 +485,18 @@ Section CORRECTNESS_EXPR.
Variable sp: val.
Variable e: env.
-Variable m tm: mem.
-Hypothesis mem_extends: Mem.extends m tm.
+Variable m: mem.
(** The proof of semantic preservation for the translation of expressions
is a simulation argument based on diagrams of the following form:
<<
I /\ P
- e, le, m, a ------------- State cs code sp ns rs m
+ e, le, m, a ------------- State cs code sp ns rs tm
|| |
|| |*
|| |
\/ v
- e, le, m', v ------------ State cs code sp nd rs' m'
+ e, le, m, v ------------ State cs code sp nd rs' tm'
I /\ Q
>>
where [tr_expr code map pr a ns nd rd] is assumed to hold.
@@ -521,27 +520,31 @@ Hypothesis mem_extends: Mem.extends m tm.
Definition transl_expr_prop
(le: letenv) (a: expr) (v: val) : Prop :=
- forall cs f map pr ns nd rd rs dst
+ forall tm cs f map pr ns nd rd rs dst
(MWF: map_wf map)
(TE: tr_expr f.(fn_code) map pr a ns nd rd dst)
- (ME: match_env map e le rs),
- exists rs',
- star step tge (State cs f sp ns rs tm) E0 (State cs f sp nd rs' tm)
+ (ME: match_env map e le rs)
+ (EXT: Mem.extends m tm),
+ exists rs', exists tm',
+ star step tge (State cs f sp ns rs tm) E0 (State cs f sp nd rs' tm')
/\ match_env map (set_optvar dst v e) le rs'
/\ Val.lessdef v rs'#rd
- /\ (forall r, In r pr -> rs'#r = rs#r).
+ /\ (forall r, In r pr -> rs'#r = rs#r)
+ /\ Mem.extends m tm'.
Definition transl_exprlist_prop
(le: letenv) (al: exprlist) (vl: list val) : Prop :=
- forall cs f map pr ns nd rl rs
+ forall tm cs f map pr ns nd rl rs
(MWF: map_wf map)
(TE: tr_exprlist f.(fn_code) map pr al ns nd rl)
- (ME: match_env map e le rs),
- exists rs',
- star step tge (State cs f sp ns rs tm) E0 (State cs f sp nd rs' tm)
+ (ME: match_env map e le rs)
+ (EXT: Mem.extends m tm),
+ exists rs', exists tm',
+ star step tge (State cs f sp ns rs tm) E0 (State cs f sp nd rs' tm')
/\ match_env map e le rs'
/\ Val.lessdef_list vl rs'##rl
- /\ (forall r, In r pr -> rs'#r = rs#r).
+ /\ (forall r, In r pr -> rs'#r = rs#r)
+ /\ Mem.extends m tm'.
(** The correctness of the translation is a huge induction over
the Cminor evaluation derivation for the source program. To keep
@@ -559,21 +562,23 @@ Proof.
intros; red; intros. inv TE.
exploit match_env_find_var; eauto. intro EQ.
exploit tr_move_correct; eauto. intros [rs' [A [B C]]].
- exists rs'; split. eauto.
+ exists rs'; exists tm; split. eauto.
destruct H2 as [[D E] | [D E]].
(* optimized case *)
subst r dst. simpl.
assert (forall r, rs'#r = rs#r).
intros. destruct (Reg.eq r rd). subst r. auto. auto.
split. eapply match_env_invariant; eauto.
- split. congruence. auto.
+ split. congruence.
+ split; auto.
(* general case *)
split.
apply match_env_invariant with (rs#rd <- (rs#r)).
apply match_env_update_dest; auto.
intros. rewrite Regmap.gsspec. destruct (peq r0 rd). congruence. auto.
split. congruence.
- intros. apply C. intuition congruence.
+ split. intros. apply C. intuition congruence.
+ auto.
Qed.
Lemma transl_expr_Eop_correct:
@@ -586,9 +591,9 @@ Lemma transl_expr_Eop_correct:
Proof.
intros; red; intros. inv TE.
(* normal case *)
- exploit H0; eauto. intros [rs1 [EX1 [ME1 [RR1 RO1]]]].
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RR1 [RO1 EXT1]]]]]].
edestruct eval_operation_lessdef as [v' []]; eauto.
- exists (rs1#rd <- v').
+ exists (rs1#rd <- v'); exists tm1.
(* Exec *)
split. eapply star_right. eexact EX1.
eapply exec_Iop; eauto.
@@ -599,7 +604,9 @@ Proof.
(* Result reg *)
split. rewrite Regmap.gss. auto.
(* Other regs *)
- intros. rewrite Regmap.gso. auto. intuition congruence.
+ split. intros. rewrite Regmap.gso. auto. intuition congruence.
+(* Mem *)
+ auto.
Qed.
Lemma transl_expr_Eload_correct:
@@ -612,10 +619,10 @@ Lemma transl_expr_Eload_correct:
transl_expr_prop le (Eload chunk addr args) v.
Proof.
intros; red; intros. inv TE.
- exploit H0; eauto. intros [rs1 [EX1 [ME1 [RES1 OTHER1]]]].
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RES1 [OTHER1 EXT1]]]]]].
edestruct eval_addressing_lessdef as [vaddr' []]; eauto.
edestruct Mem.loadv_extends as [v' []]; eauto.
- exists (rs1#rd <- v').
+ exists (rs1#rd <- v'); exists tm1.
(* Exec *)
split. eapply star_right. eexact EX1. eapply exec_Iload. eauto.
instantiate (1 := vaddr'). rewrite <- H3.
@@ -626,7 +633,9 @@ Proof.
(* Result *)
split. rewrite Regmap.gss. auto.
(* Other regs *)
- intros. rewrite Regmap.gso. auto. intuition congruence.
+ split. intros. rewrite Regmap.gso. auto. intuition congruence.
+(* Mem *)
+ auto.
Qed.
Lemma transl_expr_Econdition_correct:
@@ -640,11 +649,11 @@ Lemma transl_expr_Econdition_correct:
transl_expr_prop le (Econdition cond al ifso ifnot) v.
Proof.
intros; red; intros; inv TE. inv H14.
- exploit H0; eauto. intros [rs1 [EX1 [ME1 [RES1 OTHER1]]]].
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RES1 [OTHER1 EXT1]]]]]].
assert (tr_expr f.(fn_code) map pr (if vcond then ifso else ifnot) (if vcond then ntrue else nfalse) nd rd dst).
destruct vcond; auto.
- exploit H3; eauto. intros [rs2 [EX2 [ME2 [RES2 OTHER2]]]].
- exists rs2.
+ exploit H3; eauto. intros [rs2 [tm2 [EX2 [ME2 [RES2 [OTHER2 EXT2]]]]]].
+ exists rs2; exists tm2.
(* Exec *)
split. eapply star_trans. eexact EX1.
eapply star_left. eapply exec_Icond. eauto. eapply eval_condition_lessdef; eauto. reflexivity.
@@ -654,7 +663,9 @@ Proof.
(* Result value *)
split. assumption.
(* Other regs *)
- intros. transitivity (rs1#r); auto.
+ split. intros. transitivity (rs1#r); auto.
+(* Mem *)
+ auto.
Qed.
Lemma transl_expr_Elet_correct:
@@ -666,12 +677,12 @@ Lemma transl_expr_Elet_correct:
transl_expr_prop le (Elet a1 a2) v2.
Proof.
intros; red; intros; inv TE.
- exploit H0; eauto. intros [rs1 [EX1 [ME1 [RES1 OTHER1]]]].
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RES1 [OTHER1 EXT1]]]]]].
assert (map_wf (add_letvar map r)).
eapply add_letvar_wf; eauto.
exploit H2; eauto. eapply match_env_bind_letvar; eauto.
- intros [rs2 [EX2 [ME3 [RES2 OTHER2]]]].
- exists rs2.
+ intros [rs2 [tm2 [EX2 [ME3 [RES2 [OTHER2 EXT2]]]]]].
+ exists rs2; exists tm2.
(* Exec *)
split. eapply star_trans. eexact EX1. eexact EX2. auto.
(* Match-env *)
@@ -679,7 +690,9 @@ Proof.
(* Result *)
split. assumption.
(* Other regs *)
- intros. transitivity (rs1#r0); auto.
+ split. intros. transitivity (rs1#r0); auto.
+(* Mem *)
+ auto.
Qed.
Lemma transl_expr_Eletvar_correct:
@@ -689,7 +702,7 @@ Lemma transl_expr_Eletvar_correct:
Proof.
intros; red; intros; inv TE.
exploit tr_move_correct; eauto. intros [rs1 [EX1 [RES1 OTHER1]]].
- exists rs1.
+ exists rs1; exists tm.
(* Exec *)
split. eexact EX1.
(* Match-env *)
@@ -706,10 +719,73 @@ Proof.
(* Result *)
split. rewrite RES1. eapply match_env_find_letvar; eauto.
(* Other regs *)
- intros.
+ split. intros.
destruct H2 as [[A B] | [A B]].
destruct (Reg.eq r0 rd); subst; auto.
apply OTHER1. intuition congruence.
+(* Mem *)
+ auto.
+Qed.
+
+Lemma transl_expr_Ebuiltin_correct:
+ forall le ef al vl v,
+ eval_exprlist ge sp e m le al vl ->
+ transl_exprlist_prop le al vl ->
+ external_call ef ge vl m E0 v m ->
+ transl_expr_prop le (Ebuiltin ef al) v.
+Proof.
+ intros; red; intros. inv TE.
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RR1 [RO1 EXT1]]]]]].
+ exploit external_call_mem_extends; eauto.
+ intros [v' [tm2 [A [B [C [D E]]]]]].
+ exists (rs1#rd <- v'); exists tm2.
+(* Exec *)
+ split. eapply star_right. eexact EX1.
+ eapply exec_Ibuiltin; eauto.
+ eapply external_call_symbols_preserved; eauto. exact symbols_preserved. exact varinfo_preserved.
+ reflexivity.
+(* Match-env *)
+ split. eauto with rtlg.
+(* Result reg *)
+ split. rewrite Regmap.gss. auto.
+(* Other regs *)
+ split. intros. rewrite Regmap.gso. auto. intuition congruence.
+(* Mem *)
+ auto.
+Qed.
+
+Lemma transl_expr_Eexternal_correct:
+ forall le id sg al b ef vl v,
+ Genv.find_symbol ge id = Some b ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ ef_sig ef = sg ->
+ eval_exprlist ge sp e m le al vl ->
+ transl_exprlist_prop le al vl ->
+ external_call ef ge vl m E0 v m ->
+ transl_expr_prop le (Eexternal id sg al) v.
+Proof.
+ intros; red; intros. inv TE.
+ exploit H3; eauto. intros [rs1 [tm1 [EX1 [ME1 [RR1 [RO1 EXT1]]]]]].
+ exploit external_call_mem_extends; eauto.
+ intros [v' [tm2 [A [B [C [D E]]]]]].
+ exploit function_ptr_translated; eauto. simpl. intros [tf [P Q]]. inv Q.
+ exists (rs1#rd <- v'); exists tm2.
+(* Exec *)
+ split. eapply star_trans. eexact EX1.
+ eapply star_left. eapply exec_Icall; eauto.
+ simpl. rewrite symbols_preserved. rewrite H. eauto. auto.
+ eapply star_left. eapply exec_function_external.
+ eapply external_call_symbols_preserved; eauto. exact symbols_preserved. exact varinfo_preserved.
+ apply star_one. apply exec_return.
+ reflexivity. reflexivity. reflexivity.
+(* Match-env *)
+ split. eauto with rtlg.
+(* Result reg *)
+ split. rewrite Regmap.gss. auto.
+(* Other regs *)
+ split. intros. rewrite Regmap.gso. auto. intuition congruence.
+(* Mem *)
+ auto.
Qed.
Lemma transl_exprlist_Enil_correct:
@@ -717,7 +793,7 @@ Lemma transl_exprlist_Enil_correct:
transl_exprlist_prop le Enil nil.
Proof.
intros; red; intros; inv TE.
- exists rs.
+ exists rs; exists tm.
split. apply star_refl.
split. assumption.
split. constructor.
@@ -734,9 +810,9 @@ Lemma transl_exprlist_Econs_correct:
transl_exprlist_prop le (Econs a1 al) (v1 :: vl).
Proof.
intros; red; intros; inv TE.
- exploit H0; eauto. intros [rs1 [EX1 [ME1 [RES1 OTHER1]]]].
- exploit H2; eauto. intros [rs2 [EX2 [ME2 [RES2 OTHER2]]]].
- exists rs2.
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RES1 [OTHER1 EXT1]]]]]].
+ exploit H2; eauto. intros [rs2 [tm2 [EX2 [ME2 [RES2 [OTHER2 EXT2]]]]]].
+ exists rs2; exists tm2.
(* Exec *)
split. eapply star_trans. eexact EX1. eexact EX2. auto.
(* Match-env *)
@@ -746,9 +822,11 @@ Proof.
simpl; tauto.
auto.
(* Other regs *)
- intros. transitivity (rs1#r).
+ split. intros. transitivity (rs1#r).
apply OTHER2; auto. simpl; tauto.
apply OTHER1; auto.
+(* Mem *)
+ auto.
Qed.
Theorem transl_expr_correct:
@@ -765,6 +843,8 @@ Proof
transl_expr_Econdition_correct
transl_expr_Elet_correct
transl_expr_Eletvar_correct
+ transl_expr_Ebuiltin_correct
+ transl_expr_Eexternal_correct
transl_exprlist_Enil_correct
transl_exprlist_Econs_correct).
@@ -782,6 +862,8 @@ Proof
transl_expr_Econdition_correct
transl_expr_Elet_correct
transl_expr_Eletvar_correct
+ transl_expr_Ebuiltin_correct
+ transl_expr_Eexternal_correct
transl_exprlist_Enil_correct
transl_exprlist_Econs_correct).
@@ -1019,37 +1101,25 @@ Proof.
(* assign *)
inv TS.
- (* optimized translation (not 2 addr) *)
exploit transl_expr_correct; eauto.
- intros [rs' [A [B [C D]]]].
+ intros [rs' [tm' [A [B [C [D E]]]]]].
econstructor; split.
right; split. eauto. Lt_state.
econstructor; eauto. constructor.
- (* alternate translation (2 addr) *)
- exploit transl_expr_correct; eauto.
- intros [rs' [A [B [C D]]]].
- exploit tr_move_correct; eauto.
- intros [rs'' [P [Q R]]].
- econstructor; split.
- right; split. eapply star_trans. eexact A. eexact P. traceEq. Lt_state.
- econstructor; eauto. constructor.
- simpl in B. apply match_env_invariant with (rs'#r <- (rs'#rd)).
- apply match_env_update_var; auto.
- intros. rewrite Regmap.gsspec. destruct (peq r0 r). congruence. auto.
(* store *)
inv TS.
exploit transl_exprlist_correct; eauto.
- intros [rs' [A [B [C D]]]].
+ intros [rs' [tm' [A [B [C [D E]]]]]].
exploit transl_expr_correct; eauto.
- intros [rs'' [E [F [G J]]]].
+ intros [rs'' [tm'' [F [G [J [K L]]]]]].
assert (Val.lessdef_list vl rs''##rl).
replace (rs'' ## rl) with (rs' ## rl). auto.
- apply list_map_exten. intros. apply J. auto.
+ apply list_map_exten. intros. apply K. auto.
edestruct eval_addressing_lessdef as [vaddr' []]; eauto.
- edestruct Mem.storev_extends as [tm' []]; eauto.
+ edestruct Mem.storev_extends as [tm''' []]; eauto.
econstructor; split.
- left; eapply plus_right. eapply star_trans. eexact A. eexact E. reflexivity.
+ left; eapply plus_right. eapply star_trans. eexact A. eexact F. reflexivity.
eapply exec_Istore with (a := vaddr'). eauto.
rewrite <- H4. apply eval_addressing_preserved. exact symbols_preserved.
eauto. traceEq.
@@ -1059,9 +1129,9 @@ Proof.
inv TS; inv H.
(* indirect *)
exploit transl_expr_correct; eauto.
- intros [rs' [A [B [C D]]]].
+ intros [rs' [tm' [A [B [C [D X]]]]]].
exploit transl_exprlist_correct; eauto.
- intros [rs'' [E [F [G J]]]].
+ intros [rs'' [tm'' [E [F [G [J Y]]]]]].
exploit functions_translated; eauto. intros [tf' [P Q]].
econstructor; split.
left; eapply plus_right. eapply star_trans. eexact A. eexact E. reflexivity.
@@ -1071,7 +1141,7 @@ Proof.
constructor; auto. econstructor; eauto.
(* direct *)
exploit transl_exprlist_correct; eauto.
- intros [rs'' [E [F [G J]]]].
+ intros [rs'' [tm'' [E [F [G [J Y]]]]]].
exploit functions_translated; eauto. intros [tf' [P Q]].
econstructor; split.
left; eapply plus_right. eexact E.
@@ -1085,13 +1155,13 @@ Proof.
inv TS; inv H.
(* indirect *)
exploit transl_expr_correct; eauto.
- intros [rs' [A [B [C D]]]].
+ intros [rs' [tm' [A [B [C [D X]]]]]].
exploit transl_exprlist_correct; eauto.
- intros [rs'' [E [F [G J]]]].
+ intros [rs'' [tm'' [E [F [G [J Y]]]]]].
exploit functions_translated; eauto. intros [tf' [P Q]].
exploit match_stacks_call_cont; eauto. intros [U V].
assert (fn_stacksize tf = fn_stackspace f). inv TF; auto.
- edestruct Mem.free_parallel_extends as [tm' []]; eauto.
+ edestruct Mem.free_parallel_extends as [tm''' []]; eauto.
econstructor; split.
left; eapply plus_right. eapply star_trans. eexact A. eexact E. reflexivity.
eapply exec_Itailcall; eauto. simpl. rewrite J. destruct C. eauto. discriminate P. simpl; auto.
@@ -1101,11 +1171,11 @@ Proof.
constructor; auto.
(* direct *)
exploit transl_exprlist_correct; eauto.
- intros [rs'' [E [F [G J]]]].
+ intros [rs'' [tm'' [E [F [G [J Y]]]]]].
exploit functions_translated; eauto. intros [tf' [P Q]].
exploit match_stacks_call_cont; eauto. intros [U V].
assert (fn_stacksize tf = fn_stackspace f). inv TF; auto.
- edestruct Mem.free_parallel_extends as [tm' []]; eauto.
+ edestruct Mem.free_parallel_extends as [tm''' []]; eauto.
econstructor; split.
left; eapply plus_right. eexact E.
eapply exec_Itailcall; eauto. simpl. rewrite symbols_preserved. rewrite H5.
@@ -1118,8 +1188,8 @@ Proof.
(* builtin *)
inv TS.
exploit transl_exprlist_correct; eauto.
- intros [rs' [E [F [G J]]]].
- edestruct external_call_mem_extends as [tv [tm' [A [B [C D]]]]]; eauto.
+ intros [rs' [tm' [E [F [G [J K]]]]]].
+ edestruct external_call_mem_extends as [tv [tm'' [A [B [C D]]]]]; eauto.
econstructor; split.
left. eapply plus_right. eexact E.
eapply exec_Ibuiltin. eauto.
@@ -1137,7 +1207,7 @@ Proof.
(* ifthenelse *)
inv TS. inv H13.
- exploit transl_exprlist_correct; eauto. intros [rs' [A [B [C D]]]].
+ exploit transl_exprlist_correct; eauto. intros [rs' [tm' [A [B [C [D E]]]]]].
econstructor; split.
left. eapply plus_right. eexact A. eapply exec_Icond; eauto.
eapply eval_condition_lessdef; eauto. traceEq.
@@ -1179,7 +1249,7 @@ Proof.
inv TS.
exploit validate_switch_correct; eauto. intro CTM.
exploit transl_expr_correct; eauto.
- intros [rs' [A [B [C D]]]].
+ intros [rs' [tm' [A [B [C [D X]]]]]].
exploit transl_switch_correct; eauto. inv C. auto.
intros [nd [rs'' [E [F G]]]].
econstructor; split.
@@ -1200,10 +1270,10 @@ Proof.
(* return some *)
inv TS.
exploit transl_expr_correct; eauto.
- intros [rs' [A [B [C D]]]].
+ intros [rs' [tm' [A [B [C [D E]]]]]].
exploit match_stacks_call_cont; eauto. intros [U V].
inversion TF.
- edestruct Mem.free_parallel_extends as [tm' []]; eauto.
+ edestruct Mem.free_parallel_extends as [tm'' []]; eauto.
econstructor; split.
left; eapply plus_right. eexact A. eapply exec_Ireturn; eauto.
rewrite H4; eauto. traceEq.
diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v
index c50c702..d8d5dc8 100644
--- a/backend/RTLgenspec.v
+++ b/backend/RTLgenspec.v
@@ -500,32 +500,6 @@ Proof.
right; eauto with rtlg.
Qed.
-Lemma alloc_regs_2addr_valid:
- forall al rd s1 s2 map rl i,
- map_valid map s1 ->
- reg_valid rd s1 ->
- alloc_regs_2addr map al rd s1 = OK rl s2 i ->
- regs_valid rl s2.
-Proof.
- unfold alloc_regs_2addr; intros.
- destruct al; monadInv H1.
- apply regs_valid_nil.
- apply regs_valid_cons. eauto with rtlg. eauto with rtlg.
-Qed.
-Hint Resolve alloc_regs_2addr_valid: rtlg.
-
-Lemma alloc_regs_2addr_fresh_or_in_map:
- forall map al rd s rl s' i,
- map_valid map s ->
- alloc_regs_2addr map al rd s = OK rl s' i ->
- forall r, In r rl -> r = rd \/ reg_in_map map r \/ reg_fresh r s.
-Proof.
- unfold alloc_regs_2addr; intros.
- destruct al; monadInv H0.
- elim H1.
- simpl in H1; destruct H1. auto. right. eapply alloc_regs_fresh_or_in_map; eauto.
-Qed.
-
(** 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
@@ -628,24 +602,8 @@ Proof.
apply regs_valid_cons; eauto with rtlg.
Qed.
-Lemma alloc_regs_2addr_target_ok:
- forall map al rd pr s1 rl s2 i,
- map_valid map s1 ->
- regs_valid pr s1 ->
- reg_valid rd s1 ->
- ~(reg_in_map map rd) -> ~In rd pr ->
- alloc_regs_2addr map al rd s1 = OK rl s2 i ->
- target_regs_ok map pr al rl.
-Proof.
- unfold alloc_regs_2addr; intros. destruct al; monadInv H4.
- constructor.
- constructor. constructor; auto.
- eapply alloc_regs_target_ok; eauto.
- apply regs_valid_cons; auto.
-Qed.
-
Hint Resolve new_reg_target_ok alloc_reg_target_ok
- alloc_regs_target_ok alloc_regs_2addr_target_ok: rtlg.
+ 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
@@ -760,6 +718,17 @@ Inductive tr_expr (c: code):
((rd = r /\ dst = None) \/ (reg_map_ok map rd dst /\ ~In rd pr)) ->
tr_move c ns r nd rd ->
tr_expr c map pr (Eletvar n) ns nd rd dst
+ | tr_Ebuiltin: forall map pr ef al ns nd rd dst n1 rl,
+ tr_exprlist c map pr al ns n1 rl ->
+ c!n1 = Some (Ibuiltin ef rl rd nd) ->
+ reg_map_ok map rd dst -> ~In rd pr ->
+ tr_expr c map pr (Ebuiltin ef al) ns nd rd dst
+ | tr_Eexternal: forall map pr id sg al ns nd rd dst n1 rl,
+ tr_exprlist c map pr al ns n1 rl ->
+ c!n1 = Some (Icall sg (inr _ id) rl rd nd) ->
+ reg_map_ok map rd dst -> ~In rd pr ->
+ tr_expr c map pr (Eexternal id sg al) ns nd rd dst
+
(** [tr_condition c map pr cond al ns ntrue nfalse rd] holds if the graph [c],
starting at node [ns], contains instructions that compute the truth
@@ -834,13 +803,8 @@ Inductive tr_stmt (c: code) (map: mapping):
| 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 r,
- map.(map_vars)!id = Some r -> expr_is_2addr_op a = false ->
- tr_expr c map nil a ns nd r (Some id) ->
- tr_stmt c map (Sassign id a) ns nd nexits ngoto nret rret
- | tr_Sassign_2: forall id a ns n1 nd nexits ngoto nret rret rd r,
map.(map_vars)!id = Some r ->
- tr_expr c map nil a ns n1 rd None ->
- tr_move c n1 rd nd r ->
+ tr_expr c map nil a ns nd r (Some id) ->
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 ->
@@ -1008,9 +972,7 @@ Proof.
inv OK. left; split; congruence. right; eauto with rtlg.
eapply add_move_charact; eauto.
(* Eop *)
- inv OK. destruct (two_address_op o).
- econstructor; eauto with rtlg.
- eapply transl_exprlist_charact; eauto with rtlg.
+ inv OK.
econstructor; eauto with rtlg.
eapply transl_exprlist_charact; eauto with rtlg.
(* Eload *)
@@ -1045,6 +1007,14 @@ Proof.
inv OK. left; split; congruence. right; eauto with rtlg.
eapply add_move_charact; eauto.
monadInv EQ1.
+ (* Ebuiltin *)
+ inv OK.
+ econstructor; eauto with rtlg.
+ eapply transl_exprlist_charact; eauto with rtlg.
+ (* Eexternal *)
+ inv OK.
+ econstructor; eauto with rtlg.
+ eapply transl_exprlist_charact; eauto with rtlg.
(* Lists *)
@@ -1070,25 +1040,21 @@ Lemma transl_expr_assign_charact:
forall id a map rd nd s ns s' INCR
(TR: transl_expr map a rd nd s = OK ns s' INCR)
(WF: map_valid map s)
- (OK: reg_map_ok map rd (Some id))
- (NOT2ADDR: expr_is_2addr_op a = false),
+ (OK: reg_map_ok map rd (Some id)),
tr_expr s'.(st_code) map nil a ns nd rd (Some id).
Proof.
-Opaque two_address_op.
induction a; intros; monadInv TR; saturateTrans.
(* Evar *)
generalize EQ; unfold find_var. caseEq (map_vars map)!i; intros; inv EQ1.
econstructor; eauto.
eapply add_move_charact; eauto.
(* Eop *)
- simpl in NOT2ADDR. rewrite NOT2ADDR in EQ.
econstructor; eauto with rtlg.
eapply transl_exprlist_charact; eauto with rtlg.
(* Eload *)
econstructor; eauto with rtlg.
eapply transl_exprlist_charact; eauto with rtlg.
(* Econdition *)
- simpl in NOT2ADDR. destruct (orb_false_elim _ _ NOT2ADDR).
econstructor; eauto with rtlg.
econstructor; eauto with rtlg. eapply transl_exprlist_charact; eauto with rtlg.
apply tr_expr_incr with s2; auto.
@@ -1096,7 +1062,6 @@ Opaque two_address_op.
apply tr_expr_incr with s1; auto.
eapply IHa2; eauto 2 with rtlg.
(* Elet *)
- simpl in NOT2ADDR.
econstructor. eapply new_reg_not_in_map; eauto with rtlg.
eapply transl_expr_charact; eauto 3 with rtlg.
apply tr_expr_incr with s1; auto.
@@ -1109,6 +1074,12 @@ Opaque two_address_op.
econstructor; eauto with rtlg.
eapply add_move_charact; eauto.
monadInv EQ1.
+ (* Ebuiltin *)
+ econstructor; eauto with rtlg.
+ eapply transl_exprlist_charact; eauto with rtlg.
+ (* Eexternal *)
+ econstructor; eauto with rtlg.
+ eapply transl_exprlist_charact; eauto with rtlg.
Qed.
Lemma alloc_optreg_map_ok:
@@ -1141,7 +1112,6 @@ Proof.
generalize tr_expr_incr tr_condition_incr tr_exprlist_incr; intros I1 I2 I3.
pose (AT := fun pc i => instr_at_incr s1 s2 pc i EXT).
induction 1; try (econstructor; eauto; fail).
- eapply tr_Sassign_2; eauto. eapply tr_move_incr; eauto.
econstructor; eauto. eapply tr_switch_incr; eauto.
Qed.
@@ -1210,10 +1180,6 @@ Proof.
constructor.
(* Sassign *)
revert EQ. unfold find_var. case_eq (map_vars map)!i; intros; monadInv EQ.
- remember (expr_is_2addr_op e) as is2a. destruct is2a.
- monadInv EQ0. eapply tr_Sassign_2; eauto.
- eapply transl_expr_charact; eauto with rtlg.
- apply tr_move_incr with s1; auto. eapply add_move_charact; eauto.
eapply tr_Sassign; eauto.
eapply transl_expr_assign_charact; eauto with rtlg.
constructor. auto.
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index a1f9518..7ed7344 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -22,6 +22,7 @@ Require Import Registers.
Require Import Globalenvs.
Require Import Values.
Require Import Integers.
+Require Import Memory.
Require Import Events.
Require Import RTL.
Require Import Conventions.
@@ -109,7 +110,6 @@ Inductive wt_instr : instruction -> Prop :=
forall ef args res s,
List.map env args = (ef_sig ef).(sig_args) ->
env res = proj_sig_res (ef_sig ef) ->
- arity_ok (ef_sig ef).(sig_args) = true \/ ef_reloads ef = false ->
valid_successor s ->
wt_instr (Ibuiltin ef args res s)
| wt_Icond:
@@ -282,10 +282,7 @@ Definition type_instr (e: typenv) (i: instruction) : res typenv :=
let sig := ef_sig ef in
do x <- check_successor s;
do e1 <- type_regs e args sig.(sig_args);
- do e2 <- type_reg e1 res (proj_sig_res sig);
- if (negb (ef_reloads ef)) || arity_ok sig.(sig_args)
- then OK e2
- else Error(msg "cannot reload builtin")
+ type_reg e1 res (proj_sig_res sig)
| Icond cond args s1 s2 =>
do x1 <- check_successor s1;
do x2 <- check_successor s2;
@@ -518,9 +515,6 @@ Proof.
destruct (opt_typ_eq (sig_res s) (sig_res (fn_sig f))); try discriminate.
destruct (tailcall_is_possible s) eqn:TCIP; inv EQ2.
eauto with ty.
-- (* builtin *)
- destruct (negb (ef_reloads e0) || arity_ok (sig_args (ef_sig e0))) eqn:E; inv EQ3.
- eauto with ty.
- (* jumptable *)
destruct (zle (list_length_z l * 4) Int.max_unsigned); inv EQ2.
eauto with ty.
@@ -572,11 +566,9 @@ Proof.
eapply type_regs_sound; eauto with ty.
apply tailcall_is_possible_correct; auto.
- (* builtin *)
- destruct (negb (ef_reloads e0) || arity_ok (sig_args (ef_sig e0))) eqn:E; inv EQ3.
constructor.
eapply type_regs_sound; eauto with ty.
eapply type_reg_sound; eauto.
- destruct (ef_reloads e0); auto.
eauto with ty.
- (* cond *)
constructor.
@@ -833,10 +825,7 @@ Proof.
exploit type_reg_complete. eexact B. eauto. intros [e2 [C D]].
exists e2; split; auto.
rewrite check_successor_complete by auto; simpl.
- rewrite A; simpl; rewrite C; simpl.
- destruct H2; rewrite H2.
- rewrite orb_true_r. auto.
- auto.
+ rewrite A; simpl; rewrite C; auto.
- (* cond *)
exploit type_regs_complete. eauto. eauto. intros [e1 [A B]].
exists e1; split; auto.
@@ -982,6 +971,34 @@ Proof.
apply wt_regset_assign; auto.
Qed.
+Lemma wt_exec_Iop:
+ forall (ge: genv) env f sp op args res s rs m v,
+ wt_instr env f (Iop op args res s) ->
+ eval_operation ge sp op rs##args m = Some v ->
+ wt_regset env rs ->
+ wt_regset env (rs#res <- v).
+Proof.
+ intros. inv H.
+ simpl in H0. inv H0. apply wt_regset_assign; auto. rewrite <- H4; apply H1.
+ apply wt_regset_assign; auto.
+ replace (env res) with (snd (type_of_operation op)).
+ eapply type_of_operation_sound; eauto.
+ rewrite <- H7; auto.
+Qed.
+
+Lemma wt_exec_Iload:
+ forall env f chunk addr args dst s m a v rs,
+ wt_instr env f (Iload chunk addr args dst s) ->
+ Mem.loadv chunk m a = Some v ->
+ wt_regset env rs ->
+ wt_regset env (rs#dst <- v).
+Proof.
+ intros. destruct a; simpl in H0; try discriminate.
+ apply wt_regset_assign; auto.
+ inv H. rewrite H8.
+ eapply Mem.load_type; eauto.
+Qed.
+
Inductive wt_stackframes: list stackframe -> option typ -> Prop :=
| wt_stackframes_nil:
wt_stackframes nil (Some Tint)
diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml
new file mode 100644
index 0000000..fe981e3
--- /dev/null
+++ b/backend/Regalloc.ml
@@ -0,0 +1,986 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(* Register allocation by coloring of an interference graph *)
+
+(* The algorithm in a nutshell:
+
+ - Split live ranges
+ - Convert from RTL to XTL
+ - Eliminate dead code
+ - Repeat:
+ . Construct interference graph
+ . Color interference graph using IRC algorithm
+ . Check for variables that were spilled and must be in registers
+ . If none, convert to LTL and exit.
+ . If some, insert spill and reload instructions and try again
+ End Repeat
+*)
+
+open Format
+open Clflags
+open Camlcoq
+open Datatypes
+open Coqlib
+open Maps
+open AST
+open Memdata
+open Kildall
+open Registers
+open Op
+open Machregs
+open Locations
+open Conventions1
+open Conventions
+open IRC
+open XTL
+
+(* Detection of 2-address operations *)
+
+let is_two_address op args =
+ if two_address_op op then
+ match args with
+ | [] -> assert false
+ | arg1 :: argl -> Some(arg1, argl)
+ else None
+
+(* For tracing *)
+
+let destination_alloctrace : string option ref = ref None
+let pp = ref std_formatter
+
+let init_trace () =
+ if !option_dalloctrace && !pp == std_formatter then begin
+ match !destination_alloctrace with
+ | None -> () (* should not happen *)
+ | Some f -> pp := formatter_of_out_channel (open_out f)
+ end
+
+
+(**************** Initial conversion from RTL to XTL **************)
+
+let vreg tyenv r = V(r, tyenv r)
+
+let vregs tyenv rl = List.map (vreg tyenv) rl
+
+let rec expand_regs tyenv = function
+ | [] -> []
+ | r :: rl ->
+ match tyenv r with
+ | Tlong -> V(r, Tint) :: V(twin_reg r, Tint) :: expand_regs tyenv rl
+ | ty -> V(r, ty) :: expand_regs tyenv rl
+
+let constrain_reg v c =
+ match c with
+ | None -> v
+ | Some mr -> L(R mr)
+
+let rec constrain_regs vl cl =
+ match vl, cl with
+ | [], _ -> []
+ | v1 :: vl', [] -> vl
+ | v1 :: vl', Some mr1 :: cl' -> L(R mr1) :: constrain_regs vl' cl'
+ | v1 :: vl', None :: cl' -> v1 :: constrain_regs vl' cl'
+
+let move v1 v2 k =
+ if v1 = v2 then k else Xmove(v1, v2) :: k
+
+let rec movelist vl1 vl2 k =
+ match vl1, vl2 with
+ | [], [] -> k
+ | v1 :: vl1, v2 :: vl2 -> move v1 v2 (movelist vl1 vl2 k)
+ | _, _ -> assert false
+
+let xparmove srcs dsts k =
+ assert (List.length srcs = List.length dsts);
+ match srcs, dsts with
+ | [], [] -> k
+ | [src], [dst] -> Xmove(src, dst) :: k
+ | _, _ -> Xparmove(srcs, dsts, new_temp Tint, new_temp Tfloat) :: k
+
+(* Return the XTL basic block corresponding to the given RTL instruction.
+ Move and parallel move instructions are introduced to honor calling
+ conventions and register constraints on some operations.
+ 64-bit integer variables are split in two 32-bit halves. *)
+
+let block_of_RTL_instr funsig tyenv = function
+ | RTL.Inop s ->
+ [Xbranch s]
+ | RTL.Iop(Omove, [arg], res, s) ->
+ if tyenv arg = Tlong then
+ [Xmove(V(arg, Tint), V(res, Tint));
+ Xmove(V(twin_reg arg, Tint), V(twin_reg res, Tint));
+ Xbranch s]
+ else
+ [Xmove(vreg tyenv arg, vreg tyenv res); Xbranch s]
+ | RTL.Iop(Omakelong, [arg1; arg2], res, s) ->
+ [Xmove(V(arg1, Tint), V(res, Tint));
+ Xmove(V(arg2, Tint), V(twin_reg res, Tint));
+ Xbranch s]
+ | RTL.Iop(Olowlong, [arg], res, s) ->
+ [Xmove(V(twin_reg arg, Tint), V(res, Tint)); Xbranch s]
+ | RTL.Iop(Ohighlong, [arg], res, s) ->
+ [Xmove(V(arg, Tint), V(res, Tint)); Xbranch s]
+ | RTL.Iop(op, args, res, s) ->
+ let (cargs, cres) = mregs_for_operation op in
+ let args1 = vregs tyenv args and res1 = vreg tyenv res in
+ let args2 = constrain_regs args1 cargs and res2 = constrain_reg res1 cres in
+ let (args3, res3) =
+ match is_two_address op args2 with
+ | None ->
+ (args2, res2)
+ | Some(arg, args2') ->
+ if arg = res2 || not (List.mem res2 args2') then
+ (args2, res2)
+ else
+ let t = new_temp (tyenv res) in (t :: args2', t) in
+ movelist args1 args3 (Xop(op, args3, res3) :: move res3 res1 [Xbranch s])
+ | RTL.Iload(chunk, addr, args, dst, s) ->
+ if chunk = Mint64 then begin
+ match offset_addressing addr (coqint_of_camlint 4l) with
+ | None -> assert false
+ | Some addr' ->
+ [Xload(Mint32, addr, vregs tyenv args,
+ V((if big_endian then dst else twin_reg dst), Tint));
+ Xload(Mint32, addr', vregs tyenv args,
+ V((if big_endian then twin_reg dst else dst), Tint));
+ Xbranch s]
+ end else
+ [Xload(chunk, addr, vregs tyenv args, vreg tyenv dst); Xbranch s]
+ | RTL.Istore(chunk, addr, args, src, s) ->
+ if chunk = Mint64 then begin
+ match offset_addressing addr (coqint_of_camlint 4l) with
+ | None -> assert false
+ | Some addr' ->
+ [Xstore(Mint32, addr, vregs tyenv args,
+ V((if big_endian then src else twin_reg src), Tint));
+ Xstore(Mint32, addr', vregs tyenv args,
+ V((if big_endian then twin_reg src else src), Tint));
+ Xbranch s]
+ end else
+ [Xstore(chunk, addr, vregs tyenv args, vreg tyenv src); Xbranch s]
+ | RTL.Icall(sg, ros, args, res, s) ->
+ let args' = vlocs (loc_arguments sg)
+ and res' = vmregs (loc_result sg) in
+ xparmove (expand_regs tyenv args) args'
+ (Xcall(sg, sum_left_map (vreg tyenv) ros, args', res') ::
+ xparmove res' (expand_regs tyenv [res])
+ [Xbranch s])
+ | RTL.Itailcall(sg, ros, args) ->
+ let args' = vlocs (loc_arguments sg) in
+ xparmove (expand_regs tyenv args) args'
+ [Xtailcall(sg, sum_left_map (vreg tyenv) ros, args')]
+ | RTL.Ibuiltin(ef, args, res, s) ->
+ let (cargs, cres) = mregs_for_builtin ef in
+ let args1 = expand_regs tyenv args and res1 = expand_regs tyenv [res] in
+ let args2 = constrain_regs args1 cargs and res2 = constrain_regs res1 cres in
+ movelist args1 args2
+ (Xbuiltin(ef, args2, res2) :: movelist res2 res1 [Xbranch s])
+ | RTL.Icond(cond, args, s1, s2) ->
+ [Xcond(cond, vregs tyenv args, s1, s2)]
+ | RTL.Ijumptable(arg, tbl) ->
+ [Xjumptable(vreg tyenv arg, tbl)]
+ | RTL.Ireturn None ->
+ [Xreturn []]
+ | RTL.Ireturn (Some arg) ->
+ let args' = vmregs (loc_result funsig) in
+ xparmove (expand_regs tyenv [arg]) args' [Xreturn args']
+
+(* One above the [pc] nodes of the given RTL function *)
+
+let next_pc f =
+ PTree.fold
+ (fun npc pc i -> if P.lt pc npc then npc else P.succ pc)
+ f.RTL.fn_code P.one
+
+(* Translate an RTL function to an XTL function *)
+
+let function_of_RTL_function f tyenv =
+ let xc = PTree.map1 (block_of_RTL_instr f.RTL.fn_sig tyenv) f.RTL.fn_code in
+ (* Add moves for function parameters *)
+ let pc_entrypoint = next_pc f in
+ let b_entrypoint =
+ xparmove (vlocs (loc_parameters f.RTL.fn_sig))
+ (expand_regs tyenv f.RTL.fn_params)
+ [Xbranch f.RTL.fn_entrypoint] in
+ { fn_sig = f.RTL.fn_sig;
+ fn_stacksize = f.RTL.fn_stacksize;
+ fn_entrypoint = pc_entrypoint;
+ fn_code = PTree.set pc_entrypoint b_entrypoint xc }
+
+
+(***************** Liveness analysis *****************)
+
+let vset_removelist vl after = List.fold_right VSet.remove vl after
+let vset_addlist vl after = List.fold_right VSet.add vl after
+let vset_addros vos after =
+ match vos with Coq_inl v -> VSet.add v after | Coq_inr id -> after
+
+let live_before instr after =
+ match instr with
+ | Xmove(src, dst) | Xspill(src, dst) | Xreload(src, dst) ->
+ if VSet.mem dst after
+ then VSet.add src (VSet.remove dst after)
+ else after
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ vset_addlist srcs (vset_removelist dsts after)
+ | Xop(op, args, res) ->
+ if VSet.mem res after
+ then vset_addlist args (VSet.remove res after)
+ else after
+ | Xload(chunk, addr, args, dst) ->
+ if VSet.mem dst after
+ then vset_addlist args (VSet.remove dst after)
+ else after
+ | Xstore(chunk, addr, args, src) ->
+ vset_addlist args (VSet.add src after)
+ | Xcall(sg, ros, args, res) ->
+ vset_addlist args (vset_addros ros (vset_removelist res after))
+ | Xtailcall(sg, ros, args) ->
+ vset_addlist args (vset_addros ros VSet.empty)
+ | Xbuiltin(ef, args, res) ->
+ vset_addlist args (vset_removelist res after)
+ | Xbranch s ->
+ after
+ | Xcond(cond, args, s1, s2) ->
+ List.fold_right VSet.add args after
+ | Xjumptable(arg, tbl) ->
+ VSet.add arg after
+ | Xreturn args ->
+ vset_addlist args VSet.empty
+
+let rec live_before_block blk after =
+ match blk with
+ | [] -> after
+ | instr :: blk -> live_before instr (live_before_block blk after)
+
+let transfer_live f pc after =
+ match PTree.get pc f.fn_code with
+ | None -> VSet.empty
+ | Some blk -> live_before_block blk after
+
+module VSetLat = struct
+ type t = VSet.t
+ let beq = VSet.equal
+ let bot = VSet.empty
+ let lub = VSet.union
+end
+
+module Liveness_Solver = Backward_Dataflow_Solver(VSetLat)(NodeSetBackward)
+
+let liveness_analysis f =
+ match Liveness_Solver.fixpoint (successors f) (transfer_live f) [] with
+ | None -> assert false
+ | Some lv -> lv
+
+(* Pair the instructions of a block with their live-before sets *)
+
+let pair_block_live blk after =
+ let rec pair_rec accu after = function
+ | [] -> accu
+ | instr :: blk ->
+ let before = live_before instr after in
+ pair_rec ((instr, before) :: accu) before blk in
+ pair_rec [] after (List.rev blk)
+
+
+(**************** Dead code elimination **********************)
+
+(* Eliminate pure instructions whose results are not used later. *)
+
+let rec dce_parmove srcs dsts after =
+ match srcs, dsts with
+ | [], [] -> [], []
+ | src1 :: srcs, dst1 :: dsts ->
+ let (srcs', dsts') = dce_parmove srcs dsts after in
+ if VSet.mem dst1 after
+ then (src1 :: srcs', dst1 :: dsts')
+ else (srcs', dsts')
+ | _, _ -> assert false
+
+let dce_instr instr after k =
+ match instr with
+ | Xmove(src, dst) ->
+ if VSet.mem dst after
+ then instr :: k
+ else k
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ let (srcs', dsts') = dce_parmove srcs dsts after in
+ Xparmove(srcs', dsts', itmp, ftmp) :: k
+ | Xop(op, args, res) ->
+ if VSet.mem res after
+ then instr :: k
+ else k
+ | Xload(chunk, addr, args, dst) ->
+ if VSet.mem dst after
+ then instr :: k
+ else k
+ | _ ->
+ instr :: k
+
+let rec dce_block blk after =
+ match blk with
+ | [] -> (after, [])
+ | instr :: blk' ->
+ let (after', tblk') = dce_block blk' after in
+ (live_before instr after', dce_instr instr after' tblk')
+
+let dead_code_elimination f liveness =
+ { f with fn_code =
+ PTree.map (fun pc blk -> snd(dce_block blk (PMap.get pc liveness)))
+ f.fn_code }
+
+
+(*********************** Spill costs ****************************)
+
+(* Estimate spill costs and count other statistics for every variable.
+ Variables that must not be spilled are given infinite costs. *)
+
+let spill_costs f =
+
+ let costs = ref PTree.empty in
+
+ let get_stats r =
+ match PTree.get r !costs with
+ | Some st -> st
+ | None ->
+ let st = {cost = 0; usedefs = 0} in
+ costs := PTree.set r st !costs;
+ st in
+
+ let charge amount uses v =
+ match v with
+ | L l -> ()
+ | V(r, ty) ->
+ let st = get_stats r in
+ let c1 = st.cost + amount in
+ let c2 = if c1 >= 0 then c1 else max_int (* overflow *) in
+ st.cost <- c2;
+ st.usedefs <- st.usedefs + uses in
+
+ let charge_list amount uses vl =
+ List.iter (charge amount uses) vl in
+
+ let charge_ros amount ros =
+ match ros with Coq_inl v -> charge amount 1 v | Coq_inr id -> () in
+
+ let charge_instr = function
+ | Xmove(src, dst) ->
+ charge 1 1 src; charge 1 1 dst
+ | Xreload(src, dst) ->
+ charge 1 1 src; charge max_int 1 dst
+ (* dest must not be spilled! *)
+ | Xspill(src, dst) ->
+ charge max_int 1 src; charge 1 1 dst
+ (* source must not be spilled! *)
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ charge_list 1 1 srcs; charge_list 1 1 dsts;
+ charge max_int 0 itmp; charge max_int 0 ftmp
+ (* temps must not be spilled *)
+ | Xop(op, args, res) ->
+ charge_list 10 1 args; charge 10 1 res
+ | Xload(chunk, addr, args, dst) ->
+ charge_list 10 1 args; charge 10 1 dst
+ | Xstore(chunk, addr, args, src) ->
+ charge_list 10 1 args; charge 10 1 src
+ | Xcall(sg, vos, args, res) ->
+ charge_ros 10 vos
+ | Xtailcall(sg, vos, args) ->
+ charge_ros 10 vos
+ | Xbuiltin(ef, args, res) ->
+ begin match ef with
+ | EF_vstore _ | EF_vstore_global _ | EF_memcpy _ ->
+ (* result is not used but should not be spilled *)
+ charge_list 10 1 args; charge_list max_int 0 res
+ | EF_annot _ ->
+ (* arguments are not actually used, so don't charge;
+ result is never used but should not be spilled *)
+ charge_list max_int 0 res
+ | EF_annot_val _ ->
+ (* like a move *)
+ charge_list 1 1 args; charge_list 1 1 res
+ | _ ->
+ charge_list 10 1 args; charge_list 10 1 res
+ end
+ | Xbranch _ -> ()
+ | Xcond(cond, args, _, _) ->
+ charge_list 10 1 args
+ | Xjumptable(arg, _) ->
+ charge 10 1 arg
+ | Xreturn optarg ->
+ () in
+
+ let charge_block blk = List.iter charge_instr blk in
+
+ PTree.fold
+ (fun () pc blk -> charge_block blk)
+ f.fn_code ();
+ if !option_dalloctrace then begin
+ fprintf !pp "------------------ Unspillable variables --------------@ @.";
+ fprintf !pp "@[<hov 1>";
+ PTree.fold
+ (fun () r st ->
+ if st.cost = max_int then fprintf !pp "@ x%ld" (P.to_int32 r))
+ !costs ();
+ fprintf !pp "@]@ @."
+ end;
+ (* Result is cost function: pseudoreg -> stats *)
+ get_stats
+
+
+(********* Construction and coloring of the interference graph **************)
+
+let add_interfs_def g res live =
+ VSet.iter (fun v -> if v <> res then IRC.add_interf g v res) live
+
+let add_interfs_move g src dst live =
+ VSet.iter (fun v -> if v <> src && v <> dst then IRC.add_interf g v dst) live
+
+let add_interfs_destroyed g live mregs =
+ List.iter
+ (fun mr -> VSet.iter (IRC.add_interf g (L (R mr))) live)
+ mregs
+
+let add_interfs_live g live v =
+ VSet.iter (fun v' -> IRC.add_interf g v v') live
+
+let add_interfs_list g v vl =
+ List.iter (IRC.add_interf g v) vl
+
+let rec add_interfs_pairwise g = function
+ | [] -> ()
+ | v1 :: vl -> add_interfs_list g v1 vl; add_interfs_pairwise g vl
+
+let add_interfs_instr g instr live =
+ match instr with
+ | Xmove(src, dst) | Xspill(src, dst) | Xreload(src, dst) ->
+ IRC.add_pref g src dst;
+ add_interfs_move g src dst live
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ List.iter2 (IRC.add_pref g) srcs dsts;
+ (* Interferences with live across *)
+ let across = vset_removelist dsts live in
+ List.iter (add_interfs_live g across) dsts;
+ add_interfs_live g across itmp; add_interfs_live g across ftmp;
+ (* All destinations must be pairwise different *)
+ add_interfs_pairwise g dsts;
+ (* The temporaries must be different from sources and dests *)
+ add_interfs_list g itmp srcs; add_interfs_list g itmp dsts;
+ add_interfs_list g ftmp srcs; add_interfs_list g ftmp dsts;
+ (* Take into account destroyed reg when accessing Incoming param *)
+ if List.exists (function (L(S(Incoming, _, _))) -> true | _ -> false) srcs
+ then add_interfs_list g (vmreg temp_for_parent_frame) dsts
+ | Xop(op, args, res) ->
+ begin match is_two_address op args with
+ | None ->
+ add_interfs_def g res live;
+ | Some(arg1, argl) ->
+ (* Treat as "res := arg1; res := op(res, argl)" *)
+ add_interfs_def g res live;
+ IRC.add_pref g arg1 res;
+ add_interfs_move g arg1 res
+ (vset_addlist (res :: argl) (VSet.remove res live))
+ end;
+ add_interfs_destroyed g (VSet.remove res live) (destroyed_by_op op);
+ | Xload(chunk, addr, args, dst) ->
+ add_interfs_def g dst live;
+ add_interfs_destroyed g (VSet.remove dst live)
+ (destroyed_by_load chunk addr)
+ | Xstore(chunk, addr, args, src) ->
+ add_interfs_destroyed g live (destroyed_by_store chunk addr)
+ | Xcall(sg, vos, args, res) ->
+ add_interfs_destroyed g (vset_removelist res live) destroyed_at_call
+ | Xtailcall(sg, Coq_inl v, args) ->
+ List.iter (fun r -> IRC.add_interf g (vmreg r) v) int_callee_save_regs
+ | Xtailcall(sg, Coq_inr id, args) ->
+ ()
+ | Xbuiltin(ef, args, res) ->
+ (* Interferences with live across *)
+ let across = vset_removelist res live in
+ List.iter (add_interfs_live g across) res;
+ (* All results must be pairwise different *)
+ add_interfs_pairwise g res;
+ add_interfs_destroyed g across (destroyed_by_builtin ef);
+ begin match ef, args, res with
+ | EF_annot_val _, [arg], [res] -> IRC.add_pref g arg res (* like a move *)
+ | _ -> ()
+ end
+ | Xbranch s ->
+ ()
+ | Xcond(cond, args, s1, s2) ->
+ add_interfs_destroyed g live (destroyed_by_cond cond)
+ | Xjumptable(arg, tbl) ->
+ add_interfs_destroyed g live destroyed_by_jumptable
+ | Xreturn optarg ->
+ ()
+
+let rec add_interfs_block g blk live =
+ match blk with
+ | [] -> live
+ | instr :: blk' ->
+ let live' = add_interfs_block g blk' live in
+ add_interfs_instr g instr live';
+ live_before instr live'
+
+let find_coloring f liveness =
+ (*type_function f; (* for debugging *)*)
+ let g = IRC.init (spill_costs f) in
+ PTree.fold
+ (fun () pc blk -> ignore (add_interfs_block g blk (PMap.get pc liveness)))
+ f.fn_code ();
+ IRC.coloring g
+
+
+(*********** Determination of variables that need spill code insertion *****)
+
+let is_reg alloc v =
+ match alloc v with R _ -> true | S _ -> false
+
+let add_tospill alloc v ts =
+ match alloc v with R _ -> ts | S _ -> VSet.add v ts
+
+let addlist_tospill alloc vl ts =
+ List.fold_right (add_tospill alloc) vl ts
+
+let addros_tospill alloc ros ts =
+ match ros with Coq_inl v -> add_tospill alloc v ts | Coq_inr s -> ts
+
+let tospill_instr alloc instr ts =
+ match instr with
+ | Xmove(src, dst) ->
+ if is_reg alloc src || is_reg alloc dst || alloc src = alloc dst
+ then ts
+ else VSet.add src (VSet.add dst ts)
+ | Xreload(src, dst) ->
+ assert (is_reg alloc dst);
+ ts
+ | Xspill(src, dst) ->
+ assert (is_reg alloc src);
+ ts
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ assert (is_reg alloc itmp && is_reg alloc ftmp);
+ ts
+ | Xop(op, args, res) ->
+ addlist_tospill alloc args (add_tospill alloc res ts)
+ | Xload(chunk, addr, args, dst) ->
+ addlist_tospill alloc args (add_tospill alloc dst ts)
+ | Xstore(chunk, addr, args, src) ->
+ addlist_tospill alloc args (add_tospill alloc src ts)
+ | Xcall(sg, vos, args, res) ->
+ addros_tospill alloc vos ts
+ | Xtailcall(sg, vos, args) ->
+ addros_tospill alloc vos ts
+ | Xbuiltin(ef, args, res) ->
+ begin match ef with
+ | EF_annot _ -> ts
+ | _ -> addlist_tospill alloc args (addlist_tospill alloc res ts)
+ end
+ | Xbranch s ->
+ ts
+ | Xcond(cond, args, s1, s2) ->
+ addlist_tospill alloc args ts
+ | Xjumptable(arg, tbl) ->
+ add_tospill alloc arg ts
+ | Xreturn optarg ->
+ ts
+
+let rec tospill_block alloc blk ts =
+ match blk with
+ | [] -> ts
+ | instr :: blk' -> tospill_block alloc blk' (tospill_instr alloc instr ts)
+
+let tospill_function f alloc =
+ PTree.fold
+ (fun ts pc blk -> tospill_block alloc blk ts)
+ f.fn_code VSet.empty
+
+
+(********************* Spilling ***********************)
+
+(* We follow a semi-naive spilling strategy. By default, we spill at
+ every definition of a variable that could not be allocated a register,
+ and we reload before every use. However, we also maintain a list of
+ equations of the form [spilled-var = temp] that keep track of
+ variables that were recently spilled or reloaded. Based on these
+ equations, we can avoid reloading a spilled variable if its value
+ is still available in a temporary register. *)
+
+let rec find_reg_containing v = function
+ | [] -> None
+ | (var, temp, date) :: eqs ->
+ if var = v then Some temp else find_reg_containing v eqs
+
+let add v t eqs = (v, t, 0) :: eqs
+
+let kill x eqs =
+ List.filter (fun (v, t, date) -> v <> x && t <> x) eqs
+
+let reload_var tospill eqs v =
+ if not (VSet.mem v tospill) then
+ (v, [], eqs)
+ else
+ match find_reg_containing v eqs with
+ | Some t ->
+ (*printf "Reusing %a for %a@ @." PrintXTL.var t PrintXTL.var v;*)
+ (t, [], eqs)
+ | None ->
+ let t = new_temp (typeof v) in (t, [Xreload(v, t)], add v t eqs)
+
+let rec reload_vars tospill eqs vl =
+ match vl with
+ | [] -> ([], [], eqs)
+ | v1 :: vs ->
+ let (t1, c1, eqs1) = reload_var tospill eqs v1 in
+ let (ts, cs, eqs2) = reload_vars tospill eqs1 vs in
+ (t1 :: ts, c1 @ cs, eqs2)
+
+let save_var tospill eqs v =
+ if not (VSet.mem v tospill) then
+ (v, [], kill v eqs)
+ else begin
+ let t = new_temp (typeof v) in
+ (t, [Xspill(t, v)], add v t (kill v eqs))
+ end
+
+let rec save_vars tospill eqs vl =
+ match vl with
+ | [] -> ([], [], eqs)
+ | v1 :: vs ->
+ let (t1, c1, eqs1) = save_var tospill eqs v1 in
+ let (ts, cs, eqs2) = save_vars tospill eqs1 vs in
+ (t1 :: ts, c1 @ cs, eqs2)
+
+(* Trimming equations when we have too many or when they are too old.
+ The goal is to limit the live range of unspillable temporaries.
+ By setting [max_age] to zero, we can effectively deactivate
+ the reuse strategy and fall back to a naive "reload at every use"
+ strategy. *)
+
+let max_age = ref 3
+let max_num_eqs = ref 3
+
+let rec trim count eqs =
+ if count <= 0 then [] else
+ match eqs with
+ | [] -> []
+ | (v, t, date) :: eqs' ->
+ if date <= !max_age
+ then (v, t, date + 1) :: trim (count - 1) eqs'
+ else []
+
+(* Insertion of spill and reload instructions. *)
+
+let spill_instr tospill eqs instr =
+ let eqs = trim !max_num_eqs eqs in
+ match instr with
+ | Xmove(src, dst) ->
+ if VSet.mem src tospill && VSet.mem dst tospill then begin
+ let (src', c1, eqs1) = reload_var tospill eqs src in
+ (c1 @ [Xspill(src', dst)], add dst src' (kill dst eqs1))
+ end else begin
+ ([instr], kill dst eqs)
+ end
+ | Xreload(src, dst) ->
+ assert false
+ | Xspill(src, dst) ->
+ assert false
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ ([instr], List.fold_right kill dsts eqs)
+ | Xop(op, args, res) ->
+ begin match is_two_address op args with
+ | None ->
+ let (args', c1, eqs1) = reload_vars tospill eqs args in
+ let (res', c2, eqs2) = save_var tospill eqs1 res in
+ (c1 @ Xop(op, args', res') :: c2, eqs2)
+ | Some(arg1, argl) ->
+ begin match VSet.mem res tospill, VSet.mem arg1 tospill with
+ | false, false ->
+ let (argl', c1, eqs1) = reload_vars tospill eqs argl in
+ (c1 @ [Xop(op, arg1 :: argl', res)], kill res eqs1)
+ | true, false ->
+ let tmp = new_temp (typeof res) in
+ let (argl', c1, eqs1) = reload_vars tospill eqs argl in
+ (c1 @ [Xmove(arg1, tmp); Xop(op, tmp :: argl', tmp); Xspill(tmp, res)],
+ add res tmp (kill res eqs1))
+ | false, true ->
+ let eqs1 = add arg1 res (kill res eqs) in
+ let (argl', c1, eqs2) = reload_vars tospill eqs1 argl in
+ (Xreload(arg1, res) :: c1 @ [Xop(op, res :: argl', res)],
+ kill res eqs2)
+ | true, true ->
+ let tmp = new_temp (typeof res) in
+ let eqs1 = add arg1 tmp eqs in
+ let (argl', c1, eqs2) = reload_vars tospill eqs1 argl in
+ (Xreload(arg1, tmp) :: c1 @ [Xop(op, tmp :: argl', tmp); Xspill(tmp, res)],
+ add res tmp (kill tmp (kill res eqs2)))
+ end
+ end
+ | Xload(chunk, addr, args, dst) ->
+ let (args', c1, eqs1) = reload_vars tospill eqs args in
+ let (dst', c2, eqs2) = save_var tospill eqs1 dst in
+ (c1 @ Xload(chunk, addr, args', dst') :: c2, eqs2)
+ | Xstore(chunk, addr, args, src) ->
+ let (args', c1, eqs1) = reload_vars tospill eqs args in
+ let (src', c2, eqs2) = reload_var tospill eqs1 src in
+ (c1 @ c2 @ [Xstore(chunk, addr, args', src')], eqs2)
+ | Xcall(sg, Coq_inl v, args, res) ->
+ let (v', c1, eqs1) = reload_var tospill eqs v in
+ (c1 @ [Xcall(sg, Coq_inl v', args, res)], [])
+ | Xcall(sg, Coq_inr id, args, res) ->
+ ([instr], [])
+ | Xtailcall(sg, Coq_inl v, args) ->
+ let (v', c1, eqs1) = reload_var tospill eqs v in
+ (c1 @ [Xtailcall(sg, Coq_inl v', args)], [])
+ | Xtailcall(sg, Coq_inr id, args) ->
+ ([instr], [])
+ | Xbuiltin(ef, args, res) ->
+ begin match ef with
+ | EF_annot _ ->
+ ([instr], eqs)
+ | _ ->
+ let (args', c1, eqs1) = reload_vars tospill eqs args in
+ let (res', c2, eqs2) = save_vars tospill eqs1 res in
+ (c1 @ Xbuiltin(ef, args', res') :: c2, eqs2)
+ end
+ | Xbranch s ->
+ ([instr], eqs)
+ | Xcond(cond, args, s1, s2) ->
+ let (args', c1, eqs1) = reload_vars tospill eqs args in
+ (c1 @ [Xcond(cond, args', s1, s2)], eqs1)
+ | Xjumptable(arg, tbl) ->
+ let (arg', c1, eqs1) = reload_var tospill eqs arg in
+ (c1 @ [Xjumptable(arg', tbl)], eqs1)
+ | Xreturn optarg ->
+ ([instr], [])
+
+let rec spill_block tospill pc blk eqs =
+ match blk with
+ | [] -> ([], eqs)
+ | instr :: blk' ->
+ let (c1, eqs1) = spill_instr tospill eqs instr in
+ let (c2, eqs2) = spill_block tospill pc blk' eqs1 in
+ (c1 @ c2, eqs2)
+
+(*
+let spill_block tospill pc blk eqs =
+ printf "@[<hov 2>spill_block: at %ld: " (camlint_of_positive pc);
+ List.iter (fun (x,y,d) -> printf "@ %a=%a" PrintXTL.var x PrintXTL.var y) eqs;
+ printf "@]@.";
+ spill_block tospill pc blk eqs
+*)
+
+let spill_function f tospill round =
+ max_num_eqs := 3;
+ max_age := (if round <= 10 then 3 else if round <= 20 then 1 else 0);
+ transform_basic_blocks (spill_block tospill) [] f
+
+
+(***************** Generation of LTL from XTL ***********************)
+
+(** Apply a register allocation to an XTL function, producing an LTL function.
+ Raise [Bad_LTL] if some pseudoregisters were mapped to stack locations
+ while machine registers were expected, or in other words if spilling
+ and reloading code must be inserted. *)
+
+exception Bad_LTL
+
+let mreg_of alloc v = match alloc v with R mr -> mr | S _ -> raise Bad_LTL
+
+let mregs_of alloc vl = List.map (mreg_of alloc) vl
+
+let mros_of alloc vos = sum_left_map (mreg_of alloc) vos
+
+let make_move src dst k =
+ match src, dst with
+ | R rsrc, R rdst ->
+ if rsrc = rdst then k else LTL.Lop(Omove, [rsrc], rdst) :: k
+ | R rsrc, S(sl, ofs, ty) ->
+ LTL.Lsetstack(rsrc, sl, ofs, ty) :: k
+ | S(sl, ofs, ty), R rdst ->
+ LTL.Lgetstack(sl, ofs, ty, rdst) :: k
+ | S _, S _ ->
+ if src = dst then k else raise Bad_LTL
+
+type parmove_status = To_move | Being_moved | Moved
+
+let make_parmove srcs dsts itmp ftmp k =
+ let src = Array.of_list srcs
+ and dst = Array.of_list dsts in
+ let n = Array.length src in
+ assert (Array.length dst = n);
+ let status = Array.make n To_move in
+ let temp_for =
+ function Tint -> itmp | Tfloat -> ftmp | Tlong -> assert false in
+ let code = ref [] in
+ let add_move s d =
+ match s, d with
+ | R rs, R rd ->
+ code := LTL.Lop(Omove, [rs], rd) :: !code
+ | R rs, S(sl, ofs, ty) ->
+ code := LTL.Lsetstack(rs, sl, ofs, ty) :: !code
+ | S(sl, ofs, ty), R rd ->
+ code := LTL.Lgetstack(sl, ofs, ty, rd) :: !code
+ | S(sls, ofss, tys), S(sld, ofsd, tyd) ->
+ let tmp = temp_for tys in
+ (* code will be reversed at the end *)
+ code := LTL.Lsetstack(tmp, sld, ofsd, tyd) ::
+ LTL.Lgetstack(sls, ofss, tys, tmp) :: !code
+ in
+ let rec move_one i =
+ if src.(i) <> dst.(i) then begin
+ status.(i) <- Being_moved;
+ for j = 0 to n - 1 do
+ if src.(j) = dst.(i) then
+ match status.(j) with
+ | To_move ->
+ move_one j
+ | Being_moved ->
+ let tmp = R (temp_for (Loc.coq_type src.(j))) in
+ add_move src.(j) tmp;
+ src.(j) <- tmp
+ | Moved ->
+ ()
+ done;
+ add_move src.(i) dst.(i);
+ status.(i) <- Moved
+ end in
+ for i = 0 to n - 1 do
+ if status.(i) = To_move then move_one i
+ done;
+ List.rev_append !code k
+
+let transl_instr alloc instr k =
+ match instr with
+ | Xmove(src, dst) | Xreload(src, dst) | Xspill(src, dst) ->
+ make_move (alloc src) (alloc dst) k
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ make_parmove (List.map alloc srcs) (List.map alloc dsts)
+ (mreg_of alloc itmp) (mreg_of alloc ftmp) k
+ | Xop(op, args, res) ->
+ let rargs = mregs_of alloc args
+ and rres = mreg_of alloc res in
+ begin match is_two_address op rargs with
+ | None ->
+ LTL.Lop(op, rargs, rres) :: k
+ | Some(rarg1, rargl) ->
+ if rarg1 = rres then
+ LTL.Lop(op, rargs, rres) :: k
+ else
+ LTL.Lop(Omove, [rarg1], rres) ::
+ LTL.Lop(op, rres :: rargl, rres) :: k
+ end
+ | Xload(chunk, addr, args, dst) ->
+ LTL.Lload(chunk, addr, mregs_of alloc args, mreg_of alloc dst) :: k
+ | Xstore(chunk, addr, args, src) ->
+ LTL.Lstore(chunk, addr, mregs_of alloc args, mreg_of alloc src) :: k
+ | Xcall(sg, vos, args, res) ->
+ LTL.Lcall(sg, mros_of alloc vos) :: k
+ | Xtailcall(sg, vos, args) ->
+ LTL.Ltailcall(sg, mros_of alloc vos) :: []
+ | Xbuiltin(ef, args, res) ->
+ begin match ef with
+ | EF_annot _ ->
+ LTL.Lannot(ef, List.map alloc args) :: k
+ | _ ->
+ LTL.Lbuiltin(ef, mregs_of alloc args, mregs_of alloc res) :: k
+ end
+ | Xbranch s ->
+ LTL.Lbranch s :: []
+ | Xcond(cond, args, s1, s2) ->
+ LTL.Lcond(cond, mregs_of alloc args, s1, s2) :: []
+ | Xjumptable(arg, tbl) ->
+ LTL.Ljumptable(mreg_of alloc arg, tbl) :: []
+ | Xreturn optarg ->
+ LTL.Lreturn :: []
+
+let rec transl_block alloc blk =
+ match blk with
+ | [] -> []
+ | instr :: blk' -> transl_instr alloc instr (transl_block alloc blk')
+
+let transl_function fn alloc =
+ { LTL.fn_sig = fn.fn_sig;
+ LTL.fn_stacksize = fn.fn_stacksize;
+ LTL.fn_entrypoint = fn.fn_entrypoint;
+ LTL.fn_code = PTree.map1 (transl_block alloc) fn.fn_code
+ }
+
+
+(******************* All together *********************)
+
+exception Timeout
+
+let rec first_round f liveness =
+ let alloc = find_coloring f liveness in
+ if !option_dalloctrace then begin
+ fprintf !pp "-------------- After initial register allocation@ @.";
+ PrintXTL.print_function !pp ~alloc: alloc ~live: liveness f
+ end;
+ let ts = tospill_function f alloc in
+ if VSet.is_empty ts then success f alloc else more_rounds f ts 1
+
+and more_rounds f ts count =
+ if count >= 40 then raise Timeout;
+ let f' = spill_function f ts count in
+ let liveness = liveness_analysis f' in
+ let alloc = find_coloring f' liveness in
+ if !option_dalloctrace then begin
+ fprintf !pp "-------------- After register allocation (round %d)@ @." count;
+ PrintXTL.print_function !pp ~alloc: alloc ~live: liveness f'
+ end;
+ let ts' = tospill_function f' alloc in
+ if VSet.is_empty ts'
+ then success f' alloc
+ else begin
+ if !option_dalloctrace then begin
+ fprintf !pp "--- Remain to be spilled:@ @.";
+ VSet.iter (fun v -> fprintf !pp "%a " PrintXTL.var v) ts';
+ fprintf !pp "@ @."
+ end;
+ more_rounds f (VSet.union ts ts') (count + 1)
+ end
+
+and success f alloc =
+ let f' = transl_function f alloc in
+ if !option_dalloctrace then begin
+ fprintf !pp "-------------- Candidate allocation@ @.";
+ PrintLTL.print_function !pp P.one f'
+ end;
+ f'
+
+open Errors
+
+let regalloc f =
+ init_trace();
+ reset_temps();
+ let f1 = Splitting.rename_function f in
+ match RTLtyping.type_function f1 with
+ | Error msg ->
+ Error(MSG (coqstring_of_camlstring "RTL code after splitting is ill-typed:") :: msg)
+ | OK tyenv ->
+ let f2 = function_of_RTL_function f1 tyenv in
+ let liveness = liveness_analysis f2 in
+ let f3 = dead_code_elimination f2 liveness in
+ if !option_dalloctrace then begin
+ fprintf !pp "-------------- Initial XTL@ @.";
+ PrintXTL.print_function !pp f3
+ end;
+ try
+ OK(first_round f3 liveness)
+ with
+ | Timeout ->
+ Error(msg (coqstring_of_camlstring "Spilling fails to converge"))
+ | Type_error_at pc ->
+ Error [MSG(coqstring_of_camlstring "Ill-typed XTL code at PC ");
+ POS pc]
+ | Bad_LTL ->
+ Error(msg (coqstring_of_camlstring "Bad LTL after spilling"))
diff --git a/backend/Reload.v b/backend/Reload.v
deleted file mode 100644
index be844b3..0000000
--- a/backend/Reload.v
+++ /dev/null
@@ -1,274 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Reloading, spilling, and explication of calling conventions. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Op.
-Require Import Locations.
-Require Import LTLin.
-Require Import Conventions.
-Require Import Parallelmove.
-Require Import Linear.
-
-Open Local Scope error_monad_scope.
-
-(** * Spilling and reloading *)
-
-(** Operations in the Linear language, like those of the target processor,
- operate only over machine registers, but not over stack slots.
- Consider the LTLin instruction
-<<
- r1 <- Lop(Oadd, r1 :: r2 :: nil)
->>
- and assume that [r1] and [r2] are assigned to stack locations [S s1]
- and [S s2], respectively. The translated LTL code must load these
- stack locations into temporary integer registers (this is called
- ``reloading''), perform the [Oadd] operation over these temporaries,
- leave the result in a temporary, then store the temporary back to
- stack location [S s1] (this is called ``spilling''). In other term,
- the generated Linear code has the following shape:
-<<
- IT1 <- Lgetstack s1;
- IT2 <- Lgetstack s2;
- IT1 <- Lop(Oadd, IT1 :: IT2 :: nil);
- Lsetstack s1 IT1;
->>
- This section provides functions that assist in choosing appropriate
- temporaries and inserting the required spilling and reloading
- operations. *)
-
-(** ** Allocation of temporary registers for reloading and spilling. *)
-
-(** [reg_for l] returns a machine register appropriate for working
- over the location [l]: either the machine register [m] if [l = R m],
- or a temporary register of [l]'s type if [l] is a stack slot. *)
-
-Definition reg_for (l: loc) : mreg :=
- match l with
- | R r => r
- | S s => match slot_type s with Tint => IT1 | Tfloat => FT1 end
- end.
-
-(** [regs_for ll] is similar, for a list of locations [ll].
- We ensure that distinct temporaries are used for
- different elements of [ll].
- The result is correct only if [enough_temporaries ll = true]
- (see below). *)
-
-Fixpoint regs_for_rec (locs: list loc) (itmps ftmps: list mreg)
- {struct locs} : list mreg :=
- match locs with
- | nil => nil
- | R r :: ls => r :: regs_for_rec ls itmps ftmps
- | S s :: ls =>
- match slot_type s with
- | Tint =>
- match itmps with
- | nil => nil
- | it1 :: its => it1 :: regs_for_rec ls its ftmps
- end
- | Tfloat =>
- match ftmps with
- | nil => nil
- | ft1 :: fts => ft1 :: regs_for_rec ls itmps fts
- end
- end
- end.
-
-Definition regs_for (locs: list loc) :=
- regs_for_rec locs int_temporaries float_temporaries.
-
-(** Detect the situations where not enough temporaries are available
- for reloading. *)
-
-Fixpoint enough_temporaries_rec (locs: list loc) (itmps ftmps: list mreg)
- {struct locs} : bool :=
- match locs with
- | nil => true
- | R r :: ls => enough_temporaries_rec ls itmps ftmps
- | S s :: ls =>
- match slot_type s with
- | Tint =>
- match itmps with
- | nil => false
- | it1 :: its => enough_temporaries_rec ls its ftmps
- end
- | Tfloat =>
- match ftmps with
- | nil => false
- | ft1 :: fts => enough_temporaries_rec ls itmps fts
- end
- end
- end.
-
-Definition enough_temporaries (locs: list loc) :=
- enough_temporaries_rec locs int_temporaries float_temporaries.
-
-(** ** Insertion of Linear reloads, stores and moves *)
-
-(** [add_spill src dst k] prepends to [k] the instructions needed
- to assign location [dst] the value of machine register [mreg]. *)
-
-Definition add_spill (src: mreg) (dst: loc) (k: code) :=
- match dst with
- | R rd => if mreg_eq src rd then k else Lop Omove (src :: nil) rd :: k
- | S sl => Lsetstack src sl :: k
- end.
-
-(** [add_reload src dst k] prepends to [k] the instructions needed
- to assign machine register [mreg] the value of the location [src]. *)
-
-Definition add_reload (src: loc) (dst: mreg) (k: code) :=
- match src with
- | R rs => if mreg_eq rs dst then k else Lop Omove (rs :: nil) dst :: k
- | S sl => Lgetstack sl dst :: k
- end.
-
-(** [add_reloads] is similar for a list of locations (as sources)
- and a list of machine registers (as destinations). *)
-
-Fixpoint add_reloads (srcs: list loc) (dsts: list mreg) (k: code)
- {struct srcs} : code :=
- match srcs, dsts with
- | s1 :: sl, t1 :: tl => add_reload s1 t1 (add_reloads sl tl k)
- | _, _ => k
- end.
-
-(** [add_move src dst k] prepends to [k] the instructions that copy
- the value of location [src] into location [dst]. *)
-
-Definition add_move (src dst: loc) (k: code) :=
- if Loc.eq src dst then k else
- match src, dst with
- | R rs, _ =>
- add_spill rs dst k
- | _, R rd =>
- add_reload src rd k
- | S ss, S sd =>
- let tmp :=
- match slot_type ss with Tint => IT1 | Tfloat => FT1 end in
- add_reload src tmp (add_spill tmp dst k)
- end.
-
-(** [parallel_move srcs dsts k] is similar, but for a list of
- source locations and a list of destination locations of the same
- length. This is a parallel move, meaning that some of the
- destinations can also occur as sources. *)
-
-Definition parallel_move (srcs dsts: list loc) (k: code) : code :=
- List.fold_right
- (fun p k => add_move (fst p) (snd p) k)
- k (parmove srcs dsts).
-
-(** * Code transformation *)
-
-(** We insert appropriate reload, spill and parallel move operations
- around each of the instructions of the source code. *)
-
-Definition transf_instr
- (f: LTLin.function) (instr: LTLin.instruction) (k: code) : code :=
- match instr with
- | LTLin.Lop op args res =>
- match is_move_operation op args with
- | Some src =>
- add_move src res k
- | None =>
- let rargs := regs_for args in
- let rres := reg_for res in
- add_reloads args rargs (Lop op rargs rres :: add_spill rres res k)
- end
- | LTLin.Lload chunk addr args dst =>
- let rargs := regs_for args in
- let rdst := reg_for dst in
- add_reloads args rargs
- (Lload chunk addr rargs rdst :: add_spill rdst dst k)
- | LTLin.Lstore chunk addr args src =>
- if enough_temporaries (src :: args) then
- match regs_for (src :: args) with
- | nil => k (* never happens *)
- | rsrc :: rargs =>
- add_reloads (src :: args) (rsrc :: rargs)
- (Lstore chunk addr rargs rsrc :: k)
- end
- else
- let rargs := regs_for args in
- let rsrc := reg_for src in
- add_reloads args rargs
- (Lop (op_for_binary_addressing addr) rargs IT2 ::
- add_reload src rsrc
- (Lstore chunk (Aindexed Int.zero) (IT2 :: nil) rsrc :: k))
- | LTLin.Lcall sig los args res =>
- let largs := loc_arguments sig in
- let rres := loc_result sig in
- match los with
- | inl fn =>
- parallel_move args largs
- (add_reload fn (reg_for fn)
- (Lcall sig (inl _ (reg_for fn)) :: add_spill rres res k))
- | inr id =>
- parallel_move args largs
- (Lcall sig (inr _ id) :: add_spill rres res k)
- end
- | LTLin.Ltailcall sig los args =>
- let largs := loc_arguments sig in
- match los with
- | inl fn =>
- parallel_move args largs
- (add_reload fn IT1
- (Ltailcall sig (inl _ IT1) :: k))
- | inr id =>
- parallel_move args largs
- (Ltailcall sig (inr _ id) :: k)
- end
- | LTLin.Lbuiltin ef args dst =>
- if ef_reloads ef then
- (let rargs := regs_for args in
- let rdst := reg_for dst in
- add_reloads args rargs
- (Lbuiltin ef rargs rdst :: add_spill rdst dst k))
- else
- Lannot ef args :: k
- | LTLin.Llabel lbl =>
- Llabel lbl :: k
- | LTLin.Lgoto lbl =>
- Lgoto lbl :: k
- | LTLin.Lcond cond args lbl =>
- let rargs := regs_for args in
- add_reloads args rargs (Lcond cond rargs lbl :: k)
- | LTLin.Ljumptable arg tbl =>
- let rarg := reg_for arg in
- add_reload arg rarg (Ljumptable rarg tbl :: k)
- | LTLin.Lreturn None =>
- Lreturn :: k
- | LTLin.Lreturn (Some loc) =>
- add_reload loc (loc_result (LTLin.fn_sig f)) (Lreturn :: k)
- end.
-
-Definition transf_code (f: LTLin.function) (c: LTLin.code) : code :=
- list_fold_right (transf_instr f) c nil.
-
-Definition transf_function (f: LTLin.function) : function :=
- mkfunction
- (LTLin.fn_sig f)
- (LTLin.fn_stacksize f)
- (parallel_move (loc_parameters (LTLin.fn_sig f)) (LTLin.fn_params f)
- (transf_code f (LTLin.fn_code f))).
-
-Definition transf_fundef (fd: LTLin.fundef) : Linear.fundef :=
- transf_fundef transf_function fd.
-
-Definition transf_program (p: LTLin.program) : Linear.program :=
- transform_program transf_fundef p.
-
diff --git a/backend/Reloadproof.v b/backend/Reloadproof.v
deleted file mode 100644
index fe6e475..0000000
--- a/backend/Reloadproof.v
+++ /dev/null
@@ -1,1487 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Correctness proof for the [Reload] pass. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Values.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Op.
-Require Import Locations.
-Require Import Conventions.
-Require Import RTLtyping.
-Require Import LTLin.
-Require Import LTLintyping.
-Require Import Linear.
-Require Import Parallelmove.
-Require Import Reload.
-
-(** * Exploitation of the typing hypothesis *)
-
-Remark arity_ok_rec_incr_1:
- forall tys it itmps ftmps,
- arity_ok_rec tys itmps ftmps = true ->
- arity_ok_rec tys (it :: itmps) ftmps = true.
-Proof.
- induction tys; intros until ftmps; simpl.
- tauto.
- destruct a.
- destruct itmps. congruence. auto.
- destruct ftmps. congruence. auto.
-Qed.
-
-Remark arity_ok_rec_incr_2:
- forall tys ft itmps ftmps,
- arity_ok_rec tys itmps ftmps = true ->
- arity_ok_rec tys itmps (ft :: ftmps) = true.
-Proof.
- induction tys; intros until ftmps; simpl.
- tauto.
- destruct a.
- destruct itmps. congruence. auto.
- destruct ftmps. congruence. auto.
-Qed.
-
-Remark arity_ok_rec_decr:
- forall tys ty itmps ftmps,
- arity_ok_rec (ty :: tys) itmps ftmps = true ->
- arity_ok_rec tys itmps ftmps = true.
-Proof.
- intros until ftmps. simpl. destruct ty.
- destruct itmps. congruence. intros. apply arity_ok_rec_incr_1; auto.
- destruct ftmps. congruence. intros. apply arity_ok_rec_incr_2; auto.
-Qed.
-
-Lemma arity_ok_enough_rec:
- forall locs itmps ftmps,
- arity_ok_rec (List.map Loc.type locs) itmps ftmps = true ->
- enough_temporaries_rec locs itmps ftmps = true.
-Proof.
- induction locs; intros until ftmps.
- simpl. auto.
- simpl enough_temporaries_rec. simpl map.
- destruct a. intros. apply IHlocs. eapply arity_ok_rec_decr; eauto.
- simpl. destruct (slot_type s).
- destruct itmps; auto.
- destruct ftmps; auto.
-Qed.
-
-Lemma arity_ok_enough:
- forall locs,
- arity_ok (List.map Loc.type locs) = true ->
- enough_temporaries locs = true.
-Proof.
- unfold arity_ok, enough_temporaries. intros.
- apply arity_ok_enough_rec; auto.
-Qed.
-
-Lemma enough_temporaries_op_args:
- forall (op: operation) (args: list loc) (res: loc),
- (List.map Loc.type args, Loc.type res) = type_of_operation op ->
- enough_temporaries args = true.
-Proof.
- intros. apply arity_ok_enough.
- replace (map Loc.type args) with (fst (type_of_operation op)).
- destruct op; try (destruct c); try (destruct a); compute; reflexivity.
- rewrite <- H. auto.
-Qed.
-
-Lemma enough_temporaries_addr:
- forall (addr: addressing) (args: list loc),
- List.map Loc.type args = type_of_addressing addr ->
- enough_temporaries args = true.
-Proof.
- intros. apply arity_ok_enough. rewrite H.
- destruct addr; compute; reflexivity.
-Qed.
-
-Lemma enough_temporaries_cond:
- forall (cond: condition) (args: list loc),
- List.map Loc.type args = type_of_condition cond ->
- enough_temporaries args = true.
-Proof.
- intros. apply arity_ok_enough. rewrite H.
- destruct cond; compute; reflexivity.
-Qed.
-
-Lemma arity_ok_rec_length:
- forall tys itmps ftmps,
- (length tys <= length itmps)%nat ->
- (length tys <= length ftmps)%nat ->
- arity_ok_rec tys itmps ftmps = true.
-Proof.
- induction tys; intros until ftmps; simpl.
- auto.
- intros. destruct a.
- destruct itmps; simpl in H. omegaContradiction. apply IHtys; omega.
- destruct ftmps; simpl in H0. omegaContradiction. apply IHtys; omega.
-Qed.
-
-Lemma enough_temporaries_length:
- forall args,
- (length args <= 2)%nat -> enough_temporaries args = true.
-Proof.
- intros. apply arity_ok_enough. unfold arity_ok.
- apply arity_ok_rec_length.
- rewrite list_length_map. simpl. omega.
- rewrite list_length_map. simpl. omega.
-Qed.
-
-Lemma not_enough_temporaries_length:
- forall src args,
- enough_temporaries (src :: args) = false ->
- (length args >= 2)%nat.
-Proof.
- intros.
- assert (length (src :: args) <= 2 \/ length (src :: args) > 2)%nat by omega.
- destruct H0.
- exploit enough_temporaries_length. eauto. congruence.
- simpl in H0. omega.
-Qed.
-
-Lemma not_enough_temporaries_addr:
- forall (ge: genv) sp addr src args ls v m,
- enough_temporaries (src :: args) = false ->
- eval_addressing ge sp addr (List.map ls args) = Some v ->
- eval_operation ge sp (op_for_binary_addressing addr) (List.map ls args) m = Some v.
-Proof.
- intros.
- apply eval_op_for_binary_addressing; auto.
- rewrite list_length_map. eapply not_enough_temporaries_length; eauto.
-Qed.
-
-(** Some additional properties of [reg_for] and [regs_for]. *)
-
-Lemma regs_for_cons:
- forall src args,
- exists rsrc, exists rargs, regs_for (src :: args) = rsrc :: rargs.
-Proof.
- intros. unfold regs_for. simpl.
- destruct src. econstructor; econstructor; reflexivity.
- destruct (slot_type s); econstructor; econstructor; reflexivity.
-Qed.
-
-Lemma reg_for_not_IT2:
- forall l, loc_acceptable l -> reg_for l <> IT2.
-Proof.
- intros. destruct l; simpl.
- red; intros; subst m. simpl in H. intuition congruence.
- destruct (slot_type s); congruence.
-Qed.
-
-(** * Correctness of the Linear constructors *)
-
-(** This section proves theorems that establish the correctness of the
- Linear constructor functions such as [add_move]. The theorems are of
- the general form ``the generated Linear instructions execute and
- modify the location set in the expected way: the result location(s)
- contain the expected values; other, non-temporary locations keep
- their values''. *)
-
-Section LINEAR_CONSTRUCTORS.
-
-Variable ge: genv.
-Variable stk: list stackframe.
-Variable f: function.
-Variable sp: val.
-
-Lemma reg_for_spec:
- forall l,
- R(reg_for l) = l \/ In (R (reg_for l)) temporaries.
-Proof.
- intros. unfold reg_for. destruct l. tauto.
- case (slot_type s); simpl; tauto.
-Qed.
-
-Lemma reg_for_diff:
- forall l l',
- Loc.diff l l' -> Loc.notin l' temporaries ->
- Loc.diff (R (reg_for l)) l'.
-Proof.
- intros. destruct (reg_for_spec l).
- rewrite H1; auto.
- apply Loc.diff_sym. eapply Loc.in_notin_diff; eauto.
-Qed.
-
-Lemma add_reload_correct:
- forall src dst k rs m,
- exists rs',
- star step ge (State stk f sp (add_reload src dst k) rs m)
- E0 (State stk f sp k rs' m) /\
- rs' (R dst) = rs src /\
- forall l,
- Loc.diff (R dst) l ->
- loc_acceptable src \/ Loc.diff (R IT1) l ->
- Loc.notin l destroyed_at_move ->
- rs' l = rs l.
-Proof.
- intros. unfold add_reload. destruct src.
- destruct (mreg_eq m0 dst).
- subst dst. exists rs. split. apply star_refl. tauto.
- econstructor.
- split. apply star_one; apply exec_Lop. simpl; reflexivity.
- unfold undef_op. split. apply Locmap.gss.
- intros. rewrite Locmap.gso; auto; apply Locmap.guo; auto.
- econstructor.
- split. apply star_one; apply exec_Lgetstack.
- split. apply Locmap.gss.
- intros. rewrite Locmap.gso; auto.
- destruct s; unfold undef_getstack; unfold loc_acceptable in H0; auto.
- apply Locmap.gso. tauto.
-Qed.
-
-Lemma add_reload_correct_2:
- forall src k rs m,
- loc_acceptable src ->
- exists rs',
- star step ge (State stk f sp (add_reload src (reg_for src) k) rs m)
- E0 (State stk f sp k rs' m) /\
- rs' (R (reg_for src)) = rs src /\
- (forall l, Loc.notin l temporaries -> rs' l = rs l) /\
- rs' (R IT2) = rs (R IT2).
-Proof.
- intros. unfold reg_for, add_reload; destruct src.
- rewrite dec_eq_true. exists rs; split. constructor. auto.
- set (t := match slot_type s with
- | Tint => IT1
- | Tfloat => FT1
- end).
- exists (Locmap.set (R t) (rs (S s)) (undef_getstack s rs)).
- split. apply star_one; apply exec_Lgetstack.
- split. apply Locmap.gss.
- split. intros. rewrite Locmap.gso; auto.
- destruct s; unfold undef_getstack; unfold loc_acceptable in H; auto.
- apply Locmap.gso. tauto.
- apply Loc.diff_sym. simpl in H0; unfold t; destruct (slot_type s); tauto.
- rewrite Locmap.gso. unfold undef_getstack.
- destruct s; simpl in H; reflexivity || contradiction.
- unfold t; destruct (slot_type s); red; congruence.
-Qed.
-
-Lemma add_spill_correct:
- forall src dst k rs m,
- exists rs',
- star step ge (State stk f sp (add_spill src dst k) rs m)
- E0 (State stk f sp k rs' m) /\
- rs' dst = rs (R src) /\
- forall l, Loc.diff dst l -> Loc.notin l destroyed_at_move -> rs' l = rs l.
-Proof.
- intros. unfold add_spill. destruct dst.
- destruct (mreg_eq src m0).
- subst src. exists rs. split. apply star_refl. tauto.
- econstructor.
- split. apply star_one. apply exec_Lop. simpl; reflexivity.
- split. apply Locmap.gss.
- intros. rewrite Locmap.gso; auto; unfold undef_op; apply Locmap.guo; auto.
- econstructor.
- split. apply star_one. apply exec_Lsetstack.
- split. apply Locmap.gss.
- intros. rewrite Locmap.gso; auto; unfold undef_setstack; apply Locmap.guo; auto.
-Qed.
-
-Remark notin_destroyed_move_1:
- forall r, ~In r destroyed_at_move_regs -> Loc.notin (R r) destroyed_at_move.
-Proof.
- intros. simpl in *. intuition congruence.
-Qed.
-
-Remark notin_destroyed_move_2:
- forall s, Loc.notin (S s) destroyed_at_move.
-Proof.
- intros. simpl in *. destruct s; auto.
-Qed.
-
-Lemma add_reloads_correct_rec:
- forall srcs itmps ftmps k rs m,
- locs_acceptable srcs ->
- enough_temporaries_rec srcs itmps ftmps = true ->
- (forall r, In (R r) srcs -> In r itmps -> False) ->
- (forall r, In (R r) srcs -> In r ftmps -> False) ->
- (forall r, In (R r) srcs -> ~In r destroyed_at_move_regs) ->
- list_disjoint itmps ftmps ->
- list_norepet itmps ->
- list_norepet ftmps ->
- list_disjoint itmps destroyed_at_move_regs ->
- list_disjoint ftmps destroyed_at_move_regs ->
- exists rs',
- star step ge
- (State stk f sp (add_reloads srcs (regs_for_rec srcs itmps ftmps) k) rs m)
- E0 (State stk f sp k rs' m) /\
- reglist rs' (regs_for_rec srcs itmps ftmps) = map rs srcs /\
- (forall r, ~(In r itmps) -> ~(In r ftmps) -> ~(In r destroyed_at_move_regs) -> rs' (R r) = rs (R r)) /\
- (forall s, rs' (S s) = rs (S s)).
-Proof.
-Opaque destroyed_at_move_regs.
- induction srcs; simpl; intros.
- (* base case *)
- exists rs. split. apply star_refl. tauto.
- (* inductive case *)
- simpl in H0.
- assert (ACC1: loc_acceptable a) by (auto with coqlib).
- assert (ACC2: locs_acceptable srcs) by (red; auto with coqlib).
- destruct a as [r | s].
- (* a is a register *)
- simpl add_reload. rewrite dec_eq_true.
- exploit IHsrcs; eauto.
- intros [rs' [EX [RES [OTH1 OTH2]]]].
- exists rs'. split. eauto.
- split. simpl. decEq.
- apply OTH1. red; intros; eauto. red; intros; eauto. auto.
- auto.
- auto.
- (* a is a stack slot *)
- destruct (slot_type s).
- (* ... of integer type *)
- destruct itmps as [ | it1 itmps ]. discriminate. inv H5.
- destruct (add_reload_correct (S s) it1 (add_reloads srcs (regs_for_rec srcs itmps ftmps) k) rs m)
- as [rs1 [A [B C]]].
- exploit IHsrcs; eauto with coqlib.
- eapply list_disjoint_cons_left; eauto.
- eapply list_disjoint_cons_left; eauto.
- intros [rs' [P [Q [T U]]]].
- exists rs'. split. eapply star_trans; eauto.
- split. simpl. decEq. rewrite <- B. apply T.
- auto.
- eapply list_disjoint_notin. eexact H4. eauto with coqlib.
- eapply list_disjoint_notin. eapply H7. auto with coqlib.
- rewrite Q. apply list_map_exten. intros. symmetry. apply C.
- simpl. destruct x; auto. red; intro; subst m0. apply H1 with it1; auto with coqlib.
- auto.
- destruct x. apply notin_destroyed_move_1. auto. apply notin_destroyed_move_2.
- split. simpl. intros. transitivity (rs1 (R r)).
- apply T; tauto. apply C. simpl. tauto. auto.
- apply notin_destroyed_move_1; auto.
- intros. transitivity (rs1 (S s0)). auto. apply C. simpl. auto. auto. apply notin_destroyed_move_2.
- (* ... of float type *)
- destruct ftmps as [ | ft1 ftmps ]. discriminate. inv H6.
- destruct (add_reload_correct (S s) ft1 (add_reloads srcs (regs_for_rec srcs itmps ftmps) k) rs m)
- as [rs1 [A [B C]]].
- exploit IHsrcs; eauto with coqlib.
- eapply list_disjoint_cons_right; eauto.
- eapply list_disjoint_cons_left; eauto.
- intros [rs' [P [Q [T U]]]].
- exists rs'. split. eapply star_trans; eauto.
- split. simpl. decEq. rewrite <- B. apply T.
- eapply list_disjoint_notin; eauto. apply list_disjoint_sym. apply H4. auto with coqlib.
- auto.
- eapply list_disjoint_notin. eexact H8. auto with coqlib.
- rewrite Q. apply list_map_exten. intros. symmetry. apply C.
- simpl. destruct x; auto. red; intro; subst m0. apply H2 with ft1; auto with coqlib. auto.
- destruct x. apply notin_destroyed_move_1. auto. apply notin_destroyed_move_2.
- split. simpl. intros. transitivity (rs1 (R r)).
- apply T; tauto. apply C. simpl. tauto. auto.
- apply notin_destroyed_move_1; auto.
- intros. transitivity (rs1 (S s0)). auto. apply C. simpl. auto. auto. apply notin_destroyed_move_2; auto.
-Qed.
-
-Lemma add_reloads_correct:
- forall srcs k rs m,
- enough_temporaries srcs = true ->
- locs_acceptable srcs ->
- exists rs',
- star step ge (State stk f sp (add_reloads srcs (regs_for srcs) k) rs m)
- E0 (State stk f sp k rs' m) /\
- reglist rs' (regs_for srcs) = List.map rs srcs /\
- forall l, Loc.notin l temporaries -> rs' l = rs l.
-Proof.
-Transparent destroyed_at_move_regs.
- intros.
- unfold enough_temporaries in H.
- exploit add_reloads_correct_rec. eauto. eauto.
- intros. generalize (H0 _ H1). unfold loc_acceptable. generalize H2.
- simpl. intuition congruence.
- intros. generalize (H0 _ H1). unfold loc_acceptable. generalize H2.
- simpl. intuition congruence.
- intros. generalize (H0 _ H1). unfold loc_acceptable.
- simpl. intuition congruence.
- red; simpl; intros. intuition congruence.
- unfold int_temporaries. NoRepet.
- unfold float_temporaries. NoRepet.
- red; simpl; intros. intuition congruence.
- red; simpl; intros. intuition congruence.
- intros [rs' [EX [RES [OTH1 OTH2]]]].
- exists rs'. split. eexact EX.
- split. exact RES.
- intros. destruct l. generalize (Loc.notin_not_in _ _ H1); simpl; intro.
- apply OTH1; simpl; intuition congruence.
- apply OTH2.
-Qed.
-
-Lemma add_move_correct:
- forall src dst k rs m,
- exists rs',
- star step ge (State stk f sp (add_move src dst k) rs m)
- E0 (State stk f sp k rs' m) /\
- rs' dst = rs src /\
- forall l,
- Loc.diff l dst -> Loc.diff l (R IT1) -> Loc.diff l (R FT1) -> Loc.notin l destroyed_at_move ->
- rs' l = rs l.
-Proof.
- intros; unfold add_move.
- destruct (Loc.eq src dst).
- subst dst. exists rs. split. apply star_refl. tauto.
- destruct src.
- (* src is a register *)
- generalize (add_spill_correct m0 dst k rs m); intros [rs' [EX [RES OTH]]].
- exists rs'; intuition. apply OTH. apply Loc.diff_sym; auto. auto.
- destruct dst.
- (* src is a stack slot, dst a register *)
- generalize (add_reload_correct (S s) m0 k rs m); intros [rs' [EX [RES OTH]]].
- exists rs'; intuition. apply OTH. apply Loc.diff_sym; auto. right; apply Loc.diff_sym; auto. auto.
- (* src and dst are stack slots *)
- set (tmp := match slot_type s with Tint => IT1 | Tfloat => FT1 end).
- generalize (add_reload_correct (S s) tmp (add_spill tmp (S s0) k) rs m);
- intros [rs1 [EX1 [RES1 OTH1]]].
- generalize (add_spill_correct tmp (S s0) k rs1 m);
- intros [rs2 [EX2 [RES2 OTH2]]].
- exists rs2. split.
- eapply star_trans; eauto. traceEq.
- split. congruence.
- intros. rewrite OTH2. apply OTH1.
- apply Loc.diff_sym. unfold tmp; case (slot_type s); auto.
- right. apply Loc.diff_sym; auto. auto.
- apply Loc.diff_sym; auto. auto.
-Qed.
-
-Lemma effect_move_sequence:
- forall k moves rs m,
- let k' := List.fold_right (fun p k => add_move (fst p) (snd p) k) k moves in
- exists rs',
- star step ge (State stk f sp k' rs m)
- E0 (State stk f sp k rs' m) /\
- effect_seqmove moves rs rs'.
-Proof.
- induction moves; intros until m; simpl.
- exists rs; split. constructor. constructor.
- destruct a as [src dst]; simpl.
- set (k1 := fold_right
- (fun (p : loc * loc) (k : code) => add_move (fst p) (snd p) k)
- k moves) in *.
- destruct (add_move_correct src dst k1 rs m) as [rs1 [A [B C]]].
- destruct (IHmoves rs1 m) as [rs' [D E]].
- exists rs'; split.
- eapply star_trans; eauto.
- econstructor; eauto. red. tauto.
-Qed.
-
-Lemma parallel_move_correct:
- forall srcs dsts k rs m,
- List.length srcs = List.length dsts ->
- Loc.no_overlap srcs dsts ->
- Loc.norepet dsts ->
- Loc.disjoint srcs temporaries ->
- Loc.disjoint dsts temporaries ->
- exists rs',
- star step ge (State stk f sp (parallel_move srcs dsts k) rs m)
- E0 (State stk f sp k rs' m) /\
- List.map rs' dsts = List.map rs srcs /\
- forall l, Loc.notin l dsts -> Loc.notin l temporaries -> rs' l = rs l.
-Proof.
- intros.
- generalize (effect_move_sequence k (parmove srcs dsts) rs m).
- intros [rs' [EXEC EFFECT]].
- exists rs'. split. exact EXEC.
- apply effect_parmove; auto.
-Qed.
-
-Lemma parallel_move_arguments_correct:
- forall args sg k rs m,
- List.map Loc.type args = sg.(sig_args) ->
- locs_acceptable args ->
- exists rs',
- star step ge (State stk f sp (parallel_move args (loc_arguments sg) k) rs m)
- E0 (State stk f sp k rs' m) /\
- List.map rs' (loc_arguments sg) = List.map rs args /\
- forall l, Loc.notin l (loc_arguments sg) -> Loc.notin l temporaries -> rs' l = rs l.
-Proof.
- intros. apply parallel_move_correct.
- transitivity (length sg.(sig_args)).
- rewrite <- H. symmetry; apply list_length_map.
- symmetry. apply loc_arguments_length.
- apply no_overlap_arguments; auto.
- apply loc_arguments_norepet.
- apply locs_acceptable_disj_temporaries; auto.
- apply loc_arguments_not_temporaries.
-Qed.
-
-Lemma parallel_move_parameters_correct:
- forall params sg k rs m,
- List.map Loc.type params = sg.(sig_args) ->
- locs_acceptable params ->
- Loc.norepet params ->
- exists rs',
- star step ge (State stk f sp (parallel_move (loc_parameters sg) params k) rs m)
- E0 (State stk f sp k rs' m) /\
- List.map rs' params = List.map rs (loc_parameters sg) /\
- forall l, Loc.notin l params -> Loc.notin l temporaries -> rs' l = rs l.
-Proof.
- intros. apply parallel_move_correct.
- transitivity (length sg.(sig_args)).
- apply loc_parameters_length.
- rewrite <- H. apply list_length_map.
- apply no_overlap_parameters; auto.
- auto. apply loc_parameters_not_temporaries.
- apply locs_acceptable_disj_temporaries; auto.
-Qed.
-
-End LINEAR_CONSTRUCTORS.
-
-(** * Agreement between values of locations *)
-
-(** The predicate [agree] states that two location maps
- give compatible values to all acceptable locations,
- that is, non-temporary registers and [Local] stack slots.
- The notion of compatibility used is the [Val.lessdef] ordering,
- which enables a [Vundef] value in the original program to be refined
- into any value in the transformed program.
-
- A typical situation where this refinement of values occurs is at
- function entry point. In LTLin, all registers except those
- belonging to the function parameters are set to [Vundef]. In
- Linear, these registers have whatever value they had in the caller
- function. This difference is harmless: if the original LTLin code
- does not get stuck, we know that it does not use any of these
- [Vundef] values. *)
-
-Definition agree (rs1 rs2: locset) : Prop :=
- forall l, loc_acceptable l -> Val.lessdef (rs1 l) (rs2 l).
-
-Lemma agree_loc:
- forall rs1 rs2 l,
- agree rs1 rs2 -> loc_acceptable l -> Val.lessdef (rs1 l) (rs2 l).
-Proof.
- auto.
-Qed.
-
-Lemma agree_locs:
- forall rs1 rs2 ll,
- agree rs1 rs2 -> locs_acceptable ll ->
- Val.lessdef_list (map rs1 ll) (map rs2 ll).
-Proof.
- induction ll; simpl; intros.
- constructor.
- constructor. apply H. apply H0; auto with coqlib.
- apply IHll; auto. red; intros. apply H0; auto with coqlib.
-Qed.
-
-Lemma agree_exten:
- forall rs ls1 ls2,
- agree rs ls1 ->
- (forall l, Loc.notin l temporaries -> ls2 l = ls1 l) ->
- agree rs ls2.
-Proof.
- intros; red; intros. rewrite H0. auto.
- apply temporaries_not_acceptable; auto.
-Qed.
-
-Remark undef_temps_others:
- forall rs l,
- Loc.notin l temporaries -> LTL.undef_temps rs l = rs l.
-Proof.
- intros. apply Locmap.guo; auto.
-Qed.
-
-Remark undef_op_others:
- forall op rs l,
- Loc.notin l temporaries -> undef_op op rs l = rs l.
-Proof.
- intros. generalize (undef_temps_others rs l H); intro.
- unfold undef_op; destruct op; auto; apply Locmap.guo; simpl in *; tauto.
-Qed.
-
-Lemma agree_undef_temps:
- forall rs1 rs2,
- agree rs1 rs2 -> agree (LTL.undef_temps rs1) rs2.
-Proof.
- intros; red; intros. rewrite undef_temps_others; auto.
- apply Conventions.temporaries_not_acceptable. auto.
-Qed.
-
-Lemma agree_undef_temps2:
- forall rs1 rs2,
- agree rs1 rs2 -> agree (LTL.undef_temps rs1) (LTL.undef_temps rs2).
-Proof.
- intros. apply agree_exten with rs2. apply agree_undef_temps; auto.
- intros. apply undef_temps_others; auto.
-Qed.
-
-Lemma agree_set:
- forall rs1 rs2 rs2' l v,
- loc_acceptable l ->
- Val.lessdef v (rs2' l) ->
- (forall l', Loc.diff l l' -> Loc.notin l' temporaries -> rs2' l' = rs2 l') ->
- agree rs1 rs2 -> agree (Locmap.set l v rs1) rs2'.
-Proof.
- intros; red; intros.
- destruct (Loc.eq l l0).
- subst l0. rewrite Locmap.gss. auto.
- assert (Loc.diff l l0). eapply loc_acceptable_noteq_diff; eauto.
- rewrite Locmap.gso; auto. rewrite H1. auto. auto.
- apply temporaries_not_acceptable; auto.
-Qed.
-
-Lemma agree_set2:
- forall rs1 rs2 rs2' l v,
- loc_acceptable l ->
- Val.lessdef v (rs2' l) ->
- (forall l', Loc.diff l l' -> Loc.notin l' temporaries -> rs2' l' = rs2 l') ->
- agree rs1 rs2 -> agree (Locmap.set l v (LTL.undef_temps rs1)) rs2'.
-Proof.
- intros. eapply agree_set; eauto. apply agree_undef_temps; auto.
-Qed.
-
-Lemma agree_find_funct:
- forall (ge: Linear.genv) rs ls r f,
- Genv.find_funct ge (rs r) = Some f ->
- agree rs ls ->
- loc_acceptable r ->
- Genv.find_funct ge (ls r) = Some f.
-Proof.
- intros.
- assert (Val.lessdef (rs r) (ls r)). eapply agree_loc; eauto.
- exploit Genv.find_funct_inv; eauto. intros [b EQ]. rewrite EQ in H2.
- inv H2. rewrite <- EQ. auto.
-Qed.
-
-Lemma agree_postcall_1:
- forall rs ls,
- agree rs ls ->
- agree (LTL.postcall_locs rs) ls.
-Proof.
- intros; red; intros. unfold LTL.postcall_locs.
- destruct l; auto.
- destruct (In_dec Loc.eq (R m) temporaries). constructor.
- destruct (In_dec Loc.eq (R m) destroyed_at_call). constructor.
- auto.
-Qed.
-
-Lemma agree_postcall_2:
- forall rs ls ls',
- agree (LTL.postcall_locs rs) ls ->
- (forall l,
- loc_acceptable l -> ~In l destroyed_at_call -> ~In l temporaries ->
- ls' l = ls l) ->
- agree (LTL.postcall_locs rs) ls'.
-Proof.
- intros; red; intros. generalize (H l H1). unfold LTL.postcall_locs.
- destruct l.
- destruct (In_dec Loc.eq (R m) temporaries). intro; constructor.
- destruct (In_dec Loc.eq (R m) destroyed_at_call). intro; constructor.
- intro. rewrite H0; auto.
- intro. rewrite H0; auto.
- simpl. intuition congruence.
- simpl. intuition congruence.
-Qed.
-
-Lemma agree_postcall_call:
- forall rs ls ls' sig,
- agree rs ls ->
- (forall l,
- Loc.notin l (loc_arguments sig) -> Loc.notin l temporaries ->
- ls' l = ls l) ->
- agree (LTL.postcall_locs rs) ls'.
-Proof.
- intros. apply agree_postcall_2 with ls. apply agree_postcall_1; auto.
- intros. apply H0.
- apply arguments_not_preserved; auto.
- destruct l; simpl. simpl in H2. intuition congruence.
- destruct s; tauto.
- apply temporaries_not_acceptable; auto.
-Qed.
-
-Lemma agree_init_locs:
- forall ls dsts vl,
- locs_acceptable dsts ->
- Loc.norepet dsts ->
- Val.lessdef_list vl (map ls dsts) ->
- agree (LTL.init_locs vl dsts) ls.
-Proof.
- induction dsts; intros; simpl.
- red; intros. unfold Locmap.init. constructor.
- simpl in H1. inv H1. inv H0.
- apply agree_set with ls. apply H; auto with coqlib. auto. auto.
- apply IHdsts; auto. red; intros; apply H; auto with coqlib.
-Qed.
-
-Lemma call_regs_parameters:
- forall ls sig,
- map (call_regs ls) (loc_parameters sig) = map ls (loc_arguments sig).
-Proof.
- intros. unfold loc_parameters. rewrite list_map_compose.
- apply list_map_exten; intros.
- unfold call_regs, parameter_of_argument.
- generalize (loc_arguments_acceptable _ _ H).
- unfold loc_argument_acceptable.
- destruct x.
- intros. destruct (in_dec Loc.eq (R m) temporaries). contradiction. auto.
- destruct s; intros; try contradiction. auto.
-Qed.
-
-Lemma return_regs_preserve:
- forall ls1 ls2 l,
- ~ In l temporaries ->
- ~ In l destroyed_at_call ->
- return_regs ls1 ls2 l = ls1 l.
-Proof.
- intros. unfold return_regs. destruct l; auto.
- destruct (In_dec Loc.eq (R m) temporaries). contradiction.
- destruct (In_dec Loc.eq (R m) destroyed_at_call). contradiction.
- auto.
-Qed.
-
-Lemma return_regs_arguments:
- forall ls1 ls2 sig,
- tailcall_possible sig ->
- map (return_regs ls1 ls2) (loc_arguments sig) = map ls2 (loc_arguments sig).
-Proof.
- intros. apply list_map_exten; intros.
- unfold return_regs. generalize (H x H0). destruct x; intros.
- destruct (In_dec Loc.eq (R m) temporaries). auto.
- destruct (In_dec Loc.eq (R m) destroyed_at_call). auto.
- elim n0. eapply arguments_caller_save; eauto.
- contradiction.
-Qed.
-
-Lemma return_regs_result:
- forall ls1 ls2 sig,
- return_regs ls1 ls2 (R (loc_result sig)) = ls2 (R (loc_result sig)).
-Proof.
- intros. unfold return_regs.
- destruct (In_dec Loc.eq (R (loc_result sig)) temporaries). auto.
- destruct (In_dec Loc.eq (R (loc_result sig)) destroyed_at_call). auto.
- generalize (loc_result_caller_save sig). tauto.
-Qed.
-
-(** * Preservation of labels and gotos *)
-
-Lemma find_label_add_spill:
- forall lbl src dst k,
- find_label lbl (add_spill src dst k) = find_label lbl k.
-Proof.
- intros. destruct dst; simpl; auto.
- destruct (mreg_eq src m); auto.
-Qed.
-
-Lemma find_label_add_reload:
- forall lbl src dst k,
- find_label lbl (add_reload src dst k) = find_label lbl k.
-Proof.
- intros. destruct src; simpl; auto.
- destruct (mreg_eq m dst); auto.
-Qed.
-
-Lemma find_label_add_reloads:
- forall lbl srcs dsts k,
- find_label lbl (add_reloads srcs dsts k) = find_label lbl k.
-Proof.
- induction srcs; intros; simpl. auto.
- destruct dsts; auto. rewrite find_label_add_reload. auto.
-Qed.
-
-Lemma find_label_add_move:
- forall lbl src dst k,
- find_label lbl (add_move src dst k) = find_label lbl k.
-Proof.
- intros; unfold add_move.
- destruct (Loc.eq src dst); auto.
- destruct src. apply find_label_add_spill.
- destruct dst. apply find_label_add_reload.
- rewrite find_label_add_reload. apply find_label_add_spill.
-Qed.
-
-Lemma find_label_parallel_move:
- forall lbl srcs dsts k,
- find_label lbl (parallel_move srcs dsts k) = find_label lbl k.
-Proof.
- intros. unfold parallel_move. generalize (parmove srcs dsts).
- induction m; simpl. auto.
- rewrite find_label_add_move. auto.
-Qed.
-
-Hint Rewrite find_label_add_spill find_label_add_reload
- find_label_add_reloads find_label_add_move
- find_label_parallel_move: labels.
-
-Opaque reg_for.
-
-Ltac FL := simpl; autorewrite with labels; auto.
-
-Lemma find_label_transf_instr:
- forall lbl sg instr k,
- find_label lbl (transf_instr sg instr k) =
- if LTLin.is_label lbl instr then Some k else find_label lbl k.
-Proof.
- intros. destruct instr; FL.
- destruct (is_move_operation o l); FL; FL.
- FL.
- destruct (enough_temporaries (l0 :: l)).
- destruct (regs_for (l0 :: l)); FL.
- FL. FL.
- destruct s0; FL; FL; FL.
- destruct s0; FL; FL; FL.
- destruct (ef_reloads e). FL. FL. FL.
- destruct o; FL.
-Qed.
-
-Lemma transf_code_cons:
- forall f i c, transf_code f (i :: c) = transf_instr f i (transf_code f c).
-Proof.
- unfold transf_code; intros. rewrite list_fold_right_eq; auto.
-Qed.
-
-Lemma find_label_transf_code:
- forall sg lbl c,
- find_label lbl (transf_code sg c) =
- option_map (transf_code sg) (LTLin.find_label lbl c).
-Proof.
- induction c; simpl.
- auto.
- rewrite transf_code_cons.
- rewrite find_label_transf_instr.
- destruct (LTLin.is_label lbl a); auto.
-Qed.
-
-Lemma find_label_transf_function:
- forall lbl f c,
- LTLin.find_label lbl (LTLin.fn_code f) = Some c ->
- find_label lbl (Linear.fn_code (transf_function f)) =
- Some (transf_code f c).
-Proof.
- intros. destruct f; simpl in *. FL.
- rewrite find_label_transf_code. rewrite H; auto.
-Qed.
-
-(** * Semantic preservation *)
-
-Section PRESERVATION.
-
-Variable prog: LTLin.program.
-Let tprog := transf_program prog.
-Hypothesis WT_PROG: LTLintyping.wt_program prog.
-
-Let ge := Genv.globalenv prog.
-Let tge := Genv.globalenv tprog.
-
-Lemma functions_translated:
- forall v f,
- 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 function_ptr_translated:
- forall v f,
- Genv.find_funct_ptr ge v = Some f ->
- Genv.find_funct_ptr tge v = Some (transf_fundef f).
-Proof (@Genv.find_funct_ptr_transf _ _ _ transf_fundef prog).
-
-Lemma symbols_preserved:
- forall id,
- Genv.find_symbol tge id = Genv.find_symbol ge id.
-Proof (@Genv.find_symbol_transf _ _ _ transf_fundef prog).
-
-Lemma varinfo_preserved:
- forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
-Proof (@Genv.find_var_info_transf _ _ _ transf_fundef prog).
-
-Lemma sig_preserved:
- forall f, funsig (transf_fundef f) = LTLin.funsig f.
-Proof.
- destruct f; reflexivity.
-Qed.
-
-Lemma find_function_wt:
- forall ros rs f,
- LTLin.find_function ge ros rs = Some f -> wt_fundef f.
-Proof.
- intros until f. destruct ros; simpl.
- intro. eapply Genv.find_funct_prop with (p := prog); eauto.
- caseEq (Genv.find_symbol ge i); intros.
- eapply Genv.find_funct_ptr_prop with (p := prog); eauto.
- congruence.
-Qed.
-
-(** The [match_state] predicate relates states in the original LTLin
- program and the transformed Linear program. The main property
- it enforces are:
-- Agreement between the values of locations in the two programs,
- according to the [agree] or [agree_arguments] predicates.
-- Agreement between the memory states of the two programs,
- according to the [Mem.lessdef] predicate.
-- Lists of LTLin instructions appearing in the source state
- are always suffixes of the code for the corresponding functions.
-- Well-typedness of the source code, which ensures that
- only acceptable locations are accessed by this code.
-- Agreement over return types during calls: the return type of a function
- is always equal to the return type of the signature of the corresponding
- call. This invariant is necessary since the conventional location
- used for passing return values depend on the return type. This invariant
- is enforced through the third parameter of the [match_stackframes]
- predicate, which is the signature of the called function.
-*)
-
-Inductive match_stackframes:
- list LTLin.stackframe -> list Linear.stackframe -> signature -> Prop :=
- | match_stackframes_nil:
- forall sig,
- sig.(sig_res) = Some Tint ->
- match_stackframes nil nil sig
- | match_stackframes_cons:
- forall res f sp c rs s s' c' ls sig,
- match_stackframes s s' (LTLin.fn_sig f) ->
- c' = add_spill (loc_result sig) res (transf_code f c) ->
- agree (LTL.postcall_locs rs) ls ->
- loc_acceptable res ->
- wt_function f ->
- is_tail c (LTLin.fn_code f) ->
- match_stackframes
- (LTLin.Stackframe res f sp (LTL.postcall_locs rs) c :: s)
- (Linear.Stackframe (transf_function f) sp ls c' :: s')
- sig.
-
-Inductive match_states: LTLin.state -> Linear.state -> Prop :=
- | match_states_intro:
- forall s f sp c rs m s' ls tm
- (STACKS: match_stackframes s s' (LTLin.fn_sig f))
- (AG: agree rs ls)
- (WT: wt_function f)
- (TL: is_tail c (LTLin.fn_code f))
- (MMD: Mem.extends m tm),
- match_states (LTLin.State s f sp c rs m)
- (Linear.State s' (transf_function f) sp (transf_code f c) ls tm)
- | match_states_call:
- forall s f args m s' ls tm
- (STACKS: match_stackframes s s' (LTLin.funsig f))
- (AG: Val.lessdef_list args (map ls (loc_arguments (LTLin.funsig f))))
- (PRES: forall l, ~In l temporaries -> ~In l destroyed_at_call ->
- ls l = parent_locset s' l)
- (WT: wt_fundef f)
- (MMD: Mem.extends m tm),
- match_states (LTLin.Callstate s f args m)
- (Linear.Callstate s' (transf_fundef f) ls tm)
- | match_states_return:
- forall s res m s' ls tm sig
- (STACKS: match_stackframes s s' sig)
- (AG: Val.lessdef res (ls (R (loc_result sig))))
- (PRES: forall l, ~In l temporaries -> ~In l destroyed_at_call ->
- ls l = parent_locset s' l)
- (MMD: Mem.extends m tm),
- match_states (LTLin.Returnstate s res m)
- (Linear.Returnstate s' ls tm).
-
-Lemma match_stackframes_change_sig:
- forall s s' sig1 sig2,
- match_stackframes s s' sig1 ->
- sig2.(sig_res) = sig1.(sig_res) ->
- match_stackframes s s' sig2.
-Proof.
- intros. inv H. constructor. congruence.
- econstructor; eauto. unfold loc_result. rewrite H0. auto.
-Qed.
-
-Ltac ExploitWT :=
- match goal with
- | [ WT: wt_function _, TL: is_tail _ _ |- _ ] =>
- generalize (wt_instrs _ WT _ (is_tail_in TL)); intro WTI
- end.
-
-(** The proof of semantic preservation is a simulation argument
- based on diagrams of the following form:
-<<
- st1 --------------- st2
- | |
- t| *|t
- | |
- v v
- st1'--------------- st2'
->>
- It is possible for the transformed code to take no transition,
- remaining in the same state; for instance, if the source transition
- corresponds to a move instruction that was eliminated.
- To ensure that this cannot occur infinitely often in a row,
- we use the following [measure] function that just counts the
- remaining number of instructions in the source code sequence. *)
-
-Definition measure (st: LTLin.state) : nat :=
- match st with
- | LTLin.State s f sp c ls m => List.length c
- | LTLin.Callstate s f ls m => 0%nat
- | LTLin.Returnstate s ls m => 0%nat
- end.
-
-Theorem transf_step_correct:
- forall s1 t s2, LTLin.step ge s1 t s2 ->
- forall s1' (MS: match_states s1 s1'),
- (exists s2', plus Linear.step tge s1' t s2' /\ match_states s2 s2')
- \/ (measure s2 < measure s1 /\ t = E0 /\ match_states s2 s1')%nat.
-Proof.
- Opaque regs_for. Opaque add_reloads.
- induction 1; intros; try inv MS; try rewrite transf_code_cons; simpl.
-
- (* Lop *)
- ExploitWT. inv WTI.
- (* move *)
- simpl.
- destruct (add_move_correct tge s' (transf_function f) sp r1 res (transf_code f b) ls tm)
- as [ls2 [A [B C]]].
- inv A.
- right. split. omega. split. auto.
- rewrite H1. rewrite H1. econstructor; eauto with coqlib.
- apply agree_set2 with ls2; auto.
- rewrite B. simpl in H; inversion H. auto.
- left; econstructor; split. eapply plus_left; eauto.
- econstructor; eauto with coqlib.
- apply agree_set2 with ls; auto.
- rewrite B. simpl in H; inversion H. auto.
- intros. apply C. apply Loc.diff_sym; auto.
- simpl in H7; tauto. simpl in H7; tauto. simpl in *; tauto.
- (* other ops *)
- assert (is_move_operation op args = None).
- caseEq (is_move_operation op args); intros.
- destruct (is_move_operation_correct _ _ H0). congruence.
- auto.
- rewrite H0.
- exploit add_reloads_correct.
- eapply enough_temporaries_op_args; eauto. auto.
- intros [ls2 [A [B C]]]. instantiate (1 := ls) in B.
- assert (exists tv, eval_operation tge sp op (reglist ls2 (regs_for args)) tm = Some tv
- /\ Val.lessdef v tv).
- apply eval_operation_lessdef with (map rs args) m; auto.
- rewrite B. eapply agree_locs; eauto.
- rewrite <- H. apply eval_operation_preserved. exact symbols_preserved.
- destruct H1 as [tv [P Q]].
- exploit add_spill_correct.
- intros [ls3 [D [E F]]].
- left; econstructor; split.
- eapply star_plus_trans. eexact A.
- eapply plus_left. eapply exec_Lop with (v := tv); eauto.
- eexact D. eauto. traceEq.
- econstructor; eauto with coqlib.
- apply agree_set2 with ls; auto.
- rewrite E. rewrite Locmap.gss. auto.
- intros. rewrite F; auto. rewrite Locmap.gso. rewrite undef_op_others; auto.
- apply reg_for_diff; auto. simpl in *; tauto.
-
- (* Lload *)
- ExploitWT; inv WTI.
- exploit add_reloads_correct.
- eapply enough_temporaries_addr; eauto. auto.
- intros [ls2 [A [B C]]].
- assert (exists ta, eval_addressing tge sp addr (reglist ls2 (regs_for args)) = Some ta
- /\ Val.lessdef a ta).
- apply eval_addressing_lessdef with (map rs args).
- rewrite B. eapply agree_locs; eauto.
- rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
- destruct H1 as [ta [P Q]].
- exploit Mem.loadv_extends; eauto. intros [tv [R S]].
- exploit add_spill_correct.
- intros [ls3 [D [E F]]].
- left; econstructor; split.
- eapply star_plus_trans. eauto.
- eapply plus_left. eapply exec_Lload; eauto.
- eauto. auto. traceEq.
- econstructor; eauto with coqlib.
- apply agree_set2 with ls; auto.
- rewrite E. rewrite Locmap.gss. auto.
- intros. rewrite F; auto. rewrite Locmap.gso. rewrite undef_temps_others; auto.
- apply reg_for_diff; auto. simpl in *; tauto.
-
- (* Lstore *)
- ExploitWT; inv WTI.
- caseEq (enough_temporaries (src :: args)); intro ENOUGH.
- destruct (regs_for_cons src args) as [rsrc [rargs EQ]]. rewrite EQ.
- exploit add_reloads_correct.
- eauto. red; simpl; intros. destruct H1. congruence. auto.
- intros [ls2 [A [B C]]]. rewrite EQ in A. rewrite EQ in B.
- injection B. intros D E.
- simpl in B.
- assert (exists ta, eval_addressing tge sp addr (reglist ls2 rargs) = Some ta
- /\ Val.lessdef a ta).
- apply eval_addressing_lessdef with (map rs args).
- rewrite D. eapply agree_locs; eauto.
- rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
- destruct H1 as [ta [P Q]].
- assert (X: Val.lessdef (rs src) (ls2 (R rsrc))).
- rewrite E. eapply agree_loc; eauto.
- exploit Mem.storev_extends. eexact MMD. eauto. eexact Q. eexact X.
- intros [tm2 [Y Z]].
- left; econstructor; split.
- eapply plus_right. eauto.
- eapply exec_Lstore with (a := ta); eauto.
- traceEq.
- econstructor; eauto with coqlib.
- apply agree_undef_temps2. apply agree_exten with ls; auto.
- (* not enough temporaries *)
- destruct (add_reloads_correct tge s' (transf_function f) sp args
- (Lop (op_for_binary_addressing addr) (regs_for args) IT2
- :: add_reload src (reg_for src)
- (Lstore chunk (Aindexed Int.zero) (IT2 :: nil) (reg_for src)
- :: transf_code f b)) ls tm)
- as [ls2 [A [B C]]].
- eapply enough_temporaries_addr; eauto. auto.
- assert (exists ta, eval_addressing tge sp addr (reglist ls2 (regs_for args)) = Some ta
- /\ Val.lessdef a ta).
- apply eval_addressing_lessdef with (map rs args).
- rewrite B. eapply agree_locs; eauto.
- rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
- destruct H1 as [ta [P Q]].
- set (ls3 := Locmap.set (R IT2) ta (undef_op (op_for_binary_addressing addr) ls2)).
- destruct (add_reload_correct_2 tge s' (transf_function f) sp src
- (Lstore chunk (Aindexed Int.zero) (IT2 :: nil) (reg_for src)
- :: transf_code f b)
- ls3 tm H8)
- as [ls4 [D [E [F G]]]].
- assert (NT: Loc.notin src temporaries) by (apply temporaries_not_acceptable; auto).
- assert (X: Val.lessdef (rs src) (ls4 (R (reg_for src)))).
- rewrite E. unfold ls3. rewrite Locmap.gso.
- rewrite undef_op_others; auto. rewrite C; auto.
- apply Loc.diff_sym. simpl in NT; tauto.
- exploit Mem.storev_extends. eexact MMD. eauto. eexact Q. eexact X.
- intros [tm2 [Y Z]].
- left; econstructor; split.
- eapply star_plus_trans. eauto.
- eapply plus_left. eapply exec_Lop with (v := ta).
- rewrite B. eapply not_enough_temporaries_addr; eauto.
- rewrite <- B; auto.
- eapply star_right. eauto.
- eapply exec_Lstore with (a := ta); eauto.
- simpl reglist. rewrite G. unfold ls3. rewrite Locmap.gss. simpl.
- destruct ta; simpl in Y; try discriminate. simpl; rewrite Int.add_zero; auto.
- reflexivity. reflexivity. traceEq.
- econstructor; eauto with coqlib.
- apply agree_undef_temps2. apply agree_exten with ls; auto.
- intros. rewrite F; auto. unfold ls3. rewrite Locmap.gso.
- rewrite undef_op_others; auto.
- apply Loc.diff_sym. simpl in H1; tauto.
-
- (* Lcall *)
- ExploitWT. inversion WTI. subst ros0 args0 res0. rewrite <- H0.
- assert (WTF': wt_fundef f'). eapply find_function_wt; eauto.
- destruct ros as [fn | id].
- (* indirect call *)
- red in H6. destruct H6 as [OK1 [OK2 OK3]].
- rewrite <- H0 in H4. rewrite <- H0 in OK3.
- destruct (parallel_move_arguments_correct tge s' (transf_function f) sp
- args sig
- (add_reload fn (reg_for fn)
- (Lcall sig (inl ident (reg_for fn))
- :: add_spill (loc_result sig) res (transf_code f b)))
- ls tm H4 H7)
- as [ls2 [A [B C]]].
- destruct (add_reload_correct_2 tge s' (transf_function f) sp fn
- (Lcall sig (inl ident (reg_for fn))
- :: add_spill (loc_result sig) res (transf_code f b))
- ls2 tm OK2)
- as [ls3 [D [E [F G]]]].
- assert (ARGS: Val.lessdef_list (map rs args)
- (map ls3 (loc_arguments sig))).
- replace (map ls3 (loc_arguments sig)) with (map ls2 (loc_arguments sig)).
- rewrite B. apply agree_locs; auto.
- apply list_map_exten; intros. apply F.
- apply Loc.disjoint_notin with (loc_arguments sig).
- apply loc_arguments_not_temporaries. auto.
- left; econstructor; split.
- eapply star_plus_trans. eexact A. eapply plus_right. eexact D.
- eapply exec_Lcall; eauto.
- simpl. rewrite E. rewrite C. eapply agree_find_funct; eauto.
- apply functions_translated. eauto.
- apply loc_acceptable_notin_notin; auto.
- apply temporaries_not_acceptable; auto.
- rewrite H0; symmetry; apply sig_preserved.
- eauto. traceEq.
- econstructor; eauto.
- econstructor; eauto with coqlib.
- rewrite H0. auto.
- apply agree_postcall_call with ls sig; auto.
- intros. rewrite F; auto. congruence.
- (* direct call *)
- rewrite <- H0 in H4.
- destruct (parallel_move_arguments_correct tge s' (transf_function f) sp
- args sig
- (Lcall sig (inr mreg id)
- :: add_spill (loc_result sig) res (transf_code f b))
- ls tm H4 H7)
- as [ls3 [D [E F]]].
- assert (ARGS: Val.lessdef_list (map rs args) (map ls3 (loc_arguments sig))).
- rewrite E. apply agree_locs; auto.
- left; econstructor; split.
- eapply plus_right. eauto.
- eapply exec_Lcall; eauto.
- simpl. rewrite symbols_preserved.
- generalize H; simpl. destruct (Genv.find_symbol ge id).
- apply function_ptr_translated; auto. congruence.
- rewrite H0. symmetry; apply sig_preserved.
- traceEq.
- econstructor; eauto.
- econstructor; eauto with coqlib. rewrite H0; auto.
- apply agree_postcall_call with ls sig; auto.
- congruence.
-
- (* Ltailcall *)
- ExploitWT. inversion WTI. subst ros0 args0.
- assert (WTF': wt_fundef f'). eapply find_function_wt; eauto.
- rewrite <- H0.
- exploit Mem.free_parallel_extends; eauto. intros [tm' [FREE MMD']].
- destruct ros as [fn | id].
- (* indirect call *)
- red in H5. destruct H5 as [OK1 [OK2 OK3]].
- rewrite <- H0 in H4. rewrite <- H0 in OK3.
- destruct (parallel_move_arguments_correct tge s' (transf_function f) (Vptr stk Int.zero)
- args sig
- (add_reload fn IT1
- (Ltailcall sig (inl ident IT1) :: transf_code f b))
- ls tm H4 H6)
- as [ls2 [A [B C]]].
- destruct (add_reload_correct tge s' (transf_function f) (Vptr stk Int.zero) fn IT1
- (Ltailcall sig (inl ident IT1) :: transf_code f b)
- ls2 tm)
- as [ls3 [D [E F]]].
- assert (ARGS: Val.lessdef_list (map rs args)
- (map ls3 (loc_arguments sig))).
- replace (map ls3 (loc_arguments sig)) with (map ls2 (loc_arguments sig)).
- rewrite B. apply agree_locs; auto.
- apply list_map_exten; intros.
- exploit Loc.disjoint_notin. apply loc_arguments_not_temporaries. eauto. simpl; intros.
- apply F.
- apply Loc.diff_sym; tauto.
- auto.
- simpl; tauto.
- left; econstructor; split.
- eapply star_plus_trans. eexact A. eapply plus_right. eexact D.
- eapply exec_Ltailcall; eauto.
- simpl. rewrite E. rewrite C. eapply agree_find_funct; eauto.
- apply functions_translated. eauto.
- apply loc_acceptable_notin_notin; auto.
- apply temporaries_not_acceptable; auto.
- rewrite H0; symmetry; apply sig_preserved.
- eauto. traceEq.
- econstructor; eauto.
- eapply match_stackframes_change_sig; eauto.
- rewrite return_regs_arguments; auto. congruence.
- exact (return_regs_preserve (parent_locset s') ls3).
- (* direct call *)
- rewrite <- H0 in H4.
- destruct (parallel_move_arguments_correct tge s' (transf_function f) (Vptr stk Int.zero)
- args sig
- (Ltailcall sig (inr mreg id) :: transf_code f b)
- ls tm H4 H6)
- as [ls3 [D [E F]]].
- assert (ARGS: Val.lessdef_list (map rs args)
- (map ls3 (loc_arguments sig))).
- rewrite E. apply agree_locs. apply agree_exten with ls; auto. auto.
- left; econstructor; split.
- eapply plus_right. eauto.
- eapply exec_Ltailcall; eauto.
- simpl. rewrite symbols_preserved.
- generalize H; simpl. destruct (Genv.find_symbol ge id).
- apply function_ptr_translated; auto. congruence.
- rewrite H0. symmetry; apply sig_preserved.
- traceEq.
- econstructor; eauto.
- eapply match_stackframes_change_sig; eauto.
- rewrite return_regs_arguments; auto. congruence.
- exact (return_regs_preserve (parent_locset s') ls3).
-
- (* Lbuiltin *)
- ExploitWT; inv WTI.
- case_eq (ef_reloads ef); intro RELOADS.
- exploit add_reloads_correct.
- instantiate (1 := args). apply arity_ok_enough. rewrite H3. destruct H5. auto. congruence. auto.
- intros [ls2 [A [B C]]].
- exploit external_call_mem_extends; eauto.
- apply agree_locs; eauto.
- intros [v' [tm' [P [Q [R S]]]]].
- exploit add_spill_correct.
- intros [ls3 [D [E F]]].
- left; econstructor; split.
- eapply star_plus_trans. eauto.
- eapply plus_left. eapply exec_Lbuiltin. rewrite B.
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact varinfo_preserved.
- eauto. reflexivity. traceEq.
- econstructor; eauto with coqlib.
- apply agree_set with ls; auto.
- rewrite E. rewrite Locmap.gss. auto.
- intros. rewrite F; auto. rewrite Locmap.gso. rewrite undef_temps_others; auto.
- apply reg_for_diff; auto. simpl in *; tauto.
- (* no reload *)
- exploit external_call_mem_extends; eauto.
- apply agree_locs; eauto.
- intros [v' [tm' [P [Q [R S]]]]].
- assert (EQ: v = Vundef).
- destruct ef; simpl in RELOADS; try congruence. simpl in H; inv H. auto.
- subst v.
- left; econstructor; split.
- apply plus_one. eapply exec_Lannot.
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact varinfo_preserved.
- econstructor; eauto with coqlib.
- apply agree_set with ls; auto.
-
- (* Llabel *)
- left; econstructor; split.
- apply plus_one. apply exec_Llabel.
- econstructor; eauto with coqlib.
-
- (* Lgoto *)
- left; econstructor; split.
- apply plus_one. apply exec_Lgoto. apply find_label_transf_function; eauto.
- econstructor; eauto.
- eapply LTLin.find_label_is_tail; eauto.
-
- (* Lcond true *)
- ExploitWT; inv WTI.
- exploit add_reloads_correct.
- eapply enough_temporaries_cond; eauto. auto.
- intros [ls2 [A [B C]]].
- left; econstructor; split.
- eapply plus_right. eauto. eapply exec_Lcond_true; eauto.
- rewrite B. apply eval_condition_lessdef with (map rs args) m; auto.
- eapply agree_locs; eauto.
- apply find_label_transf_function; eauto.
- traceEq.
- econstructor; eauto.
- apply agree_undef_temps2. apply agree_exten with ls; auto.
- eapply LTLin.find_label_is_tail; eauto.
-
- (* Lcond false *)
- ExploitWT; inv WTI.
- exploit add_reloads_correct.
- eapply enough_temporaries_cond; eauto. auto.
- intros [ls2 [A [B C]]].
- left; econstructor; split.
- eapply plus_right. eauto. eapply exec_Lcond_false; eauto.
- rewrite B. apply eval_condition_lessdef with (map rs args) m; auto.
- eapply agree_locs; eauto.
- traceEq.
- econstructor; eauto with coqlib.
- apply agree_undef_temps2. apply agree_exten with ls; auto.
-
- (* Ljumptable *)
- ExploitWT; inv WTI.
- exploit add_reload_correct_2; eauto.
- intros [ls2 [A [B [C D]]]].
- left; econstructor; split.
- eapply plus_right. eauto. eapply exec_Ljumptable; eauto.
- assert (Val.lessdef (rs arg) (ls arg)). apply AG. auto.
- rewrite H in H2. inv H2. congruence.
- apply find_label_transf_function; eauto.
- traceEq.
- econstructor; eauto with coqlib.
- apply agree_undef_temps2. apply agree_exten with ls; auto.
- eapply LTLin.find_label_is_tail; eauto.
-
- (* Lreturn *)
- ExploitWT; inv WTI.
- exploit Mem.free_parallel_extends; eauto. intros [tm' [FREE MMD']].
- destruct or; simpl.
- (* with an argument *)
- exploit add_reload_correct.
- intros [ls2 [A [B C]]].
- left; econstructor; split.
- eapply plus_right. eauto. eapply exec_Lreturn; eauto.
- traceEq.
- econstructor; eauto.
- rewrite return_regs_result. rewrite B. apply agree_loc; auto.
- apply return_regs_preserve.
- (* without an argument *)
- left; econstructor; split.
- apply plus_one. eapply exec_Lreturn; eauto.
- econstructor; eauto.
- apply return_regs_preserve.
-
- (* internal function *)
- simpl in WT. inversion_clear WT. inversion H0. simpl in AG.
- exploit Mem.alloc_extends. eauto. eauto.
- instantiate (1 := 0); omega. instantiate (1 := LTLin.fn_stacksize f); omega.
- intros [tm' [ALLOC MMD']].
- destruct (parallel_move_parameters_correct tge s' (transf_function f)
- (Vptr stk Int.zero) (LTLin.fn_params f) (LTLin.fn_sig f)
- (transf_code f (LTLin.fn_code f)) (call_regs ls) tm'
- wt_params wt_acceptable wt_norepet)
- as [ls2 [A [B C]]].
- assert (AG2: agree (LTL.init_locs args (fn_params f)) ls2).
- apply agree_init_locs; auto.
- rewrite B. rewrite call_regs_parameters. auto.
- left; econstructor; split.
- eapply plus_left.
- eapply exec_function_internal; eauto.
- simpl. eauto. traceEq.
- econstructor; eauto with coqlib.
-
- (* external function *)
- exploit external_call_mem_extends; eauto.
- intros [res' [tm' [A [B [C D]]]]].
- left; econstructor; split.
- apply plus_one. eapply exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact varinfo_preserved.
- econstructor; eauto.
- simpl. rewrite Locmap.gss. auto.
- intros. rewrite Locmap.gso. auto.
- simpl. destruct l; auto. red; intro. elim H1. subst m0.
- generalize (loc_result_caller_save (ef_sig ef)). tauto.
-
- (* return *)
- inv STACKS.
- exploit add_spill_correct. intros [ls2 [A [B C]]].
- left; econstructor; split.
- eapply plus_left. eapply exec_return; eauto.
- eauto. traceEq.
- econstructor; eauto.
- apply agree_set with ls; auto.
- rewrite B. auto.
- intros. apply C; auto. simpl in *; tauto.
- unfold parent_locset in PRES.
- apply agree_postcall_2 with ls0; auto.
-Qed.
-
-Lemma transf_initial_states:
- forall st1, LTLin.initial_state prog st1 ->
- exists st2, Linear.initial_state tprog st2 /\ match_states st1 st2.
-Proof.
- intros. inversion H.
- econstructor; split.
- econstructor.
- apply Genv.init_mem_transf; eauto.
- rewrite symbols_preserved. eauto.
- apply function_ptr_translated; eauto.
- rewrite sig_preserved. auto.
- econstructor; eauto. constructor. rewrite H3; auto.
- rewrite H3. simpl. constructor.
- eapply Genv.find_funct_ptr_prop; eauto.
- apply Mem.extends_refl.
-Qed.
-
-Lemma transf_final_states:
- forall st1 st2 r,
- match_states st1 st2 -> LTLin.final_state st1 r -> Linear.final_state st2 r.
-Proof.
- intros. inv H0. inv H. inv STACKS. econstructor.
- unfold loc_result in AG. rewrite H in AG. inv AG. auto.
-Qed.
-
-Theorem transf_program_correct:
- forward_simulation (LTLin.semantics prog) (Linear.semantics tprog).
-Proof.
- eapply forward_simulation_star.
- eexact symbols_preserved.
- eexact transf_initial_states.
- eexact transf_final_states.
- eexact transf_step_correct.
-Qed.
-
-End PRESERVATION.
diff --git a/backend/Reloadtyping.v b/backend/Reloadtyping.v
deleted file mode 100644
index 70d79bc..0000000
--- a/backend/Reloadtyping.v
+++ /dev/null
@@ -1,353 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Proof of type preservation for Reload and of
- correctness of computation of stack bounds for Linear. *)
-
-Require Import Coqlib.
-Require Import AST.
-Require Import Op.
-Require Import Locations.
-Require Import LTLin.
-Require Import LTLintyping.
-Require Import Linear.
-Require Import Lineartyping.
-Require Import Conventions.
-Require Import Parallelmove.
-Require Import Reload.
-Require Import Reloadproof.
-
-(** * Typing Linear constructors *)
-
-(** We show that the Linear constructor functions defined in [Reload]
- generate well-typed instruction sequences,
- given sufficient typing and well-formedness hypotheses over the locations involved. *)
-
-Hint Constructors wt_instr: reloadty.
-
-Remark wt_code_cons:
- forall f i c, wt_instr f i -> wt_code f c -> wt_code f (i :: c).
-Proof.
- intros; red; simpl; intros. elim H1; intro.
- subst i; auto. auto.
-Qed.
-
-Hint Resolve wt_code_cons: reloadty.
-
-Definition loc_valid (f: function) (l: loc) :=
- match l with R _ => True | S s => slot_valid f s end.
-
-Lemma loc_acceptable_valid:
- forall f l, loc_acceptable l -> loc_valid f l.
-Proof.
- destruct l; simpl; intro. auto.
- destruct s; simpl. omega. tauto. tauto.
-Qed.
-
-Definition loc_writable (l: loc) :=
- match l with R _ => True | S s => slot_writable s end.
-
-Lemma loc_acceptable_writable:
- forall l, loc_acceptable l -> loc_writable l.
-Proof.
- destruct l; simpl; intro. auto.
- destruct s; simpl; tauto.
-Qed.
-
-Hint Resolve loc_acceptable_valid loc_acceptable_writable: reloadty.
-
-Definition locs_valid (f: function) (ll: list loc) :=
- forall l, In l ll -> loc_valid f l.
-Definition locs_writable (ll: list loc) :=
- forall l, In l ll -> loc_writable l.
-
-Lemma locs_acceptable_valid:
- forall f ll, locs_acceptable ll -> locs_valid f ll.
-Proof.
- unfold locs_acceptable, locs_valid. auto with reloadty.
-Qed.
-
-Lemma locs_acceptable_writable:
- forall ll, locs_acceptable ll -> locs_writable ll.
-Proof.
- unfold locs_acceptable, locs_writable. auto with reloadty.
-Qed.
-
-Hint Resolve locs_acceptable_valid locs_acceptable_writable: reloadty.
-
-Lemma wt_add_reload:
- forall f src dst k,
- loc_valid f src -> Loc.type src = mreg_type dst ->
- wt_code f k -> wt_code f (add_reload src dst k).
-Proof.
- intros; unfold add_reload.
- destruct src; eauto with reloadty.
- destruct (mreg_eq m dst); eauto with reloadty.
-Qed.
-
-Hint Resolve wt_add_reload: reloadty.
-
-Lemma wt_add_reloads:
- forall f srcs dsts k,
- locs_valid f srcs -> map Loc.type srcs = map mreg_type dsts ->
- wt_code f k -> wt_code f (add_reloads srcs dsts k).
-Proof.
- induction srcs; destruct dsts; simpl; intros; try congruence.
- auto. inv H0. apply wt_add_reload; auto with coqlib reloadty.
- apply IHsrcs; auto. red; intros; auto with coqlib.
-Qed.
-
-Hint Resolve wt_add_reloads: reloadty.
-
-Lemma wt_add_spill:
- forall f src dst k,
- loc_valid f dst -> loc_writable dst -> Loc.type dst = mreg_type src ->
- wt_code f k -> wt_code f (add_spill src dst k).
-Proof.
- intros; unfold add_spill.
- destruct dst; eauto with reloadty.
- destruct (mreg_eq src m); eauto with reloadty.
-Qed.
-
-Hint Resolve wt_add_spill: reloadty.
-
-Lemma wt_add_move:
- forall f src dst k,
- loc_valid f src -> loc_valid f dst -> loc_writable dst ->
- Loc.type dst = Loc.type src ->
- wt_code f k -> wt_code f (add_move src dst k).
-Proof.
- intros. unfold add_move.
- destruct (Loc.eq src dst); auto.
- destruct src; auto with reloadty.
- destruct dst; auto with reloadty.
- set (tmp := match slot_type s with
- | Tint => IT1
- | Tfloat => FT1
- end).
- assert (mreg_type tmp = Loc.type (S s)).
- simpl. destruct (slot_type s); reflexivity.
- apply wt_add_reload; auto with reloadty.
- apply wt_add_spill; auto. congruence.
-Qed.
-
-Hint Resolve wt_add_move: reloadty.
-
-Lemma wt_add_moves:
- forall f b moves,
- (forall s d, In (s, d) moves ->
- loc_valid f s /\ loc_valid f d /\ loc_writable d /\ Loc.type s = Loc.type d) ->
- wt_code f b ->
- wt_code f
- (List.fold_right (fun p k => add_move (fst p) (snd p) k) b moves).
-Proof.
- induction moves; simpl; intros.
- auto.
- destruct a as [s d]. simpl.
- destruct (H s d) as [A [B [C D]]]. auto.
- auto with reloadty.
-Qed.
-
-Theorem wt_parallel_move:
- forall f srcs dsts b,
- List.map Loc.type srcs = List.map Loc.type dsts ->
- locs_valid f srcs -> locs_valid f dsts -> locs_writable dsts ->
- wt_code f b -> wt_code f (parallel_move srcs dsts b).
-Proof.
- intros. unfold parallel_move. apply wt_add_moves; auto.
- intros.
- elim (parmove_prop_2 _ _ _ _ H4); intros A B.
- split. destruct A as [C|[C|C]]; auto; subst s; exact I.
- split. destruct B as [C|[C|C]]; auto; subst d; exact I.
- split. destruct B as [C|[C|C]]; auto; subst d; exact I.
- eapply parmove_prop_3; eauto.
-Qed.
-Hint Resolve wt_parallel_move: reloadty.
-
-Lemma wt_reg_for:
- forall l, mreg_type (reg_for l) = Loc.type l.
-Proof.
- intros. destruct l; simpl. auto.
- case (slot_type s); reflexivity.
-Qed.
-Hint Resolve wt_reg_for: reloadty.
-
-Lemma wt_regs_for_rec:
- forall locs itmps ftmps,
- enough_temporaries_rec locs itmps ftmps = true ->
- (forall r, In r itmps -> mreg_type r = Tint) ->
- (forall r, In r ftmps -> mreg_type r = Tfloat) ->
- List.map mreg_type (regs_for_rec locs itmps ftmps) = List.map Loc.type locs.
-Proof.
- induction locs; intros.
- simpl. auto.
- simpl in H. simpl.
- destruct a.
- simpl. decEq. eauto.
- caseEq (slot_type s); intro SLOTTYPE; rewrite SLOTTYPE in H.
- destruct itmps. discriminate.
- simpl. decEq.
- rewrite SLOTTYPE. auto with coqlib.
- apply IHlocs; auto with coqlib.
- destruct ftmps. discriminate. simpl. decEq.
- rewrite SLOTTYPE. auto with coqlib.
- apply IHlocs; auto with coqlib.
-Qed.
-
-Lemma wt_regs_for:
- forall locs,
- enough_temporaries locs = true ->
- List.map mreg_type (regs_for locs) = List.map Loc.type locs.
-Proof.
- intros. unfold regs_for. apply wt_regs_for_rec.
- auto.
- simpl; intros; intuition; subst r; reflexivity.
- simpl; intros; intuition; subst r; reflexivity.
-Qed.
-Hint Resolve wt_regs_for: reloadty.
-
-Hint Resolve enough_temporaries_op_args enough_temporaries_cond enough_temporaries_addr: reloadty.
-
-Hint Extern 4 (_ = _) => congruence : reloadty.
-
-Lemma wt_transf_instr:
- forall f instr k,
- LTLintyping.wt_instr (LTLin.fn_sig f) instr ->
- wt_code (transf_function f) k ->
- wt_code (transf_function f) (transf_instr f instr k).
-Proof.
- Opaque reg_for regs_for.
- intros. inv H; simpl; auto with reloadty.
- caseEq (is_move_operation op args); intros.
- destruct (is_move_operation_correct _ _ H). congruence.
- assert (map mreg_type (regs_for args) = map Loc.type args).
- eauto with reloadty.
- assert (mreg_type (reg_for res) = Loc.type res). eauto with reloadty.
- auto with reloadty.
-
- assert (map mreg_type (regs_for args) = map Loc.type args).
- eauto with reloadty.
- assert (mreg_type (reg_for dst) = Loc.type dst). eauto with reloadty.
- auto with reloadty.
-
- caseEq (enough_temporaries (src :: args)); intro ENOUGH.
- caseEq (regs_for (src :: args)); intros. auto.
- assert (map mreg_type (regs_for (src :: args)) = map Loc.type (src :: args)).
- apply wt_regs_for. auto.
- rewrite H in H5. injection H5; intros.
- auto with reloadty.
- apply wt_add_reloads; auto with reloadty.
- symmetry. apply wt_regs_for. eauto with reloadty.
- assert (op_for_binary_addressing addr <> Omove).
- unfold op_for_binary_addressing; destruct addr; congruence.
- assert (type_of_operation (op_for_binary_addressing addr) = (type_of_addressing addr, Tint)).
- apply type_op_for_binary_addressing.
- rewrite <- H1. rewrite list_length_map.
- eapply not_enough_temporaries_length; eauto.
- apply wt_code_cons.
- constructor; auto. rewrite H5. decEq. rewrite <- H1. eauto with reloadty.
- apply wt_add_reload; auto with reloadty.
- apply wt_code_cons; auto. constructor. reflexivity.
- rewrite <- H2. eauto with reloadty.
-
- assert (locs_valid (transf_function f) (loc_arguments sig)).
- red; intros. generalize (loc_arguments_acceptable sig l H).
- destruct l; simpl; auto. destruct s; simpl; intuition.
- assert (locs_writable (loc_arguments sig)).
- red; intros. generalize (loc_arguments_acceptable sig l H6).
- destruct l; simpl; auto. destruct s; simpl; intuition.
- assert (map Loc.type args = map Loc.type (loc_arguments sig)).
- rewrite loc_arguments_type; auto.
- assert (Loc.type res = mreg_type (loc_result sig)).
- rewrite H2. unfold loc_result. unfold proj_sig_res.
- destruct (sig_res sig); auto. destruct t; auto.
- destruct ros.
- destruct H3 as [A [B C]].
- apply wt_parallel_move; eauto with reloadty.
- apply wt_add_reload; eauto with reloadty.
- apply wt_code_cons; eauto with reloadty.
- constructor. rewrite <- A. auto with reloadty.
- auto with reloadty.
-
- assert (locs_valid (transf_function f) (loc_arguments sig)).
- red; intros. generalize (loc_arguments_acceptable sig l H).
- destruct l; simpl; auto. destruct s; simpl; intuition.
- assert (locs_writable (loc_arguments sig)).
- red; intros. generalize (loc_arguments_acceptable sig l H6).
- destruct l; simpl; auto. destruct s; simpl; intuition.
- assert (map Loc.type args = map Loc.type (loc_arguments sig)).
- rewrite loc_arguments_type; auto.
- destruct ros. destruct H2 as [A [B C]]. auto 10 with reloadty.
- auto 10 with reloadty.
- destruct (ef_reloads ef) eqn:?.
- assert (arity_ok (sig_args (ef_sig ef)) = true) by intuition congruence.
- assert (map mreg_type (regs_for args) = map Loc.type args).
- apply wt_regs_for. apply arity_ok_enough. congruence.
- assert (mreg_type (reg_for res) = Loc.type res). eauto with reloadty.
- auto 10 with reloadty.
- auto with reloadty.
-
- assert (map mreg_type (regs_for args) = map Loc.type args).
- eauto with reloadty.
- auto with reloadty.
-
- assert (mreg_type (reg_for arg) = Loc.type arg).
- eauto with reloadty.
- auto with reloadty.
-
- destruct optres; simpl in *; auto with reloadty.
- apply wt_add_reload; auto with reloadty.
- unfold loc_result. rewrite <- H1.
- destruct (Loc.type l); reflexivity.
-Qed.
-
-Lemma wt_transf_code:
- forall f c,
- LTLintyping.wt_code (LTLin.fn_sig f) c ->
- Lineartyping.wt_code (transf_function f) (transf_code f c).
-Proof.
- induction c; simpl; intros.
- red; simpl; tauto.
- rewrite transf_code_cons.
- apply wt_transf_instr; auto with coqlib.
- apply IHc. red; auto with coqlib.
-Qed.
-
-Lemma wt_transf_fundef:
- forall fd,
- LTLintyping.wt_fundef fd ->
- Lineartyping.wt_fundef (transf_fundef fd).
-Proof.
- intros. destruct fd; simpl.
- inv H. inv H1. constructor. unfold wt_function. simpl.
- apply wt_parallel_move; auto with reloadty.
- rewrite loc_parameters_type. auto.
- red; intros.
- destruct (list_in_map_inv _ _ _ H) as [r [A B]].
- generalize (loc_arguments_acceptable _ _ B).
- destruct r; intro.
- rewrite A. simpl. auto.
- red in H0. destruct s; try tauto.
- simpl in A. subst l. simpl. auto.
- apply wt_transf_code; auto.
- constructor.
-Qed.
-
-Lemma program_typing_preserved:
- forall p,
- LTLintyping.wt_program p ->
- Lineartyping.wt_program (transf_program p).
-Proof.
- intros; red; intros.
- destruct (transform_program_function _ _ _ _ H0) as [f0 [A B]].
- subst f; apply wt_transf_fundef. eauto.
-Qed.
diff --git a/backend/SelectLong.vp b/backend/SelectLong.vp
new file mode 100644
index 0000000..8eb32c3
--- /dev/null
+++ b/backend/SelectLong.vp
@@ -0,0 +1,368 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Instruction selection for 64-bit integer operations *)
+
+Require Import Coqlib.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Op.
+Require Import CminorSel.
+Require Import SelectOp.
+
+Open Local Scope cminorsel_scope.
+
+(** Some operations on 64-bit integers are transformed into calls to
+ runtime library functions. The following record type collects
+ the names of these functions. *)
+
+Record helper_functions : Type := mk_helper_functions {
+ i64_dtos: ident; (**r float -> signed long *)
+ i64_dtou: ident; (**r float -> unsigned long *)
+ i64_stod: ident; (**r signed long -> float *)
+ i64_utod: ident; (**r unsigned long -> float *)
+ i64_neg: ident; (**r opposite *)
+ i64_add: ident; (**r addition *)
+ i64_sub: ident; (**r subtraction *)
+ i64_mul: ident; (**r multiplication 32x32->64 *)
+ i64_sdiv: ident; (**r signed division *)
+ i64_udiv: ident; (**r unsigned division *)
+ i64_smod: ident; (**r signed remainder *)
+ i64_umod: ident; (**r unsigned remainder *)
+ i64_shl: ident; (**r shift left *)
+ i64_shr: ident; (**r shift right unsigned *)
+ i64_sar: ident; (**r shift right signed *)
+ i64_scmp: ident; (**r signed comparison *)
+ i64_ucmp: ident (**r unsigned comparison *)
+}.
+
+Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong).
+Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat).
+Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong).
+Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong).
+Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong).
+Definition sig_ll_i := mksignature (Tlong :: Tlong :: nil) (Some Tint).
+Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong).
+
+Section SELECT.
+
+Variable hf: helper_functions.
+
+Definition makelong (h l: expr): expr :=
+ Eop Omakelong (h ::: l ::: Enil).
+
+Nondetfunction splitlong (e: expr) (f: expr -> expr -> expr) :=
+ match e with
+ | Eop Omakelong (h ::: l ::: Enil) => f h l
+ | _ => Elet e (f (Eop Ohighlong (Eletvar O ::: Enil)) (Eop Olowlong (Eletvar O ::: Enil)))
+ end.
+
+Nondetfunction splitlong2 (e1 e2: expr) (f: expr -> expr -> expr -> expr -> expr) :=
+ match e1, e2 with
+ | Eop Omakelong (h1 ::: l1 ::: Enil), Eop Omakelong (h2 ::: l2 ::: Enil) =>
+ f h1 l1 h2 l2
+ | Eop Omakelong (h1 ::: l1 ::: Enil), t2 =>
+ Elet t2 (f (lift h1) (lift l1)
+ (Eop Ohighlong (Eletvar O ::: Enil)) (Eop Olowlong (Eletvar O ::: Enil)))
+ | t1, Eop Omakelong (h2 ::: l2 ::: Enil) =>
+ Elet t1 (f (Eop Ohighlong (Eletvar O ::: Enil)) (Eop Olowlong (Eletvar O ::: Enil))
+ (lift h2) (lift l2))
+ | _, _ =>
+ Elet e1 (Elet (lift e2)
+ (f (Eop Ohighlong (Eletvar 1 ::: Enil)) (Eop Olowlong (Eletvar 1 ::: Enil))
+ (Eop Ohighlong (Eletvar O ::: Enil)) (Eop Olowlong (Eletvar O ::: Enil))))
+ end.
+
+Nondetfunction lowlong (e: expr) :=
+ match e with
+ | Eop Omakelong (e1 ::: e2 ::: Enil) => e2
+ | _ => Eop Olowlong (e ::: Enil)
+ end.
+
+Nondetfunction highlong (e: expr) :=
+ match e with
+ | Eop Omakelong (e1 ::: e2 ::: Enil) => e1
+ | _ => Eop Ohighlong (e ::: Enil)
+ end.
+
+Definition longconst (n: int64) : expr :=
+ makelong (Eop (Ointconst (Int64.hiword n)) Enil)
+ (Eop (Ointconst (Int64.loword n)) Enil).
+
+Nondetfunction is_longconst (e: expr) :=
+ match e with
+ | Eop Omakelong (Eop (Ointconst h) Enil ::: Eop (Ointconst l) Enil ::: Enil) =>
+ Some(Int64.ofwords h l)
+ | _ =>
+ None
+ end.
+
+Definition is_longconst_zero (e: expr) :=
+ match is_longconst e with
+ | Some n => Int64.eq n Int64.zero
+ | None => false
+ end.
+
+Definition intoflong (e: expr) := lowlong e.
+
+Definition longofint (e: expr) :=
+ Elet e (makelong (shrimm (Eletvar O) (Int.repr 31)) (Eletvar O)).
+
+Definition longofintu (e: expr) :=
+ makelong (Eop (Ointconst Int.zero) Enil) e.
+
+Definition negl (e: expr) :=
+ match is_longconst e with
+ | Some n => longconst (Int64.neg n)
+ | None => Ebuiltin (EF_builtin hf.(i64_neg) sig_l_l) (e ::: Enil)
+ end.
+
+Definition notl (e: expr) :=
+ splitlong e (fun h l => makelong (notint h) (notint l)).
+
+Definition longoffloat (arg: expr) :=
+ Eexternal hf.(i64_dtos) sig_f_l (arg ::: Enil).
+Definition longuoffloat (arg: expr) :=
+ Eexternal hf.(i64_dtou) sig_f_l (arg ::: Enil).
+Definition floatoflong (arg: expr) :=
+ Eexternal hf.(i64_stod) sig_l_f (arg ::: Enil).
+Definition floatoflongu (arg: expr) :=
+ Eexternal hf.(i64_utod) sig_l_f (arg ::: Enil).
+
+Definition andl (e1 e2: expr) :=
+ splitlong2 e1 e2 (fun h1 l1 h2 l2 => makelong (and h1 h2) (and l1 l2)).
+
+Definition orl (e1 e2: expr) :=
+ splitlong2 e1 e2 (fun h1 l1 h2 l2 => makelong (or h1 h2) (or l1 l2)).
+
+Definition xorl (e1 e2: expr) :=
+ splitlong2 e1 e2 (fun h1 l1 h2 l2 => makelong (xor h1 h2) (xor l1 l2)).
+
+Definition shllimm (e1: expr) (n: int) :=
+ if Int.eq n Int.zero then e1 else
+ if Int.ltu n Int.iwordsize then
+ splitlong e1 (fun h l =>
+ makelong (or (shlimm h n) (shruimm l (Int.sub Int.iwordsize n)))
+ (shlimm l n))
+ else if Int.ltu n Int64.iwordsize' then
+ makelong (shlimm (lowlong e1) (Int.sub n Int.iwordsize))
+ (Eop (Ointconst Int.zero) Enil)
+ else
+ Eexternal hf.(i64_shl) sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
+
+Definition shrluimm (e1: expr) (n: int) :=
+ if Int.eq n Int.zero then e1 else
+ if Int.ltu n Int.iwordsize then
+ splitlong e1 (fun h l =>
+ makelong (shruimm h n)
+ (or (shruimm l n) (shlimm h (Int.sub Int.iwordsize n))))
+ else if Int.ltu n Int64.iwordsize' then
+ makelong (Eop (Ointconst Int.zero) Enil)
+ (shruimm (highlong e1) (Int.sub n Int.iwordsize))
+ else
+ Eexternal hf.(i64_shr) sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
+
+Definition shrlimm (e1: expr) (n: int) :=
+ if Int.eq n Int.zero then e1 else
+ if Int.ltu n Int.iwordsize then
+ splitlong e1 (fun h l =>
+ makelong (shrimm h n)
+ (or (shruimm l n) (shlimm h (Int.sub Int.iwordsize n))))
+ else if Int.ltu n Int64.iwordsize' then
+ Elet (highlong e1)
+ (makelong (shrimm (Eletvar 0) (Int.repr 31))
+ (shrimm (Eletvar 0) (Int.sub n Int.iwordsize)))
+ else
+ Eexternal hf.(i64_sar) sig_li_l (e1 ::: Eop (Ointconst n) Enil ::: Enil).
+
+Definition is_intconst (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => Some n
+ | _ => None
+ end.
+
+Definition shll (e1 e2: expr) :=
+ match is_intconst e2 with
+ | Some n => shllimm e1 n
+ | None => Eexternal hf.(i64_shl) sig_li_l (e1 ::: e2 ::: Enil)
+ end.
+
+Definition shrlu (e1 e2: expr) :=
+ match is_intconst e2 with
+ | Some n => shrluimm e1 n
+ | None => Eexternal hf.(i64_shr) sig_li_l (e1 ::: e2 ::: Enil)
+ end.
+
+Definition shrl (e1 e2: expr) :=
+ match is_intconst e2 with
+ | Some n => shrlimm e1 n
+ | None => Eexternal hf.(i64_sar) sig_li_l (e1 ::: e2 ::: Enil)
+ end.
+
+Definition addl (e1 e2: expr) :=
+ let default := Ebuiltin (EF_builtin hf.(i64_add) sig_ll_l) (e1 ::: e2 ::: Enil) in
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 => longconst (Int64.add n1 n2)
+ | Some n1, _ => if Int64.eq n1 Int64.zero then e2 else default
+ | _, Some n2 => if Int64.eq n2 Int64.zero then e1 else default
+ | _, _ => default
+ end.
+
+Definition subl (e1 e2: expr) :=
+ let default := Ebuiltin (EF_builtin hf.(i64_sub) sig_ll_l) (e1 ::: e2 ::: Enil) in
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 => longconst (Int64.sub n1 n2)
+ | Some n1, _ => if Int64.eq n1 Int64.zero then negl e2 else default
+ | _, Some n2 => if Int64.eq n2 Int64.zero then e1 else default
+ | _, _ => default
+ end.
+
+Definition mull_base (e1 e2: expr) :=
+ splitlong2 e1 e2 (fun h1 l1 h2 l2 =>
+ Elet (Ebuiltin (EF_builtin hf.(i64_mul) sig_ii_l) (l1 ::: l2 ::: Enil))
+ (makelong
+ (add (add (Eop Ohighlong (Eletvar O ::: Enil))
+ (mul (lift l1) (lift h2)))
+ (mul (lift h1) (lift l2)))
+ (Eop Olowlong (Eletvar O ::: Enil)))).
+
+Definition mullimm (e: expr) (n: int64) :=
+ if Int64.eq n Int64.zero then longconst Int64.zero else
+ if Int64.eq n Int64.one then e else
+ match Int64.is_power2 n with
+ | Some l => shllimm e (Int.repr (Int64.unsigned l))
+ | None => mull_base e (longconst n)
+ end.
+
+Definition mull (e1 e2: expr) :=
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 => longconst (Int64.mul n1 n2)
+ | Some n1, _ => mullimm e2 n1
+ | _, Some n2 => mullimm e1 n2
+ | _, _ => mull_base e1 e2
+ end.
+
+Definition binop_long (id: ident) (sem: int64 -> int64 -> int64) (e1 e2: expr) :=
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 => longconst (sem n1 n2)
+ | _, _ => Eexternal id sig_ll_l (e1 ::: e2 ::: Enil)
+ end.
+
+Definition divl := binop_long hf.(i64_sdiv) Int64.divs.
+Definition modl := binop_long hf.(i64_smod) Int64.mods.
+
+Definition divlu (e1 e2: expr) :=
+ let default := Eexternal hf.(i64_udiv) sig_ll_l (e1 ::: e2 ::: Enil) in
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 => longconst (Int64.divu n1 n2)
+ | _, Some n2 =>
+ match Int64.is_power2 n2 with
+ | Some l => shrluimm e1 (Int.repr (Int64.unsigned l))
+ | None => default
+ end
+ | _, _ => default
+ end.
+
+Definition modlu (e1 e2: expr) :=
+ let default := Eexternal hf.(i64_umod) sig_ll_l (e1 ::: e2 ::: Enil) in
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 => longconst (Int64.modu n1 n2)
+ | _, Some n2 =>
+ match Int64.is_power2 n2 with
+ | Some l => andl e1 (longconst (Int64.sub n2 Int64.one))
+ | None => default
+ end
+ | _, _ => default
+ end.
+
+Definition cmpl_eq_zero (e: expr) :=
+ splitlong e (fun h l => comp Ceq (or h l) (Eop (Ointconst Int.zero) Enil)).
+
+Definition cmpl_ne_zero (e: expr) :=
+ splitlong e (fun h l => comp Cne (or h l) (Eop (Ointconst Int.zero) Enil)).
+
+Definition cmplu (c: comparison) (e1 e2: expr) :=
+ match c with
+ | Ceq =>
+ if is_longconst_zero e2
+ then cmpl_eq_zero e1
+ else cmpl_eq_zero (xorl e1 e2)
+ | Cne =>
+ if is_longconst_zero e2
+ then cmpl_ne_zero e1
+ else cmpl_ne_zero (xorl e1 e2)
+ | _ =>
+ comp c (Eexternal hf.(i64_ucmp) sig_ll_i (e1 ::: e2 ::: Enil))
+ (Eop (Ointconst Int.zero) Enil)
+ end.
+
+Definition cmpl (c: comparison) (e1 e2: expr) :=
+ let default :=
+ comp c (Eexternal hf.(i64_scmp) sig_ll_i (e1 ::: e2 ::: Enil))
+ (Eop (Ointconst Int.zero) Enil) in
+ match c with
+ | Ceq =>
+ if is_longconst_zero e2
+ then cmpl_eq_zero e1
+ else cmpl_eq_zero (xorl e1 e2)
+ | Cne =>
+ if is_longconst_zero e2
+ then cmpl_ne_zero e1
+ else cmpl_ne_zero (xorl e1 e2)
+ | Clt =>
+ if is_longconst_zero e2
+ then comp Clt (highlong e1) (Eop (Ointconst Int.zero) Enil)
+ else default
+ | Cge =>
+ if is_longconst_zero e2
+ then comp Cge (highlong e1) (Eop (Ointconst Int.zero) Enil)
+ else default
+ | _ =>
+ default
+ end.
+
+End SELECT.
+
+(** Setting up the helper functions *)
+
+Require Import Errors.
+
+Local Open Scope string_scope.
+Local Open Scope error_monad_scope.
+
+Parameter get_helper: Cminor.genv -> String.string -> signature -> res ident.
+Parameter get_builtin: String.string -> signature -> res ident.
+
+Definition get_helpers (ge: Cminor.genv): res helper_functions :=
+ do i64_dtos <- get_helper ge "__i64_dtos" sig_f_l ;
+ do i64_dtou <- get_helper ge "__i64_dtou" sig_f_l ;
+ do i64_stod <- get_helper ge "__i64_stod" sig_l_f ;
+ do i64_utod <- get_helper ge "__i64_utod" sig_l_f ;
+ do i64_neg <- get_builtin "__builtin_negl" sig_l_l ;
+ do i64_add <- get_builtin "__builtin_addl" sig_ll_l ;
+ do i64_sub <- get_builtin "__builtin_subl" sig_ll_l ;
+ do i64_mul <- get_builtin "__builtin_mull" sig_ll_l ;
+ do i64_sdiv <- get_helper ge "__i64_sdiv" sig_ll_l ;
+ do i64_udiv <- get_helper ge "__i64_udiv" sig_ll_l ;
+ do i64_smod <- get_helper ge "__i64_smod" sig_ll_l ;
+ do i64_umod <- get_helper ge "__i64_umod" sig_ll_l ;
+ do i64_shl <- get_helper ge "__i64_shl" sig_li_l ;
+ do i64_shr <- get_helper ge "__i64_shr" sig_li_l ;
+ do i64_sar <- get_helper ge "__i64_sar" sig_li_l ;
+ do i64_scmp <- get_helper ge "__i64_scmp" sig_ll_i ;
+ do i64_ucmp <- get_helper ge "__i64_ucmp" sig_ll_i ;
+ OK (mk_helper_functions
+ i64_dtos i64_dtou i64_stod i64_utod
+ i64_neg i64_add i64_sub i64_mul i64_sdiv i64_udiv i64_smod i64_umod
+ i64_shl i64_shr i64_sar
+ i64_scmp i64_ucmp).
diff --git a/backend/SelectLongproof.v b/backend/SelectLongproof.v
new file mode 100644
index 0000000..aca05bf
--- /dev/null
+++ b/backend/SelectLongproof.v
@@ -0,0 +1,1063 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Correctness of instruction selection for 64-bit integer operations *)
+
+Require Import Coqlib.
+Require Import AST.
+Require Import Errors.
+Require Import Integers.
+Require Import Floats.
+Require Import Values.
+Require Import Memory.
+Require Import Globalenvs.
+Require Import Events.
+Require Import Cminor.
+Require Import Op.
+Require Import CminorSel.
+Require Import SelectOp.
+Require Import SelectOpproof.
+Require Import SelectLong.
+
+Open Local Scope cminorsel_scope.
+
+(** * Axiomatization of the helper functions *)
+
+Section HELPERS.
+
+Context {F V: Type} (ge: Genv.t (AST.fundef F) V).
+
+Definition helper_implements (id: ident) (sg: signature) (vargs: list val) (vres: val) : Prop :=
+ exists b, exists ef,
+ Genv.find_symbol ge id = Some b
+ /\ Genv.find_funct_ptr ge b = Some (External ef)
+ /\ ef_sig ef = sg
+ /\ forall m, external_call ef ge vargs m E0 vres m.
+
+Definition builtin_implements (id: ident) (sg: signature) (vargs: list val) (vres: val) : Prop :=
+ forall m, external_call (EF_builtin id sg) ge vargs m E0 vres m.
+
+Definition i64_helpers_correct (hf: helper_functions) : Prop :=
+ (forall x z, Val.longoffloat x = Some z -> helper_implements hf.(i64_dtos) sig_f_l (x::nil) z)
+ /\(forall x z, Val.longuoffloat x = Some z -> helper_implements hf.(i64_dtou) sig_f_l (x::nil) z)
+ /\(forall x z, Val.floatoflong x = Some z -> helper_implements hf.(i64_stod) sig_l_f (x::nil) z)
+ /\(forall x z, Val.floatoflongu x = Some z -> helper_implements hf.(i64_utod) sig_l_f (x::nil) z)
+ /\(forall x, builtin_implements hf.(i64_neg) sig_l_l (x::nil) (Val.negl x))
+ /\(forall x y, builtin_implements hf.(i64_add) sig_ll_l (x::y::nil) (Val.addl x y))
+ /\(forall x y, builtin_implements hf.(i64_sub) sig_ll_l (x::y::nil) (Val.subl x y))
+ /\(forall x y, builtin_implements hf.(i64_mul) sig_ii_l (x::y::nil) (Val.mull' x y))
+ /\(forall x y z, Val.divls x y = Some z -> helper_implements hf.(i64_sdiv) sig_ll_l (x::y::nil) z)
+ /\(forall x y z, Val.divlu x y = Some z -> helper_implements hf.(i64_udiv) sig_ll_l (x::y::nil) z)
+ /\(forall x y z, Val.modls x y = Some z -> helper_implements hf.(i64_smod) sig_ll_l (x::y::nil) z)
+ /\(forall x y z, Val.modlu x y = Some z -> helper_implements hf.(i64_umod) sig_ll_l (x::y::nil) z)
+ /\(forall x y, helper_implements hf.(i64_shl) sig_li_l (x::y::nil) (Val.shll x y))
+ /\(forall x y, helper_implements hf.(i64_shr) sig_li_l (x::y::nil) (Val.shrlu x y))
+ /\(forall x y, helper_implements hf.(i64_sar) sig_li_l (x::y::nil) (Val.shrl x y))
+ /\(forall c x y, exists z, helper_implements hf.(i64_scmp) sig_ll_i (x::y::nil) z
+ /\ Val.cmpl c x y = Val.cmp c z Vzero)
+ /\(forall c x y, exists z, helper_implements hf.(i64_ucmp) sig_ll_i (x::y::nil) z
+ /\ Val.cmplu c x y = Val.cmp c z Vzero).
+
+End HELPERS.
+
+(** * Correctness of the instruction selection functions for 64-bit operators *)
+
+Section CMCONSTR.
+
+Variable hf: helper_functions.
+Variable ge: genv.
+Hypothesis HELPERS: i64_helpers_correct ge hf.
+Variable sp: val.
+Variable e: env.
+Variable m: mem.
+
+Ltac UseHelper :=
+ red in HELPERS;
+ repeat (eauto; match goal with | [ H: _ /\ _ |- _ ] => destruct H end).
+
+Lemma eval_helper:
+ forall le id sg args vargs vres,
+ eval_exprlist ge sp e m le args vargs ->
+ helper_implements ge id sg vargs vres ->
+ eval_expr ge sp e m le (Eexternal id sg args) vres.
+Proof.
+ intros. destruct H0 as (b & ef & A & B & C & D). econstructor; eauto.
+Qed.
+
+Corollary eval_helper_1:
+ forall le id sg arg1 varg1 vres,
+ eval_expr ge sp e m le arg1 varg1 ->
+ helper_implements ge id sg (varg1::nil) vres ->
+ eval_expr ge sp e m le (Eexternal id sg (arg1 ::: Enil)) vres.
+Proof.
+ intros. eapply eval_helper; eauto. constructor; auto. constructor.
+Qed.
+
+Corollary eval_helper_2:
+ forall le id sg arg1 arg2 varg1 varg2 vres,
+ eval_expr ge sp e m le arg1 varg1 ->
+ eval_expr ge sp e m le arg2 varg2 ->
+ helper_implements ge id sg (varg1::varg2::nil) vres ->
+ eval_expr ge sp e m le (Eexternal id sg (arg1 ::: arg2 ::: Enil)) vres.
+Proof.
+ intros. eapply eval_helper; eauto. constructor; auto. constructor; auto. constructor.
+Qed.
+
+Remark eval_builtin_1:
+ forall le id sg arg1 varg1 vres,
+ eval_expr ge sp e m le arg1 varg1 ->
+ builtin_implements ge id sg (varg1::nil) vres ->
+ eval_expr ge sp e m le (Ebuiltin (EF_builtin id sg) (arg1 ::: Enil)) vres.
+Proof.
+ intros. econstructor. econstructor. eauto. constructor. apply H0.
+Qed.
+
+Remark eval_builtin_2:
+ forall le id sg arg1 arg2 varg1 varg2 vres,
+ eval_expr ge sp e m le arg1 varg1 ->
+ eval_expr ge sp e m le arg2 varg2 ->
+ builtin_implements ge id sg (varg1::varg2::nil) vres ->
+ eval_expr ge sp e m le (Ebuiltin (EF_builtin id sg) (arg1 ::: arg2 ::: Enil)) vres.
+Proof.
+ intros. econstructor. constructor; eauto. constructor; eauto. constructor. apply H1.
+Qed.
+
+Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop :=
+ forall le a x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v.
+
+Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop :=
+ forall le a x b y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v.
+
+Ltac EvalOp :=
+ eauto;
+ match goal with
+ | [ |- eval_exprlist _ _ _ _ _ Enil _ ] => constructor
+ | [ |- eval_exprlist _ _ _ _ _ (_:::_) _ ] => econstructor; EvalOp
+ | [ |- eval_expr _ _ _ _ _ (Eletvar _) _ ] => constructor; simpl; eauto
+ | [ |- eval_expr _ _ _ _ _ (Elet _ _) _ ] => econstructor; EvalOp
+ | [ |- eval_expr _ _ _ _ _ (lift _) _ ] => apply eval_lift; EvalOp
+ | [ |- eval_expr _ _ _ _ _ _ _ ] => eapply eval_Eop; [EvalOp | simpl; eauto]
+ | _ => idtac
+ end.
+
+Lemma eval_splitlong:
+ forall le a f v sem,
+ (forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (f a b) v /\
+ (forall p q, x = Vint p -> y = Vint q -> v = sem (Vlong (Int64.ofwords p q)))) ->
+ match v with Vlong _ => True | _ => sem v = Vundef end ->
+ eval_expr ge sp e m le a v ->
+ exists v', eval_expr ge sp e m le (splitlong a f) v' /\ Val.lessdef (sem v) v'.
+Proof.
+ intros until sem; intros EXEC UNDEF.
+ unfold splitlong. case (splitlong_match a); intros.
+- InvEval. subst v.
+ exploit EXEC. eexact H2. eexact H3. intros [v' [A B]].
+ exists v'; split. auto.
+ destruct v1; simpl in *; try (rewrite UNDEF; auto).
+ destruct v0; simpl in *; try (rewrite UNDEF; auto).
+ erewrite B; eauto.
+- exploit (EXEC (v :: le) (Eop Ohighlong (Eletvar 0 ::: Enil)) (Eop Olowlong (Eletvar 0 ::: Enil))).
+ EvalOp. EvalOp.
+ intros [v' [A B]].
+ exists v'; split. econstructor; eauto.
+ destruct v; try (rewrite UNDEF; auto). erewrite B; simpl; eauto. rewrite Int64.ofwords_recompose. auto.
+Qed.
+
+Lemma eval_splitlong2:
+ forall le a b f va vb sem,
+ (forall le a1 a2 b1 b2 x1 x2 y1 y2,
+ eval_expr ge sp e m le a1 x1 ->
+ eval_expr ge sp e m le a2 x2 ->
+ eval_expr ge sp e m le b1 y1 ->
+ eval_expr ge sp e m le b2 y2 ->
+ exists v,
+ eval_expr ge sp e m le (f a1 a2 b1 b2) v /\
+ (forall p1 p2 q1 q2,
+ x1 = Vint p1 -> x2 = Vint p2 -> y1 = Vint q1 -> y2 = Vint q2 ->
+ v = sem (Vlong (Int64.ofwords p1 p2)) (Vlong (Int64.ofwords q1 q2)))) ->
+ match va, vb with Vlong _, Vlong _ => True | _, _ => sem va vb = Vundef end ->
+ eval_expr ge sp e m le a va ->
+ eval_expr ge sp e m le b vb ->
+ exists v, eval_expr ge sp e m le (splitlong2 a b f) v /\ Val.lessdef (sem va vb) v.
+Proof.
+ intros until sem; intros EXEC UNDEF.
+ unfold splitlong2. case (splitlong2_match a); intros.
+- InvEval. subst va vb.
+ exploit (EXEC le h1 l1 h2 l2); eauto. intros [v [A B]].
+ exists v; split; auto.
+ destruct v1; simpl in *; try (rewrite UNDEF; auto).
+ destruct v0; try (rewrite UNDEF; auto).
+ destruct v2; simpl in *; try (rewrite UNDEF; auto).
+ destruct v3; try (rewrite UNDEF; auto).
+ erewrite B; eauto.
+- InvEval. subst va.
+ exploit (EXEC (vb :: le) (lift h1) (lift l1)
+ (Eop Ohighlong (Eletvar 0 ::: Enil)) (Eop Olowlong (Eletvar 0 ::: Enil))).
+ EvalOp. EvalOp. EvalOp. EvalOp.
+ intros [v [A B]].
+ exists v; split.
+ econstructor; eauto.
+ destruct v1; simpl in *; try (rewrite UNDEF; auto).
+ destruct v0; try (rewrite UNDEF; auto).
+ destruct vb; try (rewrite UNDEF; auto).
+ erewrite B; simpl; eauto. rewrite Int64.ofwords_recompose. auto.
+- InvEval. subst vb.
+ exploit (EXEC (va :: le)
+ (Eop Ohighlong (Eletvar 0 ::: Enil)) (Eop Olowlong (Eletvar 0 ::: Enil))
+ (lift h2) (lift l2)).
+ EvalOp. EvalOp. EvalOp. EvalOp.
+ intros [v [A B]].
+ exists v; split.
+ econstructor; eauto.
+ destruct va; try (rewrite UNDEF; auto).
+ destruct v1; simpl in *; try (rewrite UNDEF; auto).
+ destruct v0; try (rewrite UNDEF; auto).
+ erewrite B; simpl; eauto. rewrite Int64.ofwords_recompose. auto.
+- exploit (EXEC (vb :: va :: le)
+ (Eop Ohighlong (Eletvar 1 ::: Enil)) (Eop Olowlong (Eletvar 1 ::: Enil))
+ (Eop Ohighlong (Eletvar 0 ::: Enil)) (Eop Olowlong (Eletvar 0 ::: Enil))).
+ EvalOp. EvalOp. EvalOp. EvalOp.
+ intros [v [A B]].
+ exists v; split. EvalOp.
+ destruct va; try (rewrite UNDEF; auto); destruct vb; try (rewrite UNDEF; auto).
+ erewrite B; simpl; eauto. rewrite ! Int64.ofwords_recompose; auto.
+Qed.
+
+Lemma is_longconst_sound:
+ forall le a x n,
+ is_longconst a = Some n ->
+ eval_expr ge sp e m le a x ->
+ x = Vlong n.
+Proof.
+ unfold is_longconst; intros until n; intros LC.
+ destruct (is_longconst_match a); intros.
+ inv LC. InvEval. simpl in H5. inv H5. auto.
+ discriminate.
+Qed.
+
+Lemma is_longconst_zero_sound:
+ forall le a x,
+ is_longconst_zero a = true ->
+ eval_expr ge sp e m le a x ->
+ x = Vlong Int64.zero.
+Proof.
+ unfold is_longconst_zero; intros.
+ destruct (is_longconst a) as [n|] eqn:E; try discriminate.
+ revert H. predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ intros. subst. eapply is_longconst_sound; eauto.
+ congruence.
+Qed.
+
+Lemma eval_lowlong: unary_constructor_sound lowlong Val.loword.
+Proof.
+ unfold lowlong; red. intros until x. destruct (lowlong_match a); intros.
+ InvEval. subst x. exists v0; split; auto.
+ destruct v1; simpl; auto. destruct v0; simpl; auto.
+ rewrite Int64.lo_ofwords. auto.
+ exists (Val.loword x); split; auto. EvalOp.
+Qed.
+
+Lemma eval_highlong: unary_constructor_sound highlong Val.hiword.
+Proof.
+ unfold highlong; red. intros until x. destruct (highlong_match a); intros.
+ InvEval. subst x. exists v1; split; auto.
+ destruct v1; simpl; auto. destruct v0; simpl; auto.
+ rewrite Int64.hi_ofwords. auto.
+ exists (Val.hiword x); split; auto. EvalOp.
+Qed.
+
+Lemma eval_longconst:
+ forall le n, eval_expr ge sp e m le (longconst n) (Vlong n).
+Proof.
+ intros. EvalOp. rewrite Int64.ofwords_recompose; auto.
+Qed.
+
+Theorem eval_intoflong: unary_constructor_sound intoflong Val.loword.
+Proof eval_lowlong.
+
+Theorem eval_longofintu: unary_constructor_sound longofintu Val.longofintu.
+Proof.
+ red; intros. unfold longofintu. econstructor; split. EvalOp.
+ unfold Val.longofintu. destruct x; auto.
+ replace (Int64.repr (Int.unsigned i)) with (Int64.ofwords Int.zero i); auto.
+ apply Int64.same_bits_eq; intros.
+ rewrite Int64.testbit_repr by auto.
+ rewrite Int64.bits_ofwords by auto.
+ fold (Int.testbit i i0).
+ destruct (zlt i0 Int.zwordsize).
+ auto.
+ rewrite Int.bits_zero. rewrite Int.bits_above by omega. auto.
+Qed.
+
+Theorem eval_longofint: unary_constructor_sound longofint Val.longofint.
+Proof.
+ red; intros. unfold longofint.
+ exploit (eval_shrimm ge sp e m (Int.repr 31) (x :: le) (Eletvar 0)). EvalOp.
+ intros [v1 [A B]].
+ econstructor; split. EvalOp.
+ destruct x; simpl; auto.
+ simpl in B. inv B. simpl.
+ replace (Int64.repr (Int.signed i))
+ with (Int64.ofwords (Int.shr i (Int.repr 31)) i); auto.
+ apply Int64.same_bits_eq; intros.
+ rewrite Int64.testbit_repr by auto.
+ rewrite Int64.bits_ofwords by auto.
+ rewrite Int.bits_signed by omega.
+ destruct (zlt i0 Int.zwordsize).
+ auto.
+ assert (Int64.zwordsize = 2 * Int.zwordsize) by reflexivity.
+ rewrite Int.bits_shr by omega.
+ change (Int.unsigned (Int.repr 31)) with (Int.zwordsize - 1).
+ f_equal. destruct (zlt (i0 - Int.zwordsize + (Int.zwordsize - 1)) Int.zwordsize); omega.
+Qed.
+
+Theorem eval_negl: unary_constructor_sound (negl hf) Val.negl.
+Proof.
+ unfold negl; red; intros. destruct (is_longconst a) eqn:E.
+ econstructor; split. apply eval_longconst.
+ exploit is_longconst_sound; eauto. intros EQ; subst x. simpl. auto.
+ econstructor; split. eapply eval_builtin_1; eauto. UseHelper. auto.
+Qed.
+
+Theorem eval_notl: unary_constructor_sound notl Val.notl.
+Proof.
+ red; intros. unfold notl. apply eval_splitlong; auto.
+ intros.
+ exploit eval_notint. eexact H0. intros [va [A B]].
+ exploit eval_notint. eexact H1. intros [vb [C D]].
+ exists (Val.longofwords va vb); split. EvalOp.
+ intros; subst. simpl in *. inv B; inv D.
+ simpl. unfold Int.not. rewrite <- Int64.decompose_xor. auto.
+ destruct x; auto.
+Qed.
+
+Theorem eval_longoffloat:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.longoffloat x = Some y ->
+ exists v, eval_expr ge sp e m le (longoffloat hf a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold longoffloat. econstructor; split.
+ eapply eval_helper_1; eauto. UseHelper.
+ auto.
+Qed.
+
+Theorem eval_longuoffloat:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.longuoffloat x = Some y ->
+ exists v, eval_expr ge sp e m le (longuoffloat hf a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold longuoffloat. econstructor; split.
+ eapply eval_helper_1; eauto. UseHelper.
+ auto.
+Qed.
+
+Theorem eval_floatoflong:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.floatoflong x = Some y ->
+ exists v, eval_expr ge sp e m le (floatoflong hf a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold floatoflong. econstructor; split.
+ eapply eval_helper_1; eauto. UseHelper.
+ auto.
+Qed.
+
+Theorem eval_floatoflongu:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.floatoflongu x = Some y ->
+ exists v, eval_expr ge sp e m le (floatoflongu hf a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold floatoflongu. econstructor; split.
+ eapply eval_helper_1; eauto. UseHelper.
+ auto.
+Qed.
+
+Theorem eval_andl: binary_constructor_sound andl Val.andl.
+Proof.
+ red; intros. unfold andl. apply eval_splitlong2; auto.
+ intros.
+ exploit eval_and. eexact H1. eexact H3. intros [va [A B]].
+ exploit eval_and. eexact H2. eexact H4. intros [vb [C D]].
+ exists (Val.longofwords va vb); split. EvalOp.
+ intros; subst. simpl in B; inv B. simpl in D; inv D.
+ simpl. f_equal. rewrite Int64.decompose_and. auto.
+ destruct x; auto. destruct y; auto.
+Qed.
+
+Theorem eval_orl: binary_constructor_sound orl Val.orl.
+Proof.
+ red; intros. unfold orl. apply eval_splitlong2; auto.
+ intros.
+ exploit eval_or. eexact H1. eexact H3. intros [va [A B]].
+ exploit eval_or. eexact H2. eexact H4. intros [vb [C D]].
+ exists (Val.longofwords va vb); split. EvalOp.
+ intros; subst. simpl in B; inv B. simpl in D; inv D.
+ simpl. f_equal. rewrite Int64.decompose_or. auto.
+ destruct x; auto. destruct y; auto.
+Qed.
+
+Theorem eval_xorl: binary_constructor_sound xorl Val.xorl.
+Proof.
+ red; intros. unfold xorl. apply eval_splitlong2; auto.
+ intros.
+ exploit eval_xor. eexact H1. eexact H3. intros [va [A B]].
+ exploit eval_xor. eexact H2. eexact H4. intros [vb [C D]].
+ exists (Val.longofwords va vb); split. EvalOp.
+ intros; subst. simpl in B; inv B. simpl in D; inv D.
+ simpl. f_equal. rewrite Int64.decompose_xor. auto.
+ destruct x; auto. destruct y; auto.
+Qed.
+
+Lemma is_intconst_sound:
+ forall le a x n,
+ is_intconst a = Some n ->
+ eval_expr ge sp e m le a x ->
+ x = Vint n.
+Proof.
+ unfold is_intconst; intros until n; intros LC.
+ destruct a; try discriminate. destruct o; try discriminate. destruct e0; try discriminate.
+ inv LC. intros. InvEval. auto.
+Qed.
+
+Remark eval_shift_imm:
+ forall (P: expr -> Prop) n a0 a1 a2 a3,
+ (n = Int.zero -> P a0) ->
+ (0 <= Int.unsigned n < Int.zwordsize ->
+ Int.ltu n Int.iwordsize = true ->
+ Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize = true ->
+ Int.ltu n Int64.iwordsize' = true ->
+ P a1) ->
+ (Int.zwordsize <= Int.unsigned n < Int64.zwordsize ->
+ Int.ltu (Int.sub n Int.iwordsize) Int.iwordsize = true ->
+ P a2) ->
+ P a3 ->
+ P (if Int.eq n Int.zero then a0
+ else if Int.ltu n Int.iwordsize then a1
+ else if Int.ltu n Int64.iwordsize' then a2
+ else a3).
+Proof.
+ intros until a3; intros A0 A1 A2 A3.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ apply A0; auto.
+ assert (NZ: Int.unsigned n <> 0).
+ { red; intros. elim H. rewrite <- (Int.repr_unsigned n). rewrite H0. auto. }
+ destruct (Int.ltu n Int.iwordsize) eqn:LT.
+ exploit Int.ltu_iwordsize_inv; eauto. intros RANGE.
+ assert (0 <= Int.zwordsize - Int.unsigned n < Int.zwordsize) by omega.
+ apply A1. auto. auto.
+ unfold Int.ltu, Int.sub. rewrite Int.unsigned_repr_wordsize.
+ rewrite Int.unsigned_repr. rewrite zlt_true; auto. omega.
+ generalize Int.wordsize_max_unsigned; omega.
+ unfold Int.ltu. rewrite zlt_true; auto.
+ change (Int.unsigned Int64.iwordsize') with 64.
+ change Int.zwordsize with 32 in RANGE. omega.
+ destruct (Int.ltu n Int64.iwordsize') eqn:LT'.
+ exploit Int.ltu_inv; eauto.
+ change (Int.unsigned Int64.iwordsize') with (Int.zwordsize * 2).
+ intros RANGE.
+ assert (Int.zwordsize <= Int.unsigned n).
+ unfold Int.ltu in LT. rewrite Int.unsigned_repr_wordsize in LT.
+ destruct (zlt (Int.unsigned n) Int.zwordsize). discriminate. omega.
+ apply A2. tauto. unfold Int.ltu, Int.sub. rewrite Int.unsigned_repr_wordsize.
+ rewrite Int.unsigned_repr. rewrite zlt_true; auto. omega.
+ generalize Int.wordsize_max_unsigned; omega.
+ auto.
+Qed.
+
+Lemma eval_shllimm:
+ forall n,
+ unary_constructor_sound (fun e => shllimm hf e n) (fun v => Val.shll v (Vint n)).
+Proof.
+ unfold shllimm; red; intros.
+ apply eval_shift_imm; intros.
+ + (* n = 0 *)
+ subst n. exists x; split; auto. destruct x; simpl; auto.
+ change (Int64.shl' i Int.zero) with (Int64.shl i Int64.zero).
+ rewrite Int64.shl_zero. auto.
+ + (* 0 < n < 32 *)
+ apply eval_splitlong with (sem := fun x => Val.shll x (Vint n)); auto.
+ intros.
+ exploit eval_shlimm. eexact H4. instantiate (1 := n). intros [v1 [A1 B1]].
+ exploit eval_shlimm. eexact H5. instantiate (1 := n). intros [v2 [A2 B2]].
+ exploit eval_shruimm. eexact H5. instantiate (1 := Int.sub Int.iwordsize n). intros [v3 [A3 B3]].
+ exploit eval_or. eexact A1. eexact A3. intros [v4 [A4 B4]].
+ econstructor; split. EvalOp.
+ intros. subst. simpl in *. rewrite H1 in *. rewrite H2 in *. rewrite H3.
+ inv B1; inv B2; inv B3. simpl in B4. inv B4.
+ simpl. rewrite Int64.decompose_shl_1; auto.
+ destruct x; auto.
+ + (* 32 <= n < 64 *)
+ exploit eval_lowlong. eexact H. intros [v1 [A1 B1]].
+ exploit eval_shlimm. eexact A1. instantiate (1 := Int.sub n Int.iwordsize). intros [v2 [A2 B2]].
+ econstructor; split. EvalOp.
+ destruct x; simpl; auto.
+ destruct (Int.ltu n Int64.iwordsize'); auto.
+ simpl in B1; inv B1. simpl in B2. rewrite H1 in B2. inv B2.
+ simpl. erewrite <- Int64.decompose_shl_2. instantiate (1 := Int64.hiword i).
+ rewrite Int64.ofwords_recompose. auto. auto.
+ + (* n >= 64 *)
+ econstructor; split. eapply eval_helper_2; eauto. EvalOp. UseHelper. auto.
+Qed.
+
+Theorem eval_shll: binary_constructor_sound (shll hf) Val.shll.
+Proof.
+ unfold shll; red; intros.
+ destruct (is_intconst b) as [n|] eqn:IC.
+- (* Immediate *)
+ exploit is_intconst_sound; eauto. intros EQ; subst y; clear H0.
+ eapply eval_shllimm; eauto.
+- (* General case *)
+ econstructor; split. eapply eval_helper_2; eauto. UseHelper. auto.
+Qed.
+
+Lemma eval_shrluimm:
+ forall n,
+ unary_constructor_sound (fun e => shrluimm hf e n) (fun v => Val.shrlu v (Vint n)).
+Proof.
+ unfold shrluimm; red; intros. apply eval_shift_imm; intros.
+ + (* n = 0 *)
+ subst n. exists x; split; auto. destruct x; simpl; auto.
+ change (Int64.shru' i Int.zero) with (Int64.shru i Int64.zero).
+ rewrite Int64.shru_zero. auto.
+ + (* 0 < n < 32 *)
+ apply eval_splitlong with (sem := fun x => Val.shrlu x (Vint n)); auto.
+ intros.
+ exploit eval_shruimm. eexact H5. instantiate (1 := n). intros [v1 [A1 B1]].
+ exploit eval_shruimm. eexact H4. instantiate (1 := n). intros [v2 [A2 B2]].
+ exploit eval_shlimm. eexact H4. instantiate (1 := Int.sub Int.iwordsize n). intros [v3 [A3 B3]].
+ exploit eval_or. eexact A1. eexact A3. intros [v4 [A4 B4]].
+ econstructor; split. EvalOp.
+ intros. subst. simpl in *. rewrite H1 in *. rewrite H2 in *. rewrite H3.
+ inv B1; inv B2; inv B3. simpl in B4. inv B4.
+ simpl. rewrite Int64.decompose_shru_1; auto.
+ destruct x; auto.
+ + (* 32 <= n < 64 *)
+ exploit eval_highlong. eexact H. intros [v1 [A1 B1]].
+ exploit eval_shruimm. eexact A1. instantiate (1 := Int.sub n Int.iwordsize). intros [v2 [A2 B2]].
+ econstructor; split. EvalOp.
+ destruct x; simpl; auto.
+ destruct (Int.ltu n Int64.iwordsize'); auto.
+ simpl in B1; inv B1. simpl in B2. rewrite H1 in B2. inv B2.
+ simpl. erewrite <- Int64.decompose_shru_2. instantiate (1 := Int64.loword i).
+ rewrite Int64.ofwords_recompose. auto. auto.
+ + (* n >= 64 *)
+ econstructor; split. eapply eval_helper_2; eauto. EvalOp. UseHelper. auto.
+Qed.
+
+Theorem eval_shrlu: binary_constructor_sound (shrlu hf) Val.shrlu.
+Proof.
+ unfold shrlu; red; intros.
+ destruct (is_intconst b) as [n|] eqn:IC.
+- (* Immediate *)
+ exploit is_intconst_sound; eauto. intros EQ; subst y; clear H0.
+ eapply eval_shrluimm; eauto.
+- (* General case *)
+ econstructor; split. eapply eval_helper_2; eauto. UseHelper. auto.
+Qed.
+
+Lemma eval_shrlimm:
+ forall n,
+ unary_constructor_sound (fun e => shrlimm hf e n) (fun v => Val.shrl v (Vint n)).
+Proof.
+ unfold shrlimm; red; intros. apply eval_shift_imm; intros.
+ + (* n = 0 *)
+ subst n. exists x; split; auto. destruct x; simpl; auto.
+ change (Int64.shr' i Int.zero) with (Int64.shr i Int64.zero).
+ rewrite Int64.shr_zero. auto.
+ + (* 0 < n < 32 *)
+ apply eval_splitlong with (sem := fun x => Val.shrl x (Vint n)); auto.
+ intros.
+ exploit eval_shruimm. eexact H5. instantiate (1 := n). intros [v1 [A1 B1]].
+ exploit eval_shrimm. eexact H4. instantiate (1 := n). intros [v2 [A2 B2]].
+ exploit eval_shlimm. eexact H4. instantiate (1 := Int.sub Int.iwordsize n). intros [v3 [A3 B3]].
+ exploit eval_or. eexact A1. eexact A3. intros [v4 [A4 B4]].
+ econstructor; split. EvalOp.
+ intros. subst. simpl in *. rewrite H1 in *. rewrite H2 in *. rewrite H3.
+ inv B1; inv B2; inv B3. simpl in B4. inv B4.
+ simpl. rewrite Int64.decompose_shr_1; auto.
+ destruct x; auto.
+ + (* 32 <= n < 64 *)
+ exploit eval_highlong. eexact H. intros [v1 [A1 B1]].
+ assert (eval_expr ge sp e m (v1 :: le) (Eletvar 0) v1) by EvalOp.
+ exploit eval_shrimm. eexact H2. instantiate (1 := Int.sub n Int.iwordsize). intros [v2 [A2 B2]].
+ exploit eval_shrimm. eexact H2. instantiate (1 := Int.repr 31). intros [v3 [A3 B3]].
+ econstructor; split. EvalOp.
+ destruct x; simpl; auto.
+ destruct (Int.ltu n Int64.iwordsize'); auto.
+ simpl in B1; inv B1. simpl in B2. rewrite H1 in B2. inv B2.
+ simpl in B3. inv B3.
+ change (Int.ltu (Int.repr 31) Int.iwordsize) with true. simpl.
+ erewrite <- Int64.decompose_shr_2. instantiate (1 := Int64.loword i).
+ rewrite Int64.ofwords_recompose. auto. auto.
+ + (* n >= 64 *)
+ econstructor; split. eapply eval_helper_2; eauto. EvalOp. UseHelper. auto.
+Qed.
+
+Theorem eval_shrl: binary_constructor_sound (shrl hf) Val.shrl.
+Proof.
+ unfold shrl; red; intros.
+ destruct (is_intconst b) as [n|] eqn:IC.
+- (* Immediate *)
+ exploit is_intconst_sound; eauto. intros EQ; subst y; clear H0.
+ eapply eval_shrlimm; eauto.
+- (* General case *)
+ econstructor; split. eapply eval_helper_2; eauto. UseHelper. auto.
+Qed.
+
+Theorem eval_addl: binary_constructor_sound (addl hf) Val.addl.
+Proof.
+ unfold addl; red; intros.
+ set (default := Ebuiltin (EF_builtin (i64_add hf) sig_ll_l) (a ::: b ::: Enil)).
+ assert (DEFAULT:
+ exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.addl x y) v).
+ {
+ econstructor; split. eapply eval_builtin_2; eauto. UseHelper. auto.
+ }
+ destruct (is_longconst a) as [p|] eqn:LC1;
+ destruct (is_longconst b) as [q|] eqn:LC2.
+- exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ econstructor; split. apply eval_longconst. simpl; auto.
+- predSpec Int64.eq Int64.eq_spec p Int64.zero; auto.
+ subst p. exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ exists y; split; auto. simpl. destruct y; auto. rewrite Int64.add_zero_l; auto.
+- predSpec Int64.eq Int64.eq_spec q Int64.zero; auto.
+ subst q. exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ exists x; split; auto. destruct x; simpl; auto. rewrite Int64.add_zero; auto.
+- auto.
+Qed.
+
+Theorem eval_subl: binary_constructor_sound (subl hf) Val.subl.
+Proof.
+ unfold subl; red; intros.
+ set (default := Ebuiltin (EF_builtin (i64_sub hf) sig_ll_l) (a ::: b ::: Enil)).
+ assert (DEFAULT:
+ exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.subl x y) v).
+ {
+ econstructor; split. eapply eval_builtin_2; eauto. UseHelper. auto.
+ }
+ destruct (is_longconst a) as [p|] eqn:LC1;
+ destruct (is_longconst b) as [q|] eqn:LC2.
+- exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ econstructor; split. apply eval_longconst. simpl; auto.
+- predSpec Int64.eq Int64.eq_spec p Int64.zero; auto.
+ replace (Val.subl x y) with (Val.negl y). eapply eval_negl; eauto.
+ subst p. exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ destruct y; simpl; auto.
+- predSpec Int64.eq Int64.eq_spec q Int64.zero; auto.
+ subst q. exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ exists x; split; auto. destruct x; simpl; auto. rewrite Int64.sub_zero_l; auto.
+- auto.
+Qed.
+
+Lemma eval_mull_base: binary_constructor_sound (mull_base hf) Val.mull.
+Proof.
+ unfold mull_base; red; intros. apply eval_splitlong2; auto.
+- intros.
+ set (p := Val.mull' x2 y2). set (le1 := p :: le0).
+ assert (E1: eval_expr ge sp e m le1 (Eop Olowlong (Eletvar O ::: Enil)) (Val.loword p)) by EvalOp.
+ assert (E2: eval_expr ge sp e m le1 (Eop Ohighlong (Eletvar O ::: Enil)) (Val.hiword p)) by EvalOp.
+ exploit eval_mul. apply eval_lift. eexact H2. apply eval_lift. eexact H3.
+ instantiate (1 := p). fold le1. intros [v3 [E3 L3]].
+ exploit eval_mul. apply eval_lift. eexact H1. apply eval_lift. eexact H4.
+ instantiate (1 := p). fold le1. intros [v4 [E4 L4]].
+ exploit eval_add. eexact E2. eexact E3. intros [v5 [E5 L5]].
+ exploit eval_add. eexact E5. eexact E4. intros [v6 [E6 L6]].
+ exists (Val.longofwords v6 (Val.loword p)); split.
+ EvalOp. eapply eval_builtin_2; eauto. UseHelper.
+ intros. unfold le1, p in *; subst; simpl in *.
+ inv L3. inv L4. inv L5. simpl in L6. inv L6.
+ simpl. f_equal. symmetry. apply Int64.decompose_mul.
+- destruct x; auto; destruct y; auto.
+Qed.
+
+Lemma eval_mullimm:
+ forall n, unary_constructor_sound (fun a => mullimm hf a n) (fun v => Val.mull v (Vlong n)).
+Proof.
+ unfold mullimm; red; intros.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ subst n. econstructor; split. apply eval_longconst.
+ destruct x; simpl; auto. rewrite Int64.mul_zero. auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.one.
+ subst n. exists x; split; auto.
+ destruct x; simpl; auto. rewrite Int64.mul_one. auto.
+ destruct (Int64.is_power2 n) as [l|] eqn:P2.
+ exploit eval_shllimm. eauto. instantiate (1 := Int.repr (Int64.unsigned l)).
+ intros [v [A B]].
+ exists v; split; auto.
+ destruct x; simpl; auto.
+ erewrite Int64.mul_pow2 by eauto.
+ assert (EQ: Int.unsigned (Int.repr (Int64.unsigned l)) = Int64.unsigned l).
+ { apply Int.unsigned_repr.
+ exploit Int64.is_power2_rng; eauto.
+ assert (Int64.zwordsize < Int.max_unsigned) by (compute; auto).
+ omega.
+ }
+ simpl in B.
+ replace (Int.ltu (Int.repr (Int64.unsigned l)) Int64.iwordsize')
+ with (Int64.ltu l Int64.iwordsize) in B.
+ erewrite Int64.is_power2_range in B by eauto.
+ unfold Int64.shl' in B. rewrite EQ in B. auto.
+ unfold Int64.ltu, Int.ltu. rewrite EQ. auto.
+ apply eval_mull_base; auto. apply eval_longconst.
+Qed.
+
+Theorem eval_mull: binary_constructor_sound (mull hf) Val.mull.
+Proof.
+ unfold mull; red; intros.
+ destruct (is_longconst a) as [p|] eqn:LC1;
+ destruct (is_longconst b) as [q|] eqn:LC2.
+- exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ econstructor; split. apply eval_longconst. simpl; auto.
+- exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ replace (Val.mull (Vlong p) y) with (Val.mull y (Vlong p)) in *.
+ eapply eval_mullimm; eauto.
+ destruct y; simpl; auto. rewrite Int64.mul_commut; auto.
+- exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ eapply eval_mullimm; eauto.
+- apply eval_mull_base; auto.
+Qed.
+
+Lemma eval_binop_long:
+ forall id sem le a b x y z,
+ (forall p q, x = Vlong p -> y = Vlong q -> z = Vlong (sem p q)) ->
+ helper_implements ge id sig_ll_l (x::y::nil) z ->
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (binop_long id sem a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold binop_long.
+ destruct (is_longconst a) as [p|] eqn:LC1.
+ destruct (is_longconst b) as [q|] eqn:LC2.
+ exploit is_longconst_sound. eexact LC1. eauto. intros EQ; subst x.
+ exploit is_longconst_sound. eexact LC2. eauto. intros EQ; subst y.
+ econstructor; split. EvalOp. erewrite H by eauto. rewrite Int64.ofwords_recompose. auto.
+ econstructor; split. eapply eval_helper_2; eauto. auto.
+ econstructor; split. eapply eval_helper_2; eauto. auto.
+Qed.
+
+Theorem eval_divl:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.divls x y = Some z ->
+ exists v, eval_expr ge sp e m le (divl hf a b) v /\ Val.lessdef z v.
+Proof.
+ intros. eapply eval_binop_long; eauto.
+ intros; subst; simpl in H1.
+ destruct (Int64.eq q Int64.zero
+ || Int64.eq p (Int64.repr Int64.min_signed) && Int64.eq q Int64.mone); inv H1.
+ auto.
+ UseHelper.
+Qed.
+
+Theorem eval_modl:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.modls x y = Some z ->
+ exists v, eval_expr ge sp e m le (modl hf a b) v /\ Val.lessdef z v.
+Proof.
+ intros. eapply eval_binop_long; eauto.
+ intros; subst; simpl in H1.
+ destruct (Int64.eq q Int64.zero
+ || Int64.eq p (Int64.repr Int64.min_signed) && Int64.eq q Int64.mone); inv H1.
+ auto.
+ UseHelper.
+Qed.
+
+Theorem eval_divlu:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.divlu x y = Some z ->
+ exists v, eval_expr ge sp e m le (divlu hf a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold divlu.
+ set (default := Eexternal (i64_udiv hf) sig_ll_l (a ::: b ::: Enil)).
+ assert (DEFAULT:
+ exists v, eval_expr ge sp e m le default v /\ Val.lessdef z v).
+ {
+ econstructor; split. eapply eval_helper_2; eauto. UseHelper. auto.
+ }
+ destruct (is_longconst a) as [p|] eqn:LC1;
+ destruct (is_longconst b) as [q|] eqn:LC2.
+- exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ econstructor; split. apply eval_longconst.
+ simpl in H1. destruct (Int64.eq q Int64.zero); inv H1. auto.
+- auto.
+- destruct (Int64.is_power2 q) as [l|] eqn:P2; auto.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ replace z with (Val.shrlu x (Vint (Int.repr (Int64.unsigned l)))).
+ apply eval_shrluimm. auto.
+ destruct x; simpl in H1; try discriminate.
+ destruct (Int64.eq q Int64.zero); inv H1.
+ simpl.
+ assert (EQ: Int.unsigned (Int.repr (Int64.unsigned l)) = Int64.unsigned l).
+ { apply Int.unsigned_repr.
+ exploit Int64.is_power2_rng; eauto.
+ assert (Int64.zwordsize < Int.max_unsigned) by (compute; auto).
+ omega.
+ }
+ replace (Int.ltu (Int.repr (Int64.unsigned l)) Int64.iwordsize')
+ with (Int64.ltu l Int64.iwordsize).
+ erewrite Int64.is_power2_range by eauto.
+ erewrite Int64.divu_pow2 by eauto.
+ unfold Int64.shru', Int64.shru. rewrite EQ. auto.
+ unfold Int64.ltu, Int.ltu. rewrite EQ. auto.
+- auto.
+Qed.
+
+Theorem eval_modlu:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.modlu x y = Some z ->
+ exists v, eval_expr ge sp e m le (modlu hf a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold modlu.
+ set (default := Eexternal (i64_umod hf) sig_ll_l (a ::: b ::: Enil)).
+ assert (DEFAULT:
+ exists v, eval_expr ge sp e m le default v /\ Val.lessdef z v).
+ {
+ econstructor; split. eapply eval_helper_2; eauto. UseHelper. auto.
+ }
+ destruct (is_longconst a) as [p|] eqn:LC1;
+ destruct (is_longconst b) as [q|] eqn:LC2.
+- exploit (is_longconst_sound le a); eauto. intros EQ; subst x.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ econstructor; split. apply eval_longconst.
+ simpl in H1. destruct (Int64.eq q Int64.zero); inv H1. auto.
+- auto.
+- destruct (Int64.is_power2 q) as [l|] eqn:P2; auto.
+ exploit (is_longconst_sound le b); eauto. intros EQ; subst y.
+ replace z with (Val.andl x (Vlong (Int64.sub q Int64.one))).
+ apply eval_andl. auto. apply eval_longconst.
+ destruct x; simpl in H1; try discriminate.
+ destruct (Int64.eq q Int64.zero); inv H1.
+ simpl.
+ erewrite Int64.modu_and by eauto. auto.
+- auto.
+Qed.
+
+Remark decompose_cmpl_eq_zero:
+ forall h l,
+ Int64.eq (Int64.ofwords h l) Int64.zero = Int.eq (Int.or h l) Int.zero.
+Proof.
+ intros.
+ assert (Int64.zwordsize = Int.zwordsize * 2) by reflexivity.
+ predSpec Int64.eq Int64.eq_spec (Int64.ofwords h l) Int64.zero.
+ replace (Int.or h l) with Int.zero. rewrite Int.eq_true. auto.
+ apply Int.same_bits_eq; intros.
+ rewrite Int.bits_zero. rewrite Int.bits_or by auto.
+ symmetry. apply orb_false_intro.
+ transitivity (Int64.testbit (Int64.ofwords h l) (i + Int.zwordsize)).
+ rewrite Int64.bits_ofwords by omega. rewrite zlt_false by omega. f_equal; omega.
+ rewrite H0. apply Int64.bits_zero.
+ transitivity (Int64.testbit (Int64.ofwords h l) i).
+ rewrite Int64.bits_ofwords by omega. rewrite zlt_true by omega. auto.
+ rewrite H0. apply Int64.bits_zero.
+ symmetry. apply Int.eq_false. red; intros; elim H0.
+ apply Int64.same_bits_eq; intros.
+ rewrite Int64.bits_zero. rewrite Int64.bits_ofwords by auto.
+ destruct (zlt i Int.zwordsize).
+ assert (Int.testbit (Int.or h l) i = false) by (rewrite H1; apply Int.bits_zero).
+ rewrite Int.bits_or in H3 by omega. exploit orb_false_elim; eauto. tauto.
+ assert (Int.testbit (Int.or h l) (i - Int.zwordsize) = false) by (rewrite H1; apply Int.bits_zero).
+ rewrite Int.bits_or in H3 by omega. exploit orb_false_elim; eauto. tauto.
+Qed.
+
+Lemma eval_cmpl_eq_zero:
+ unary_constructor_sound cmpl_eq_zero (fun v => Val.cmpl Ceq v (Vlong Int64.zero)).
+Proof.
+ red; intros. unfold cmpl_eq_zero.
+ apply eval_splitlong with (sem := fun x => Val.cmpl Ceq x (Vlong Int64.zero)); auto.
+- intros.
+ exploit eval_or. eexact H0. eexact H1. intros [v1 [A1 B1]].
+ exploit eval_comp. eexact A1. instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
+ instantiate (1 := Ceq). intros [v2 [A2 B2]].
+ exists v2; split. auto. intros; subst.
+ simpl in B1; inv B1. unfold Val.cmp in B2; simpl in B2.
+ unfold Val.cmpl; simpl. rewrite decompose_cmpl_eq_zero.
+ destruct (Int.eq (Int.or p q)); inv B2; auto.
+- destruct x; auto.
+Qed.
+
+Lemma eval_cmpl_ne_zero:
+ unary_constructor_sound cmpl_ne_zero (fun v => Val.cmpl Cne v (Vlong Int64.zero)).
+Proof.
+ red; intros. unfold cmpl_eq_zero.
+ apply eval_splitlong with (sem := fun x => Val.cmpl Cne x (Vlong Int64.zero)); auto.
+- intros.
+ exploit eval_or. eexact H0. eexact H1. intros [v1 [A1 B1]].
+ exploit eval_comp. eexact A1. instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
+ instantiate (1 := Cne). intros [v2 [A2 B2]].
+ exists v2; split. auto. intros; subst.
+ simpl in B1; inv B1. unfold Val.cmp in B2; simpl in B2.
+ unfold Val.cmpl; simpl. rewrite decompose_cmpl_eq_zero.
+ destruct (Int.eq (Int.or p q)); inv B2; auto.
+- destruct x; auto.
+Qed.
+
+Remark int64_eq_xor:
+ forall p q, Int64.eq p q = Int64.eq (Int64.xor p q) Int64.zero.
+Proof.
+ intros.
+ predSpec Int64.eq Int64.eq_spec p q.
+ subst q. rewrite Int64.xor_idem. rewrite Int64.eq_true. auto.
+ predSpec Int64.eq Int64.eq_spec (Int64.xor p q) Int64.zero.
+ elim H. apply Int64.xor_zero_equal; auto.
+ auto.
+Qed.
+
+Theorem eval_cmplu: forall c, binary_constructor_sound (cmplu hf c) (Val.cmplu c).
+Proof.
+ intros; red; intros. unfold cmplu.
+ set (default := comp c (Eexternal (i64_ucmp hf) sig_ll_i (a ::: b ::: Enil))
+ (Eop (Ointconst Int.zero) Enil)).
+ assert (DEFAULT:
+ exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.cmplu c x y) v).
+ {
+ assert (HELP: exists z, helper_implements ge hf.(i64_ucmp) sig_ll_i (x::y::nil) z
+ /\ Val.cmplu c x y = Val.cmp c z Vzero)
+ by UseHelper.
+ destruct HELP as [z [A B]].
+ exploit eval_comp. eapply eval_helper_2. eexact H. eexact H0. eauto.
+ instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
+ instantiate (1 := c). fold Vzero. intros [v [C D]].
+ econstructor; split; eauto. congruence.
+ }
+ destruct c; auto.
+- (* Ceq *)
+ destruct (is_longconst_zero b) eqn:LC.
++ exploit is_longconst_zero_sound; eauto. intros EQ; subst; clear H0.
+ apply eval_cmpl_eq_zero; auto.
++ exploit eval_xorl. eexact H. eexact H0. intros [v [A B]].
+ exploit eval_cmpl_eq_zero. eexact A. intros [v' [C D]].
+ exists v'; split; auto.
+ eapply Val.lessdef_trans; [idtac|eexact D].
+ destruct x; auto. destruct y; auto. simpl in B; inv B.
+ unfold Val.cmplu, Val.cmpl; simpl. rewrite int64_eq_xor; auto.
+- (* Cne *)
+ destruct (is_longconst_zero b) eqn:LC.
++ exploit is_longconst_zero_sound; eauto. intros EQ; subst; clear H0.
+ apply eval_cmpl_ne_zero; auto.
++ exploit eval_xorl. eexact H. eexact H0. intros [v [A B]].
+ exploit eval_cmpl_ne_zero. eexact A. intros [v' [C D]].
+ exists v'; split; auto.
+ eapply Val.lessdef_trans; [idtac|eexact D].
+ destruct x; auto. destruct y; auto. simpl in B; inv B.
+ unfold Val.cmplu, Val.cmpl; simpl. rewrite int64_eq_xor; auto.
+Qed.
+
+Remark decompose_cmpl_lt_zero:
+ forall h l,
+ Int64.lt (Int64.ofwords h l) Int64.zero = Int.lt h Int.zero.
+Proof.
+ intros.
+ generalize (Int64.shru_lt_zero (Int64.ofwords h l)).
+ change (Int64.shru (Int64.ofwords h l) (Int64.repr (Int64.zwordsize - 1)))
+ with (Int64.shru' (Int64.ofwords h l) (Int.repr 63)).
+ rewrite Int64.decompose_shru_2.
+ change (Int.sub (Int.repr 63) Int.iwordsize)
+ with (Int.repr (Int.zwordsize - 1)).
+ rewrite Int.shru_lt_zero.
+ destruct (Int64.lt (Int64.ofwords h l) Int64.zero); destruct (Int.lt h Int.zero); auto; intros.
+ elim Int64.one_not_zero. auto.
+ elim Int64.one_not_zero. auto.
+ vm_compute. intuition congruence.
+Qed.
+
+Theorem eval_cmpl: forall c, binary_constructor_sound (cmpl hf c) (Val.cmpl c).
+Proof.
+ intros; red; intros. unfold cmpl.
+ set (default := comp c (Eexternal (i64_scmp hf) sig_ll_i (a ::: b ::: Enil))
+ (Eop (Ointconst Int.zero) Enil)).
+ assert (DEFAULT:
+ exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.cmpl c x y) v).
+ {
+ assert (HELP: exists z, helper_implements ge hf.(i64_scmp) sig_ll_i (x::y::nil) z
+ /\ Val.cmpl c x y = Val.cmp c z Vzero)
+ by UseHelper.
+ destruct HELP as [z [A B]].
+ exploit eval_comp. eapply eval_helper_2. eexact H. eexact H0. eauto.
+ instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
+ instantiate (1 := c). fold Vzero. intros [v [C D]].
+ econstructor; split; eauto. congruence.
+ }
+ destruct c; auto.
+- (* Ceq *)
+ destruct (is_longconst_zero b) eqn:LC.
++ exploit is_longconst_zero_sound; eauto. intros EQ; subst; clear H0.
+ apply eval_cmpl_eq_zero; auto.
++ exploit eval_xorl. eexact H. eexact H0. intros [v [A B]].
+ exploit eval_cmpl_eq_zero. eexact A. intros [v' [C D]].
+ exists v'; split; auto.
+ eapply Val.lessdef_trans; [idtac|eexact D].
+ destruct x; auto. destruct y; auto. simpl in B; inv B.
+ unfold Val.cmpl; simpl. rewrite int64_eq_xor; auto.
+- (* Cne *)
+ destruct (is_longconst_zero b) eqn:LC.
++ exploit is_longconst_zero_sound; eauto. intros EQ; subst; clear H0.
+ apply eval_cmpl_ne_zero; auto.
++ exploit eval_xorl. eexact H. eexact H0. intros [v [A B]].
+ exploit eval_cmpl_ne_zero. eexact A. intros [v' [C D]].
+ exists v'; split; auto.
+ eapply Val.lessdef_trans; [idtac|eexact D].
+ destruct x; auto. destruct y; auto. simpl in B; inv B.
+ unfold Val.cmpl; simpl. rewrite int64_eq_xor; auto.
+- (* Clt *)
+ destruct (is_longconst_zero b) eqn:LC.
++ exploit is_longconst_zero_sound; eauto. intros EQ; subst; clear H0.
+ exploit eval_highlong. eexact H. intros [v1 [A1 B1]].
+ exploit eval_comp. eexact A1.
+ instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
+ instantiate (1 := Clt). intros [v2 [A2 B2]].
+ econstructor; split. eauto.
+ destruct x; simpl in *; auto. inv B1.
+ unfold Val.cmp, Val.cmp_bool, Val.of_optbool, Int.cmp in B2.
+ unfold Val.cmpl, Val.cmpl_bool, Val.of_optbool, Int64.cmp.
+ rewrite <- (Int64.ofwords_recompose i). rewrite decompose_cmpl_lt_zero.
+ auto.
++ apply DEFAULT.
+- (* Cge *)
+ destruct (is_longconst_zero b) eqn:LC.
++ exploit is_longconst_zero_sound; eauto. intros EQ; subst; clear H0.
+ exploit eval_highlong. eexact H. intros [v1 [A1 B1]].
+ exploit eval_comp. eexact A1.
+ instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
+ instantiate (1 := Cge). intros [v2 [A2 B2]].
+ econstructor; split. eauto.
+ destruct x; simpl in *; auto. inv B1.
+ unfold Val.cmp, Val.cmp_bool, Val.of_optbool, Int.cmp in B2.
+ unfold Val.cmpl, Val.cmpl_bool, Val.of_optbool, Int64.cmp.
+ rewrite <- (Int64.ofwords_recompose i). rewrite decompose_cmpl_lt_zero.
+ auto.
++ apply DEFAULT.
+Qed.
+
+End CMCONSTR.
+
diff --git a/backend/Selection.v b/backend/Selection.v
index 29bdabc..7964feb 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -24,12 +24,14 @@
Require Import Coqlib.
Require Import AST.
+Require Import Errors.
Require Import Integers.
Require Import Globalenvs.
Require Cminor.
Require Import Op.
Require Import CminorSel.
Require Import SelectOp.
+Require Import SelectLong.
Open Local Scope cminorsel_scope.
@@ -55,12 +57,17 @@ Definition store (chunk: memory_chunk) (e1 e2: expr) :=
(** Instruction selection for operator applications. Most of the work
is done by the processor-specific smart constructors defined
- in module [SelectOp]. *)
+ in modules [SelectOp] and [SelectLong]. *)
+
+Section SELECTION.
+
+Variable hf: helper_functions.
Definition sel_constant (cst: Cminor.constant) : expr :=
match cst with
| Cminor.Ointconst n => Eop (Ointconst n) Enil
| Cminor.Ofloatconst f => Eop (Ofloatconst f) Enil
+ | Cminor.Olongconst n => longconst n
| Cminor.Oaddrsymbol id ofs => addrsymbol id ofs
| Cminor.Oaddrstack ofs => addrstack ofs
end.
@@ -80,6 +87,15 @@ Definition sel_unop (op: Cminor.unary_operation) (arg: expr) : expr :=
| Cminor.Ointuoffloat => intuoffloat arg
| Cminor.Ofloatofint => floatofint arg
| Cminor.Ofloatofintu => floatofintu arg
+ | Cminor.Onegl => negl hf arg
+ | Cminor.Onotl => notl arg
+ | Cminor.Ointoflong => intoflong arg
+ | Cminor.Olongofint => longofint arg
+ | Cminor.Olongofintu => longofintu arg
+ | Cminor.Olongoffloat => longoffloat hf arg
+ | Cminor.Olonguoffloat => longuoffloat hf arg
+ | Cminor.Ofloatoflong => floatoflong hf arg
+ | Cminor.Ofloatoflongu => floatoflongu hf arg
end.
Definition sel_binop (op: Cminor.binary_operation) (arg1 arg2: expr) : expr :=
@@ -101,9 +117,24 @@ Definition sel_binop (op: Cminor.binary_operation) (arg1 arg2: expr) : expr :=
| Cminor.Osubf => subf arg1 arg2
| Cminor.Omulf => mulf arg1 arg2
| Cminor.Odivf => divf arg1 arg2
+ | Cminor.Oaddl => addl hf arg1 arg2
+ | Cminor.Osubl => subl hf arg1 arg2
+ | Cminor.Omull => mull hf arg1 arg2
+ | Cminor.Odivl => divl hf arg1 arg2
+ | Cminor.Odivlu => divlu hf arg1 arg2
+ | Cminor.Omodl => modl hf arg1 arg2
+ | Cminor.Omodlu => modlu hf arg1 arg2
+ | Cminor.Oandl => andl arg1 arg2
+ | Cminor.Oorl => orl arg1 arg2
+ | Cminor.Oxorl => xorl arg1 arg2
+ | Cminor.Oshll => shll hf arg1 arg2
+ | Cminor.Oshrl => shrl hf arg1 arg2
+ | Cminor.Oshrlu => shrlu hf arg1 arg2
| Cminor.Ocmp c => comp c arg1 arg2
| Cminor.Ocmpu c => compu c arg1 arg2
| Cminor.Ocmpf c => compf c arg1 arg2
+ | Cminor.Ocmpl c => cmpl hf c arg1 arg2
+ | Cminor.Ocmplu c => cmplu hf c arg1 arg2
end.
(** Conversion from Cminor expression to Cminorsel expressions *)
@@ -186,22 +217,26 @@ Fixpoint sel_stmt (ge: Cminor.genv) (s: Cminor.stmt) : stmt :=
| Cminor.Sgoto lbl => Sgoto lbl
end.
-(** Conversion of functions and programs. *)
+End SELECTION.
+
+(** Conversion of functions. *)
-Definition sel_function (ge: Cminor.genv) (f: Cminor.function) : function :=
+Definition sel_function (hf: helper_functions) (ge: Cminor.genv) (f: Cminor.function) : function :=
mkfunction
f.(Cminor.fn_sig)
f.(Cminor.fn_params)
f.(Cminor.fn_vars)
f.(Cminor.fn_stackspace)
- (sel_stmt ge f.(Cminor.fn_body)).
+ (sel_stmt hf ge f.(Cminor.fn_body)).
-Definition sel_fundef (ge: Cminor.genv) (f: Cminor.fundef) : fundef :=
- transf_fundef (sel_function ge) f.
+Definition sel_fundef (hf: helper_functions) (ge: Cminor.genv) (f: Cminor.fundef) : fundef :=
+ transf_fundef (sel_function hf ge) f.
-Definition sel_program (p: Cminor.program) : program :=
- let ge := Genv.globalenv p in
- transform_program (sel_fundef ge) p.
+(** Conversion of programs. *)
+Local Open Scope error_monad_scope.
+Definition sel_program (p: Cminor.program) : res program :=
+ let ge := Genv.globalenv p in
+ do hf <- get_helpers ge; OK (transform_program (sel_fundef hf ge) p).
diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v
index 0269438..525a29d 100644
--- a/backend/Selectionproof.v
+++ b/backend/Selectionproof.v
@@ -15,6 +15,7 @@
Require Import Coqlib.
Require Import Maps.
Require Import AST.
+Require Import Errors.
Require Import Integers.
Require Import Values.
Require Import Memory.
@@ -25,27 +26,114 @@ Require Import Cminor.
Require Import Op.
Require Import CminorSel.
Require Import SelectOp.
+Require Import SelectLong.
Require Import Selection.
Require Import SelectOpproof.
+Require Import SelectLongproof.
Open Local Scope cminorsel_scope.
+
(** * Correctness of the instruction selection functions for expressions *)
+Section PRESERVATION.
+
+Variable prog: Cminor.program.
+Let ge := Genv.globalenv prog.
+Variable hf: helper_functions.
+Let tprog := transform_program (sel_fundef hf ge) prog.
+Let tge := Genv.globalenv tprog.
+Hypothesis HELPERS: i64_helpers_correct tge hf.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof.
+ intros; unfold ge, tge, tprog. apply Genv.find_symbol_transf.
+Qed.
+
+Lemma function_ptr_translated:
+ forall (b: block) (f: Cminor.fundef),
+ Genv.find_funct_ptr ge b = Some f ->
+ Genv.find_funct_ptr tge b = Some (sel_fundef hf ge f).
+Proof.
+ intros.
+ exact (Genv.find_funct_ptr_transf (sel_fundef hf ge) _ _ H).
+Qed.
+
+Lemma functions_translated:
+ forall (v v': val) (f: Cminor.fundef),
+ Genv.find_funct ge v = Some f ->
+ Val.lessdef v v' ->
+ Genv.find_funct tge v' = Some (sel_fundef hf ge f).
+Proof.
+ intros. inv H0.
+ exact (Genv.find_funct_transf (sel_fundef hf ge) _ _ H).
+ simpl in H. discriminate.
+Qed.
+
+Lemma sig_function_translated:
+ forall f,
+ funsig (sel_fundef hf ge f) = Cminor.funsig f.
+Proof.
+ intros. destruct f; reflexivity.
+Qed.
+
+Lemma varinfo_preserved:
+ forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
+Proof.
+ intros; unfold ge, tge, tprog, sel_program.
+ apply Genv.find_var_info_transf.
+Qed.
+
+Lemma helper_implements_preserved:
+ forall id sg vargs vres,
+ helper_implements ge id sg vargs vres ->
+ helper_implements tge id sg vargs vres.
+Proof.
+ intros. destruct H as (b & ef & A & B & C & D).
+ exploit function_ptr_translated; eauto. simpl. intros.
+ exists b; exists ef.
+ split. rewrite symbols_preserved. auto.
+ split. auto.
+ split. auto.
+ intros. eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+Qed.
+
+Lemma builtin_implements_preserved:
+ forall id sg vargs vres,
+ builtin_implements ge id sg vargs vres ->
+ builtin_implements tge id sg vargs vres.
+Proof.
+ unfold builtin_implements; intros.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+Qed.
+
+Lemma helpers_correct_preserved:
+ forall h, i64_helpers_correct ge h -> i64_helpers_correct tge h.
+Proof.
+ unfold i64_helpers_correct; intros.
+ repeat (match goal with [ H: _ /\ _ |- _ /\ _ ] => destruct H; split end);
+ intros; try (eapply helper_implements_preserved; eauto);
+ try (eapply builtin_implements_preserved; eauto).
+ exploit H14; eauto. intros [z [A B]]. exists z; split; eauto. eapply helper_implements_preserved; eauto.
+ exploit H15; eauto. intros [z [A B]]. exists z; split; eauto. eapply helper_implements_preserved; eauto.
+Qed.
+
Section CMCONSTR.
-Variable ge: genv.
Variable sp: val.
Variable e: env.
Variable m: mem.
Lemma eval_condition_of_expr:
forall le a v b,
- eval_expr ge sp e m le a v ->
+ eval_expr tge sp e m le a v ->
Val.bool_of_val v b ->
match condition_of_expr a with (cond, args) =>
exists vl,
- eval_exprlist ge sp e m le args vl /\
+ eval_exprlist tge sp e m le args vl /\
eval_condition cond vl m = Some b
end.
Proof.
@@ -60,9 +148,9 @@ Qed.
Lemma eval_load:
forall le a v chunk v',
- eval_expr ge sp e m le a v ->
+ eval_expr tge sp e m le a v ->
Mem.loadv chunk m v = Some v' ->
- eval_expr ge sp e m le (load chunk a) v'.
+ eval_expr tge sp e m le (load chunk a) v'.
Proof.
intros. generalize H0; destruct v; simpl; intro; try discriminate.
unfold load.
@@ -73,11 +161,11 @@ Qed.
Lemma eval_store:
forall chunk a1 a2 v1 v2 f k m',
- eval_expr ge sp e m nil a1 v1 ->
- eval_expr ge sp e m nil a2 v2 ->
+ eval_expr tge sp e m nil a1 v1 ->
+ eval_expr tge sp e m nil a2 v2 ->
Mem.storev chunk m v1 v2 = Some m' ->
- step ge (State f (store chunk a1 a2) k sp e m)
- E0 (State f Sskip k sp e m').
+ step tge (State f (store chunk a1 a2) k sp e m)
+ E0 (State f Sskip k sp e m').
Proof.
intros. generalize H1; destruct v1; simpl; intro; try discriminate.
unfold store.
@@ -90,9 +178,9 @@ Qed.
Lemma eval_sel_unop:
forall le op a1 v1 v,
- eval_expr ge sp e m le a1 v1 ->
+ eval_expr tge sp e m le a1 v1 ->
eval_unop op v1 = Some v ->
- exists v', eval_expr ge sp e m le (sel_unop op a1) v' /\ Val.lessdef v v'.
+ exists v', eval_expr tge sp e m le (sel_unop hf op a1) v' /\ Val.lessdef v v'.
Proof.
destruct op; simpl; intros; FuncInv; try subst v.
apply eval_cast8unsigned; auto.
@@ -108,14 +196,23 @@ Proof.
eapply eval_intuoffloat; eauto.
eapply eval_floatofint; eauto.
eapply eval_floatofintu; eauto.
+ eapply eval_negl; eauto.
+ eapply eval_notl; eauto.
+ eapply eval_intoflong; eauto.
+ eapply eval_longofint; eauto.
+ eapply eval_longofintu; eauto.
+ eapply eval_longoffloat; eauto.
+ eapply eval_longuoffloat; eauto.
+ eapply eval_floatoflong; eauto.
+ eapply eval_floatoflongu; eauto.
Qed.
Lemma eval_sel_binop:
forall le op a1 a2 v1 v2 v,
- eval_expr ge sp e m le a1 v1 ->
- eval_expr ge sp e m le a2 v2 ->
+ eval_expr tge sp e m le a1 v1 ->
+ eval_expr tge sp e m le a2 v2 ->
eval_binop op v1 v2 m = Some v ->
- exists v', eval_expr ge sp e m le (sel_binop op a1 a2) v' /\ Val.lessdef v v'.
+ exists v', eval_expr tge sp e m le (sel_binop hf op a1 a2) v' /\ Val.lessdef v v'.
Proof.
destruct op; simpl; intros; FuncInv; try subst v.
apply eval_add; auto.
@@ -135,9 +232,24 @@ Proof.
apply eval_subf; auto.
apply eval_mulf; auto.
apply eval_divf; auto.
+ eapply eval_addl; eauto.
+ eapply eval_subl; eauto.
+ eapply eval_mull; eauto.
+ eapply eval_divl; eauto.
+ eapply eval_divlu; eauto.
+ eapply eval_modl; eauto.
+ eapply eval_modlu; eauto.
+ eapply eval_andl; eauto.
+ eapply eval_orl; eauto.
+ eapply eval_xorl; eauto.
+ eapply eval_shll; eauto.
+ eapply eval_shrl; eauto.
+ eapply eval_shrlu; eauto.
apply eval_comp; auto.
apply eval_compu; auto.
apply eval_compf; auto.
+ apply eval_cmpl; auto.
+ apply eval_cmplu; auto.
Qed.
End CMCONSTR.
@@ -156,7 +268,7 @@ Proof.
Qed.
Lemma classify_call_correct:
- forall ge sp e m a v fd,
+ forall sp e m a v fd,
Cminor.eval_expr ge sp e m a v ->
Genv.find_funct ge v = Some fd ->
match classify_call ge a with
@@ -215,57 +327,6 @@ Qed.
(** * Semantic preservation for instruction selection. *)
-Section PRESERVATION.
-
-Variable prog: Cminor.program.
-Let tprog := sel_program prog.
-Let ge := Genv.globalenv prog.
-Let tge := Genv.globalenv tprog.
-
-(** Relationship between the global environments for the original
- Cminor program and the generated CminorSel program. *)
-
-Lemma symbols_preserved:
- forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof.
- intros; unfold ge, tge, tprog, sel_program.
- apply Genv.find_symbol_transf.
-Qed.
-
-Lemma function_ptr_translated:
- forall (b: block) (f: Cminor.fundef),
- Genv.find_funct_ptr ge b = Some f ->
- Genv.find_funct_ptr tge b = Some (sel_fundef ge f).
-Proof.
- intros.
- exact (Genv.find_funct_ptr_transf (sel_fundef ge) _ _ H).
-Qed.
-
-Lemma functions_translated:
- forall (v v': val) (f: Cminor.fundef),
- Genv.find_funct ge v = Some f ->
- Val.lessdef v v' ->
- Genv.find_funct tge v' = Some (sel_fundef ge f).
-Proof.
- intros. inv H0.
- exact (Genv.find_funct_transf (sel_fundef ge) _ _ H).
- simpl in H. discriminate.
-Qed.
-
-Lemma sig_function_translated:
- forall f,
- funsig (sel_fundef ge f) = Cminor.funsig f.
-Proof.
- intros. destruct f; reflexivity.
-Qed.
-
-Lemma varinfo_preserved:
- forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
-Proof.
- intros; unfold ge, tge, tprog, sel_program.
- apply Genv.find_var_info_transf.
-Qed.
-
(** Relationship between the local environments. *)
Definition env_lessdef (e1 e2: env) : Prop :=
@@ -305,7 +366,7 @@ Lemma sel_expr_correct:
Cminor.eval_expr ge sp e m a v ->
forall e' le m',
env_lessdef e e' -> Mem.extends m m' ->
- exists v', eval_expr tge sp e' m' le (sel_expr a) v' /\ Val.lessdef v v'.
+ exists v', eval_expr tge sp e' m' le (sel_expr hf a) v' /\ Val.lessdef v v'.
Proof.
induction 1; intros; simpl.
(* Evar *)
@@ -314,6 +375,9 @@ Proof.
destruct cst; simpl in *; inv H.
exists (Vint i); split; auto. econstructor. constructor. auto.
exists (Vfloat f); split; auto. econstructor. constructor. auto.
+ exists (Val.longofwords (Vint (Int64.hiword i)) (Vint (Int64.loword i))); split.
+ eapply eval_Eop. constructor. EvalOp. simpl; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
+ simpl. rewrite Int64.ofwords_recompose. auto.
rewrite <- symbols_preserved. fold (symbol_address tge i i0). apply eval_addrsymbol.
apply eval_addrstack.
(* Eunop *)
@@ -338,7 +402,7 @@ Lemma sel_exprlist_correct:
Cminor.eval_exprlist ge sp e m a v ->
forall e' le m',
env_lessdef e e' -> Mem.extends m m' ->
- exists v', eval_exprlist tge sp e' m' le (sel_exprlist a) v' /\ Val.lessdef_list v v'.
+ exists v', eval_exprlist tge sp e' m' le (sel_exprlist hf a) v' /\ Val.lessdef_list v v'.
Proof.
induction 1; intros; simpl.
exists (@nil val); split; auto. constructor.
@@ -354,30 +418,30 @@ Inductive match_cont: Cminor.cont -> CminorSel.cont -> Prop :=
match_cont Cminor.Kstop Kstop
| match_cont_seq: forall s k k',
match_cont k k' ->
- match_cont (Cminor.Kseq s k) (Kseq (sel_stmt ge s) k')
+ match_cont (Cminor.Kseq s k) (Kseq (sel_stmt hf ge s) k')
| match_cont_block: forall k k',
match_cont k k' ->
match_cont (Cminor.Kblock k) (Kblock k')
| match_cont_call: forall id f sp e k e' k',
match_cont k k' -> env_lessdef e e' ->
- match_cont (Cminor.Kcall id f sp e k) (Kcall id (sel_function ge f) sp e' k').
+ match_cont (Cminor.Kcall id f sp e k) (Kcall id (sel_function hf ge f) sp e' k').
Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
| match_state: forall f s k s' k' sp e m e' m',
- s' = sel_stmt ge s ->
+ s' = sel_stmt hf ge s ->
match_cont k k' ->
env_lessdef e e' ->
Mem.extends m m' ->
match_states
(Cminor.State f s k sp e m)
- (State (sel_function ge f) s' k' sp e' m')
+ (State (sel_function hf ge f) s' k' sp e' m')
| match_callstate: forall f args args' k k' m m',
match_cont k k' ->
Val.lessdef_list args args' ->
Mem.extends m m' ->
match_states
(Cminor.Callstate f args k m)
- (Callstate (sel_fundef ge f) args' k' m')
+ (Callstate (sel_fundef hf ge f) args' k' m')
| match_returnstate: forall v v' k k' m m',
match_cont k k' ->
Val.lessdef v v' ->
@@ -393,7 +457,7 @@ Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
eval_exprlist tge sp e' m' nil al args' ->
match_states
(Cminor.Callstate (External ef) args (Cminor.Kcall optid f sp e k) m)
- (State (sel_function ge f) (Sbuiltin optid ef al) k' sp e' m')
+ (State (sel_function hf ge f) (Sbuiltin optid ef al) k' sp e' m')
| match_builtin_2: forall v v' optid f sp e k m e' m' k',
match_cont k k' ->
Val.lessdef v v' ->
@@ -401,7 +465,7 @@ Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
Mem.extends m m' ->
match_states
(Cminor.Returnstate v (Cminor.Kcall optid f sp e k) m)
- (State (sel_function ge f) Sskip k' sp (set_optvar optid v' e') m').
+ (State (sel_function hf ge f) Sskip k' sp (set_optvar optid v' e') m').
Remark call_cont_commut:
forall k k', match_cont k k' -> match_cont (Cminor.call_cont k) (call_cont k').
@@ -412,15 +476,15 @@ Qed.
Remark find_label_commut:
forall lbl s k k',
match_cont k k' ->
- match Cminor.find_label lbl s k, find_label lbl (sel_stmt ge s) k' with
+ match Cminor.find_label lbl s k, find_label lbl (sel_stmt hf ge s) k' with
| None, None => True
- | Some(s1, k1), Some(s1', k1') => s1' = sel_stmt ge s1 /\ match_cont k1 k1'
+ | Some(s1, k1), Some(s1', k1') => s1' = sel_stmt hf ge s1 /\ match_cont k1 k1'
| _, _ => False
end.
Proof.
induction s; intros; simpl; auto.
(* store *)
- unfold store. destruct (addressing m (sel_expr e)); simpl; auto.
+ unfold store. destruct (addressing m (sel_expr hf e)); simpl; auto.
(* call *)
destruct (classify_call ge e); simpl; auto.
(* tailcall *)
@@ -428,13 +492,13 @@ Proof.
(* seq *)
exploit (IHs1 (Cminor.Kseq s2 k)). constructor; eauto.
destruct (Cminor.find_label lbl s1 (Cminor.Kseq s2 k)) as [[sx kx] | ];
- destruct (find_label lbl (sel_stmt ge s1) (Kseq (sel_stmt ge s2) k')) as [[sy ky] | ];
+ destruct (find_label lbl (sel_stmt hf ge s1) (Kseq (sel_stmt hf ge s2) k')) as [[sy ky] | ];
intuition. apply IHs2; auto.
(* ifthenelse *)
- destruct (condition_of_expr (sel_expr e)) as [cond args]. simpl.
+ destruct (condition_of_expr (sel_expr hf e)) as [cond args]. simpl.
exploit (IHs1 k); eauto.
destruct (Cminor.find_label lbl s1 k) as [[sx kx] | ];
- destruct (find_label lbl (sel_stmt ge s1) k') as [[sy ky] | ];
+ destruct (find_label lbl (sel_stmt hf ge s1) k') as [[sy ky] | ];
intuition. apply IHs2; auto.
(* loop *)
apply IHs. constructor; auto.
@@ -531,9 +595,9 @@ Proof.
exploit sel_expr_correct; eauto. intros [v' [A B]].
assert (Val.bool_of_val v' b). inv B. auto. inv H0.
exploit eval_condition_of_expr; eauto.
- destruct (condition_of_expr (sel_expr a)) as [cond args].
+ destruct (condition_of_expr (sel_expr hf a)) as [cond args].
intros [vl' [C D]].
- left; exists (State (sel_function ge f) (if b then sel_stmt ge s1 else sel_stmt ge s2) k' sp e' m'); split.
+ left; exists (State (sel_function hf ge f) (if b then sel_stmt hf ge s1 else sel_stmt hf ge s2) k' sp e' m'); split.
econstructor; eauto.
constructor; auto. destruct b; auto.
(* Sloop *)
@@ -566,7 +630,7 @@ Proof.
exploit (find_label_commut lbl (Cminor.fn_body f) (Cminor.call_cont k)).
apply call_cont_commut; eauto.
rewrite H.
- destruct (find_label lbl (sel_stmt ge (Cminor.fn_body f)) (call_cont k'0))
+ destruct (find_label lbl (sel_stmt hf ge (Cminor.fn_body f)) (call_cont k'0))
as [[s'' k'']|] eqn:?; intros; try contradiction.
destruct H0.
left; econstructor; split.
@@ -623,14 +687,22 @@ Proof.
intros. inv H0. inv H. inv H3. inv H5. constructor.
Qed.
+End PRESERVATION.
+
+Axiom get_helpers_correct:
+ forall ge hf, get_helpers ge = OK hf -> i64_helpers_correct ge hf.
+
Theorem transf_program_correct:
+ forall prog tprog,
+ sel_program prog = OK tprog ->
forward_simulation (Cminor.semantics prog) (CminorSel.semantics tprog).
Proof.
+ intros. unfold sel_program in H.
+ destruct (get_helpers (Genv.globalenv prog)) as [hf|] eqn:E; simpl in H; try discriminate.
+ inv H.
eapply forward_simulation_opt.
- eexact symbols_preserved.
- eexact sel_initial_states.
- eexact sel_final_states.
- eexact sel_step_correct.
+ apply symbols_preserved.
+ apply sel_initial_states.
+ apply sel_final_states.
+ apply sel_step_correct. apply helpers_correct_preserved. apply get_helpers_correct. auto.
Qed.
-
-End PRESERVATION.
diff --git a/backend/Splitting.ml b/backend/Splitting.ml
new file mode 100644
index 0000000..60f6818
--- /dev/null
+++ b/backend/Splitting.ml
@@ -0,0 +1,184 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(* Live range splitting over RTL *)
+
+open Camlcoq
+open Datatypes
+open Coqlib
+open Maps
+open AST
+open Kildall
+open Registers
+open RTL
+
+(* We represent live ranges by the following unification variables.
+ Live range inference manipulates only variable live ranges.
+ Code rewriting assigns fresh RTL registers to live ranges. *)
+
+type live_range = { source: reg; mutable kind: live_range_kind }
+
+and live_range_kind =
+ | Link of live_range
+ | Var
+ | Reg of reg
+
+let new_range r = { source = r; kind = Var }
+
+let rec repr lr =
+ match lr.kind with
+ | Link lr' -> let lr'' = repr lr' in lr.kind <- Link lr''; lr''
+ | _ -> lr
+
+let same_range lr1 lr2 =
+ lr1 == lr2 || (* quick test for speed *)
+ repr lr1 == repr lr2 (* the real test *)
+
+let unify lr1 lr2 =
+ let lr1 = repr lr1 and lr2 = repr lr2 in
+ if lr1 != lr2 then begin
+ match lr1.kind, lr2.kind with
+ | Var, _ -> lr1.kind <- Link lr2
+ | _, Var -> lr2.kind <- Link lr1
+ | _, _ -> assert false
+ end
+
+let reg_for lr =
+ let lr = repr lr in
+ match lr.kind with
+ | Link _ -> assert false
+ | Reg r -> r
+ | Var -> let r = XTL.new_reg() in lr.kind <- Reg r; r
+
+(* Live range inference is a variant on liveness analysis.
+ At each PC and for each register, liveness analysis determines
+ whether the reg is live or not. Live range inference associates
+ a live range to the reg if it is live, and no live range if it
+ is dead. *)
+
+module LRMap = struct
+
+ type t = live_range PTree.t (* live register -> live range *)
+
+ let beq m1 m2 = PTree.beq same_range m1 m2
+
+ let bot : t = PTree.empty
+
+ let lub_opt_range olr1 olr2 =
+ match olr1, olr2 with
+ | Some lr1, Some lr2 -> unify lr1 lr2; olr1
+ | Some _, None -> olr1
+ | None, _ -> olr2
+
+ let lub m1 m2 =
+ PTree.combine lub_opt_range m1 m2
+
+end
+
+module Solver = Backward_Dataflow_Solver(LRMap)(NodeSetBackward)
+
+(* A cache of live ranges associated to (pc, used reg) pairs. *)
+
+let live_range_cache =
+ (Hashtbl.create 123 : (int32 * int32, live_range) Hashtbl.t)
+
+let live_range_for pc r =
+ let pc' = P.to_int32 pc
+ and r' = P.to_int32 r in
+ try
+ Hashtbl.find live_range_cache (pc',r')
+ with Not_found ->
+ let lr = new_range r in
+ Hashtbl.add live_range_cache (pc', r') lr;
+ lr
+
+(* The transfer function *)
+
+let reg_live pc r map =
+ match PTree.get r map with
+ | Some lr -> map (* already live *)
+ | None -> PTree.set r (live_range_for pc r) map (* becomes live *)
+
+let reg_list_live pc rl map = List.fold_right (reg_live pc) rl map
+
+let reg_dead r map =
+ PTree.remove r map
+
+let transfer f pc after =
+ match PTree.get pc f.fn_code with
+ | None ->
+ LRMap.bot
+ | Some i ->
+ let across =
+ match instr_defs i with
+ | None -> after
+ | Some r -> reg_dead r after in
+ reg_list_live pc (instr_uses i) across
+
+(* The live range analysis *)
+
+let analysis f = Solver.fixpoint (successors f) (transfer f) []
+
+(* Produce renamed registers for each instruction. *)
+
+let ren_reg map r =
+ match PTree.get r map with
+ | Some lr -> reg_for lr
+ | None -> XTL.new_reg()
+
+let ren_regs map rl =
+ List.map (ren_reg map) rl
+
+let ren_ros map ros =
+ sum_left_map (ren_reg map) ros
+
+(* Rename in an instruction *)
+
+let ren_instr f maps pc i =
+ let after = PMap.get pc maps in
+ let before = transfer f pc after in
+ match i with
+ | Inop s -> Inop s
+ | Iop(op, args, res, s) ->
+ Iop(op, ren_regs before args, ren_reg after res, s)
+ | Iload(chunk, addr, args, dst, s) ->
+ Iload(chunk, addr, ren_regs before args, ren_reg after dst, s)
+ | Istore(chunk, addr, args, src, s) ->
+ Istore(chunk, addr, ren_regs before args, ren_reg before src, s)
+ | Icall(sg, ros, args, res, s) ->
+ Icall(sg, ren_ros before ros, ren_regs before args, ren_reg after res, s)
+ | Itailcall(sg, ros, args) ->
+ Itailcall(sg, ren_ros before ros, ren_regs before args)
+ | Ibuiltin(ef, args, res, s) ->
+ Ibuiltin(ef, ren_regs before args, ren_reg after res, s)
+ | Icond(cond, args, s1, s2) ->
+ Icond(cond, ren_regs before args, s1, s2)
+ | Ijumptable(arg, tbl) ->
+ Ijumptable(ren_reg before arg, tbl)
+ | Ireturn optarg ->
+ Ireturn(option_map (ren_reg before) optarg)
+
+(* Rename live ranges in a function *)
+
+let rename_function f =
+ Hashtbl.clear live_range_cache;
+ let maps =
+ match analysis f with
+ | None -> assert false
+ | Some maps -> maps in
+ let before_entrypoint =
+ transfer f f.fn_entrypoint (PMap.get f.fn_entrypoint maps) in
+ { fn_sig = f.fn_sig;
+ fn_params = ren_regs before_entrypoint f.fn_params;
+ fn_stacksize = f.fn_stacksize;
+ fn_code = PTree.map (ren_instr f maps) f.fn_code;
+ fn_entrypoint = f.fn_entrypoint }
diff --git a/backend/Stacking.v b/backend/Stacking.v
index 03e882e..f7c16d1 100644
--- a/backend/Stacking.v
+++ b/backend/Stacking.v
@@ -23,6 +23,7 @@ Require Import Bounds.
Require Import Mach.
Require Import Conventions.
Require Import Stacklayout.
+Require Import Lineartyping.
(** * Layout of activation records *)
@@ -44,8 +45,7 @@ Definition offset_of_index (fe: frame_env) (idx: frame_index) :=
match idx with
| FI_link => fe.(fe_ofs_link)
| FI_retaddr => fe.(fe_ofs_retaddr)
- | FI_local x Tint => fe.(fe_ofs_int_local) + 4 * x
- | FI_local x Tfloat => fe.(fe_ofs_float_local) + 8 * x
+ | FI_local x ty => fe.(fe_ofs_local) + 4 * x
| FI_arg x ty => fe_ofs_arg + 4 * x
| FI_saved_int x => fe.(fe_ofs_int_callee_save) + 4 * x
| FI_saved_float x => fe.(fe_ofs_float_callee_save) + 8 * x
@@ -133,8 +133,8 @@ Definition transl_addr (fe: frame_env) (addr: addressing) :=
Definition transl_annot_param (fe: frame_env) (l: loc) : annot_param :=
match l with
| R r => APreg r
- | S (Local ofs ty) => APstack (chunk_of_type ty) (offset_of_index fe (FI_local ofs ty))
- | S _ => APstack Mint32 (-1) (**r never happens *)
+ | S Local ofs ty => APstack (chunk_of_type ty) (offset_of_index fe (FI_local ofs ty))
+ | S _ _ _ => APstack Mint32 (-1) (**r never happens *)
end.
@@ -150,22 +150,22 @@ Definition transl_annot_param (fe: frame_env) (l: loc) : annot_param :=
Definition transl_instr
(fe: frame_env) (i: Linear.instruction) (k: Mach.code) : Mach.code :=
match i with
- | Lgetstack s r =>
- match s with
- | Local ofs ty =>
+ | Lgetstack sl ofs ty r =>
+ match sl with
+ | Local =>
Mgetstack (Int.repr (offset_of_index fe (FI_local ofs ty))) ty r :: k
- | Incoming ofs ty =>
+ | Incoming =>
Mgetparam (Int.repr (offset_of_index fe (FI_arg ofs ty))) ty r :: k
- | Outgoing ofs ty =>
+ | Outgoing =>
Mgetstack (Int.repr (offset_of_index fe (FI_arg ofs ty))) ty r :: k
end
- | Lsetstack r s =>
- match s with
- | Local ofs ty =>
+ | Lsetstack r sl ofs ty =>
+ match sl with
+ | Local =>
Msetstack r (Int.repr (offset_of_index fe (FI_local ofs ty))) ty :: k
- | Incoming ofs ty =>
+ | Incoming =>
k (* should not happen *)
- | Outgoing ofs ty =>
+ | Outgoing =>
Msetstack r (Int.repr (offset_of_index fe (FI_arg ofs ty))) ty :: k
end
| Lop op args res =>
@@ -216,7 +216,9 @@ Open Local Scope string_scope.
Definition transf_function (f: Linear.function) : res Mach.function :=
let fe := make_env (function_bounds f) in
- if zlt Int.max_unsigned fe.(fe_size) then
+ if negb (wt_function f) then
+ Error (msg "Ill-formed Linear code")
+ else if zlt Int.max_unsigned fe.(fe_size) then
Error (msg "Too many spilled variables, stack size exceeded")
else
OK (Mach.mkfunction
diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v
index a73f0aa..1808402 100644
--- a/backend/Stackingproof.v
+++ b/backend/Stackingproof.v
@@ -25,7 +25,7 @@ Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Locations.
-Require LTL.
+Require Import LTL.
Require Import Linear.
Require Import Lineartyping.
Require Import Mach.
@@ -83,17 +83,28 @@ Lemma unfold_transf_function:
(Int.repr fe.(fe_ofs_retaddr)).
Proof.
generalize TRANSF_F. unfold transf_function.
+ destruct (wt_function f); simpl negb.
destruct (zlt Int.max_unsigned (fe_size (make_env (function_bounds f)))).
intros; discriminate.
intros. unfold fe. unfold b. congruence.
+ intros; discriminate.
+Qed.
+
+Lemma transf_function_well_typed:
+ wt_function f = true.
+Proof.
+ generalize TRANSF_F. unfold transf_function.
+ destruct (wt_function f); simpl negb. auto. intros; discriminate.
Qed.
Lemma size_no_overflow: fe.(fe_size) <= Int.max_unsigned.
Proof.
generalize TRANSF_F. unfold transf_function.
+ destruct (wt_function f); simpl negb.
destruct (zlt Int.max_unsigned (fe_size (make_env (function_bounds f)))).
intros; discriminate.
intros. unfold fe. unfold b. omega.
+ intros; discriminate.
Qed.
Remark bound_stack_data_stacksize:
@@ -109,9 +120,8 @@ Definition index_valid (idx: frame_index) :=
match idx with
| FI_link => True
| FI_retaddr => True
- | FI_local x Tint => 0 <= x < b.(bound_int_local)
- | FI_local x Tfloat => 0 <= x < b.(bound_float_local)
- | FI_arg x ty => 0 <= x /\ x + typesize ty <= b.(bound_outgoing)
+ | FI_local x ty => ty <> Tlong /\ 0 <= x /\ x + typesize ty <= b.(bound_local)
+ | FI_arg x ty => ty <> Tlong /\ 0 <= x /\ x + typesize ty <= b.(bound_outgoing)
| FI_saved_int x => 0 <= x < b.(bound_int_callee_save)
| FI_saved_float x => 0 <= x < b.(bound_float_callee_save)
end.
@@ -134,7 +144,7 @@ Definition index_diff (idx1 idx2: frame_index) : Prop :=
| FI_link, FI_link => False
| FI_retaddr, FI_retaddr => False
| FI_local x1 ty1, FI_local x2 ty2 =>
- x1 <> x2 \/ ty1 <> ty2
+ x1 + typesize ty1 <= x2 \/ x2 + typesize ty2 <= x1
| FI_arg x1 ty1, FI_arg x2 ty2 =>
x1 + typesize ty1 <= x2 \/ x2 + typesize ty2 <= x1
| FI_saved_int x1, FI_saved_int x2 => x1 <> x2
@@ -150,8 +160,7 @@ Proof.
Qed.
Ltac AddPosProps :=
- generalize (bound_int_local_pos b); intro;
- generalize (bound_float_local_pos b); intro;
+ generalize (bound_local_pos b); intro;
generalize (bound_int_callee_save_pos b); intro;
generalize (bound_float_callee_save_pos b); intro;
generalize (bound_outgoing_pos b); intro;
@@ -166,6 +175,12 @@ Qed.
Opaque function_bounds.
+Ltac InvIndexValid :=
+ match goal with
+ | [ H: ?ty <> Tlong /\ _ |- _ ] =>
+ destruct H; generalize (typesize_pos ty) (typesize_typesize ty); intros
+ end.
+
Lemma offset_of_index_disj:
forall idx1 idx2,
index_valid idx1 -> index_valid idx2 ->
@@ -177,12 +192,11 @@ Proof.
generalize (frame_env_separated b). intuition. fold fe in H.
AddPosProps.
destruct idx1; destruct idx2;
- try (destruct t); try (destruct t0);
- unfold offset_of_index, type_of_index, AST.typesize;
- simpl in V1; simpl in V2; simpl in DIFF;
- try omega.
- assert (z <> z0). intuition auto. omega.
- assert (z <> z0). intuition auto. omega.
+ simpl in V1; simpl in V2; repeat InvIndexValid; simpl in DIFF;
+ unfold offset_of_index, type_of_index;
+ change (AST.typesize Tint) with 4;
+ change (AST.typesize Tfloat) with 8;
+ omega.
Qed.
Lemma offset_of_index_disj_stack_data_1:
@@ -194,9 +208,11 @@ Proof.
intros idx V.
generalize (frame_env_separated b). intuition. fold fe in H.
AddPosProps.
- destruct idx; try (destruct t);
- unfold offset_of_index, type_of_index, AST.typesize;
- simpl in V;
+ destruct idx;
+ simpl in V; repeat InvIndexValid;
+ unfold offset_of_index, type_of_index;
+ change (AST.typesize Tint) with 4;
+ change (AST.typesize Tfloat) with 8;
omega.
Qed.
@@ -250,18 +266,22 @@ Qed.
Lemma index_local_valid:
forall ofs ty,
- slot_within_bounds f b (Local ofs ty) ->
+ slot_within_bounds b Local ofs ty -> slot_valid f Local ofs ty = true ->
index_valid (FI_local ofs ty).
Proof.
- unfold slot_within_bounds, index_valid. auto.
+ unfold slot_within_bounds, slot_valid, index_valid; intros.
+ InvBooleans.
+ split. destruct ty; congruence. auto.
Qed.
Lemma index_arg_valid:
forall ofs ty,
- slot_within_bounds f b (Outgoing ofs ty) ->
+ slot_within_bounds b Outgoing ofs ty -> slot_valid f Outgoing ofs ty = true ->
index_valid (FI_arg ofs ty).
Proof.
- unfold slot_within_bounds, index_valid. auto.
+ unfold slot_within_bounds, slot_valid, index_valid; intros.
+ InvBooleans.
+ split. destruct ty; congruence. auto.
Qed.
Lemma index_saved_int_valid:
@@ -300,9 +320,10 @@ Proof.
intros idx V.
generalize (frame_env_separated b). intros [A B]. fold fe in A. fold fe in B.
AddPosProps.
- destruct idx; try (destruct t);
- unfold offset_of_index, type_of_index, AST.typesize;
- simpl in V;
+ destruct idx; simpl in V; repeat InvIndexValid;
+ unfold offset_of_index, type_of_index;
+ change (AST.typesize Tint) with 4;
+ change (AST.typesize Tfloat) with 8;
omega.
Qed.
@@ -459,7 +480,9 @@ Proof.
apply offset_of_index_perm; auto.
replace (align_chunk (chunk_of_type (type_of_index idx))) with 4.
apply offset_of_index_aligned; auto.
- destruct (type_of_index idx); auto.
+ assert (type_of_index idx <> Tlong).
+ destruct idx; simpl in *; tauto || congruence.
+ destruct (type_of_index idx); reflexivity || congruence.
exists m'; auto.
Qed.
@@ -539,7 +562,10 @@ Proof.
apply Mem.range_perm_implies with Freeable; auto with mem.
apply offset_of_index_perm; auto.
replace (align_chunk (chunk_of_type (type_of_index idx))) with 4.
- apply offset_of_index_aligned. destruct (type_of_index idx); auto.
+ apply offset_of_index_aligned.
+ assert (type_of_index idx <> Tlong).
+ destruct idx; simpl in *; tauto || congruence.
+ destruct (type_of_index idx); reflexivity || congruence.
intros [v C].
exists v; split; auto. constructor; auto.
Qed.
@@ -570,19 +596,19 @@ Record agree_frame (j: meminj) (ls ls0: locset)
at the corresponding offsets. *)
agree_locals:
forall ofs ty,
- slot_within_bounds f b (Local ofs ty) ->
- index_contains_inj j m' sp' (FI_local ofs ty) (ls (S (Local ofs ty)));
+ slot_within_bounds b Local ofs ty -> slot_valid f Local ofs ty = true ->
+ index_contains_inj j m' sp' (FI_local ofs ty) (ls (S Local ofs ty));
agree_outgoing:
forall ofs ty,
- slot_within_bounds f b (Outgoing ofs ty) ->
- index_contains_inj j m' sp' (FI_arg ofs ty) (ls (S (Outgoing ofs ty)));
+ slot_within_bounds b Outgoing ofs ty -> slot_valid f Outgoing ofs ty = true ->
+ index_contains_inj j m' sp' (FI_arg ofs ty) (ls (S Outgoing ofs ty));
(** Incoming stack slots have the same value as the
corresponding Outgoing stack slots in the caller *)
agree_incoming:
forall ofs ty,
- In (S (Incoming ofs ty)) (loc_parameters f.(Linear.fn_sig)) ->
- ls (S (Incoming ofs ty)) = ls0 (S (Outgoing ofs ty));
+ In (S Incoming ofs ty) (loc_parameters f.(Linear.fn_sig)) ->
+ ls (S Incoming ofs ty) = ls0 (S Outgoing ofs ty);
(** The back link and return address slots of the Mach frame contain
the [parent] and [retaddr] values, respectively. *)
@@ -640,8 +666,8 @@ Hint Resolve agree_unused_reg agree_locals agree_outgoing agree_incoming
Definition agree_callee_save (ls ls0: locset) : Prop :=
forall l,
match l with
- | R r => In r int_callee_save_regs \/ In r float_callee_save_regs
- | S s => True
+ | R r => ~In r destroyed_at_call
+ | S _ _ _ => True
end ->
ls l = ls0 l.
@@ -680,6 +706,18 @@ Proof.
rewrite Locmap.gso; auto. red. auto.
Qed.
+Lemma agree_regs_set_regs:
+ forall j rl vl vl' ls rs,
+ agree_regs j ls rs ->
+ val_list_inject j vl vl' ->
+ agree_regs j (Locmap.setlist (map R rl) vl ls) (set_regs rl vl' rs).
+Proof.
+ induction rl; simpl; intros.
+ auto.
+ inv H0. auto.
+ apply IHrl; auto. apply agree_regs_set_reg; auto.
+Qed.
+
Lemma agree_regs_exten:
forall j ls rs ls' rs',
agree_regs j ls rs ->
@@ -692,52 +730,24 @@ Proof.
rewrite A; rewrite B; auto.
Qed.
-Lemma agree_regs_undef_list:
+Lemma agree_regs_undef_regs:
forall j rl ls rs,
agree_regs j ls rs ->
- agree_regs j (Locmap.undef (List.map R rl) ls) (undef_regs rl rs).
+ agree_regs j (LTL.undef_regs rl ls) (Mach.undef_regs rl rs).
Proof.
induction rl; simpl; intros.
- auto.
- apply IHrl. apply agree_regs_set_reg; auto.
-Qed.
-
-Lemma agree_regs_undef_temps:
- forall j ls rs,
- agree_regs j ls rs ->
- agree_regs j (LTL.undef_temps ls) (undef_temps rs).
-Proof.
- unfold LTL.undef_temps, undef_temps, temporaries.
- intros; apply agree_regs_undef_list; auto.
-Qed.
-
-Lemma agree_regs_undef_setstack:
- forall j ls rs,
- agree_regs j ls rs ->
- agree_regs j (Linear.undef_setstack ls) (undef_setstack rs).
-Proof.
- intros. unfold Linear.undef_setstack, undef_setstack, destroyed_at_move.
- apply agree_regs_undef_list; auto.
-Qed.
-
-Lemma agree_regs_undef_op:
- forall op j ls rs,
- agree_regs j ls rs ->
- agree_regs j (Linear.undef_op op ls) (undef_op (transl_op fe op) rs).
-Proof.
- intros. generalize (agree_regs_undef_temps _ _ _ H); intro A.
-Opaque destroyed_at_move_regs.
- destruct op; auto; simpl; apply agree_regs_undef_setstack; auto.
+ auto.
+ apply agree_regs_set_reg; auto.
Qed.
(** Preservation under assignment of stack slot *)
Lemma agree_regs_set_slot:
- forall j ls rs ss v,
+ forall j ls rs sl ofs ty v,
agree_regs j ls rs ->
- agree_regs j (Locmap.set (S ss) v ls) rs.
+ agree_regs j (Locmap.set (S sl ofs ty) v ls) rs.
Proof.
- intros; red; intros. rewrite Locmap.gso; auto. red. destruct ss; auto.
+ intros; red; intros. rewrite Locmap.gso; auto. red. auto.
Qed.
(** Preservation by increasing memory injections *)
@@ -754,16 +764,10 @@ Qed.
Lemma agree_regs_call_regs:
forall j ls rs,
agree_regs j ls rs ->
- agree_regs j (call_regs ls) (undef_temps rs).
+ agree_regs j (call_regs ls) rs.
Proof.
- intros.
- assert (agree_regs j (LTL.undef_temps ls) (undef_temps rs)).
- apply agree_regs_undef_temps; auto.
- unfold call_regs; intros; red; intros.
- destruct (in_dec Loc.eq (R r) temporaries).
- auto.
- generalize (H0 r). unfold LTL.undef_temps. rewrite Locmap.guo. auto.
- apply Loc.reg_notin; auto.
+ intros.
+ unfold call_regs; intros; red; intros; auto.
Qed.
(** ** Properties of [agree_frame] *)
@@ -782,62 +786,49 @@ Proof.
apply wt_setloc; auto.
Qed.
-Remark temporary_within_bounds:
- forall r, In (R r) temporaries -> mreg_within_bounds b r.
+Lemma agree_frame_set_regs:
+ forall j ls0 m sp m' sp' parent ra rl vl ls,
+ agree_frame j ls ls0 m sp m' sp' parent ra ->
+ (forall r, In r rl -> mreg_within_bounds b r) ->
+ Val.has_type_list vl (map Loc.type (map R rl)) ->
+ agree_frame j (Locmap.setlist (map R rl) vl ls) ls0 m sp m' sp' parent ra.
Proof.
- intros; red. destruct (mreg_type r).
- destruct (zlt (index_int_callee_save r) 0).
- generalize (bound_int_callee_save_pos b). omega.
- exploit int_callee_save_not_destroyed.
- left. eauto with coqlib. apply index_int_callee_save_pos2; auto.
- contradiction.
- destruct (zlt (index_float_callee_save r) 0).
- generalize (bound_float_callee_save_pos b). omega.
- exploit float_callee_save_not_destroyed.
- left. eauto with coqlib. apply index_float_callee_save_pos2; auto.
- contradiction.
+ induction rl; destruct vl; simpl; intros; intuition.
+ apply IHrl; auto.
+ eapply agree_frame_set_reg; eauto.
Qed.
-Lemma agree_frame_undef_locs:
+Lemma agree_frame_undef_regs:
forall j ls0 m sp m' sp' parent ra regs ls,
agree_frame j ls ls0 m sp m' sp' parent ra ->
- incl (List.map R regs) temporaries ->
- agree_frame j (Locmap.undef (List.map R regs) ls) ls0 m sp m' sp' parent ra.
+ (forall r, In r regs -> mreg_within_bounds b r) ->
+ agree_frame j (LTL.undef_regs regs ls) ls0 m sp m' sp' parent ra.
Proof.
induction regs; simpl; intros.
auto.
- apply IHregs; eauto with coqlib.
- apply agree_frame_set_reg; auto.
- apply temporary_within_bounds; eauto with coqlib.
- red; auto.
-Qed.
-
-Lemma agree_frame_undef_temps:
- forall j ls ls0 m sp m' sp' parent ra,
- agree_frame j ls ls0 m sp m' sp' parent ra ->
- agree_frame j (LTL.undef_temps ls) ls0 m sp m' sp' parent ra.
-Proof.
- intros. unfold temporaries. apply agree_frame_undef_locs; auto. apply incl_refl.
+ apply agree_frame_set_reg; auto. red; auto.
Qed.
-Lemma agree_frame_undef_setstack:
- forall j ls ls0 m sp m' sp' parent ra,
- agree_frame j ls ls0 m sp m' sp' parent ra ->
- agree_frame j (Linear.undef_setstack ls) ls0 m sp m' sp' parent ra.
+Lemma caller_save_reg_within_bounds:
+ forall r,
+ In r destroyed_at_call -> mreg_within_bounds b r.
Proof.
- intros. unfold Linear.undef_setstack, destroyed_at_move.
- apply agree_frame_undef_locs; auto.
- red; simpl; tauto.
+ intros. red.
+ destruct (zlt (index_int_callee_save r) 0).
+ destruct (zlt (index_float_callee_save r) 0).
+ generalize (bound_int_callee_save_pos b) (bound_float_callee_save_pos b); omega.
+ exfalso. eapply float_callee_save_not_destroyed; eauto. eapply index_float_callee_save_pos2; eauto.
+ exfalso. eapply int_callee_save_not_destroyed; eauto. eapply index_int_callee_save_pos2; eauto.
Qed.
-Lemma agree_frame_undef_op:
- forall j ls ls0 m sp m' sp' parent ra op,
+Lemma agree_frame_undef_locs:
+ forall j ls0 m sp m' sp' parent ra regs ls,
agree_frame j ls ls0 m sp m' sp' parent ra ->
- agree_frame j (Linear.undef_op op ls) ls0 m sp m' sp' parent ra.
+ incl regs destroyed_at_call ->
+ agree_frame j (LTL.undef_regs regs ls) ls0 m sp m' sp' parent ra.
Proof.
- intros.
- exploit agree_frame_undef_temps; eauto.
- destruct op; simpl; auto; intros; apply agree_frame_undef_setstack; auto.
+ intros. eapply agree_frame_undef_regs; eauto.
+ intros. apply caller_save_reg_within_bounds. auto.
Qed.
(** Preservation by assignment to local slot *)
@@ -845,31 +836,35 @@ Qed.
Lemma agree_frame_set_local:
forall j ls ls0 m sp m' sp' parent retaddr ofs ty v v' m'',
agree_frame j ls ls0 m sp m' sp' parent retaddr ->
- slot_within_bounds f b (Local ofs ty) ->
+ slot_within_bounds b Local ofs ty -> slot_valid f Local ofs ty = true ->
val_inject j v v' ->
Val.has_type v ty ->
Mem.store (chunk_of_type ty) m' sp' (offset_of_index fe (FI_local ofs ty)) v' = Some m'' ->
- agree_frame j (Locmap.set (S (Local ofs ty)) v ls) ls0 m sp m'' sp' parent retaddr.
+ agree_frame j (Locmap.set (S Local ofs ty) v ls) ls0 m sp m'' sp' parent retaddr.
Proof.
intros. inv H.
- change (chunk_of_type ty) with (chunk_of_type (type_of_index (FI_local ofs ty))) in H3.
+ change (chunk_of_type ty) with (chunk_of_type (type_of_index (FI_local ofs ty))) in H4.
constructor; auto; intros.
(* local *)
- unfold Locmap.set. simpl. destruct (Loc.eq (S (Local ofs ty)) (S (Local ofs0 ty0))).
- inv e. eapply gss_index_contains_inj; eauto.
- eapply gso_index_contains_inj. eauto. simpl; auto. eauto with stacking.
- simpl. destruct (zeq ofs ofs0); auto. destruct (typ_eq ty ty0); auto. congruence.
+ unfold Locmap.set.
+ destruct (Loc.eq (S Local ofs ty) (S Local ofs0 ty0)).
+ inv e. eapply gss_index_contains_inj; eauto with stacking.
+ destruct (Loc.diff_dec (S Local ofs ty) (S Local ofs0 ty0)).
+ eapply gso_index_contains_inj. eauto. eauto with stacking. eauto.
+ simpl. simpl in d. intuition.
+ apply index_contains_inj_undef. auto with stacking.
+ red; intros. eapply Mem.perm_store_1; eauto.
(* outgoing *)
rewrite Locmap.gso. eapply gso_index_contains_inj; eauto with stacking.
- simpl; auto. red; auto.
+ red; auto. red; left; congruence.
(* parent *)
- eapply gso_index_contains; eauto. red; auto.
+ eapply gso_index_contains; eauto with stacking. red; auto.
(* retaddr *)
- eapply gso_index_contains; eauto. red; auto.
+ eapply gso_index_contains; eauto with stacking. red; auto.
(* int callee save *)
- eapply gso_index_contains_inj; eauto. simpl; auto.
+ eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
(* float callee save *)
- eapply gso_index_contains_inj; eauto. simpl; auto.
+ eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
(* valid *)
eauto with mem.
(* perm *)
@@ -883,25 +878,26 @@ Qed.
Lemma agree_frame_set_outgoing:
forall j ls ls0 m sp m' sp' parent retaddr ofs ty v v' m'',
agree_frame j ls ls0 m sp m' sp' parent retaddr ->
- slot_within_bounds f b (Outgoing ofs ty) ->
+ slot_within_bounds b Outgoing ofs ty -> slot_valid f Outgoing ofs ty = true ->
val_inject j v v' ->
Val.has_type v ty ->
Mem.store (chunk_of_type ty) m' sp' (offset_of_index fe (FI_arg ofs ty)) v' = Some m'' ->
- agree_frame j (Locmap.set (S (Outgoing ofs ty)) v ls) ls0 m sp m'' sp' parent retaddr.
+ agree_frame j (Locmap.set (S Outgoing ofs ty) v ls) ls0 m sp m'' sp' parent retaddr.
Proof.
intros. inv H.
- change (chunk_of_type ty) with (chunk_of_type (type_of_index (FI_arg ofs ty))) in H3.
+ change (chunk_of_type ty) with (chunk_of_type (type_of_index (FI_arg ofs ty))) in H4.
constructor; auto; intros.
(* local *)
- rewrite Locmap.gso. eapply gso_index_contains_inj; eauto. simpl; auto. red; auto.
+ rewrite Locmap.gso. eapply gso_index_contains_inj; eauto with stacking. red; auto.
+ red; left; congruence.
(* outgoing *)
- unfold Locmap.set. simpl. destruct (Loc.eq (S (Outgoing ofs ty)) (S (Outgoing ofs0 ty0))).
- inv e. eapply gss_index_contains_inj; eauto.
- case_eq (Loc.overlap_aux ty ofs ofs0 || Loc.overlap_aux ty0 ofs0 ofs); intros.
- apply index_contains_inj_undef. auto.
+ unfold Locmap.set. destruct (Loc.eq (S Outgoing ofs ty) (S Outgoing ofs0 ty0)).
+ inv e. eapply gss_index_contains_inj; eauto with stacking.
+ destruct (Loc.diff_dec (S Outgoing ofs ty) (S Outgoing ofs0 ty0)).
+ eapply gso_index_contains_inj; eauto with stacking.
+ red. red in d. intuition.
+ apply index_contains_inj_undef. auto with stacking.
red; intros. eapply Mem.perm_store_1; eauto.
- eapply gso_index_contains_inj; eauto.
- red. eapply Loc.overlap_aux_false_1; eauto.
(* parent *)
eapply gso_index_contains; eauto with stacking. red; auto.
(* retaddr *)
@@ -1038,18 +1034,6 @@ Qed.
(** Preservation at return points (when [ls] is changed but not [ls0]). *)
-Remark mreg_not_within_bounds_callee_save:
- forall r,
- ~mreg_within_bounds b r -> In r int_callee_save_regs \/ In r float_callee_save_regs.
-Proof.
- intro r; unfold mreg_within_bounds.
- destruct (mreg_type r); intro.
- left. apply index_int_callee_save_pos2.
- generalize (bound_int_callee_save_pos b). omega.
- right. apply index_float_callee_save_pos2.
- generalize (bound_float_callee_save_pos b). omega.
-Qed.
-
Lemma agree_frame_return:
forall j ls ls0 m sp m' sp' parent retaddr ls',
agree_frame j ls ls0 m sp m' sp' parent retaddr ->
@@ -1058,7 +1042,7 @@ Lemma agree_frame_return:
agree_frame j ls' ls0 m sp m' sp' parent retaddr.
Proof.
intros. red in H0. inv H; constructor; auto; intros.
- rewrite H0; auto. apply mreg_not_within_bounds_callee_save; auto.
+ rewrite H0; auto. red; intros; elim H. apply caller_save_reg_within_bounds; auto.
rewrite H0; auto.
rewrite H0; auto.
rewrite H0; auto.
@@ -1073,13 +1057,12 @@ Lemma agree_frame_tailcall:
agree_frame j ls ls0' m sp m' sp' parent retaddr.
Proof.
intros. red in H0. inv H; constructor; auto; intros.
- rewrite <- H0; auto. apply mreg_not_within_bounds_callee_save; auto.
- rewrite <- H0; auto.
- rewrite <- H0; auto.
+ rewrite <- H0; auto. red; intros; elim H. apply caller_save_reg_within_bounds; auto.
rewrite <- H0; auto.
+ rewrite <- H0. auto. red; intros. eapply int_callee_save_not_destroyed; eauto.
+ rewrite <- H0. auto. red; intros. eapply float_callee_save_not_destroyed; eauto.
Qed.
-
(** Properties of [agree_callee_save]. *)
Lemma agree_callee_save_return_regs:
@@ -1088,21 +1071,107 @@ Lemma agree_callee_save_return_regs:
Proof.
intros; red; intros.
unfold return_regs. destruct l; auto.
- generalize (int_callee_save_not_destroyed m); intro.
- generalize (float_callee_save_not_destroyed m); intro.
- destruct (In_dec Loc.eq (R m) temporaries). tauto.
- destruct (In_dec Loc.eq (R m) destroyed_at_call). tauto.
- auto.
+ rewrite pred_dec_false; auto.
Qed.
Lemma agree_callee_save_set_result:
- forall ls1 ls2 v sg,
+ forall sg vl ls1 ls2,
agree_callee_save ls1 ls2 ->
- agree_callee_save (Locmap.set (R (loc_result sg)) v ls1) ls2.
+ agree_callee_save (Locmap.setlist (map R (loc_result sg)) vl ls1) ls2.
+Proof.
+ intros sg. generalize (loc_result_caller_save sg).
+ generalize (loc_result sg).
+Opaque destroyed_at_call.
+ induction l; simpl; intros.
+ auto.
+ destruct vl; auto.
+ apply IHl; auto.
+ red; intros. rewrite Locmap.gso. apply H0; auto.
+ destruct l0; simpl; auto.
+Qed.
+
+(** Properties of destroyed registers. *)
+
+Lemma check_mreg_list_incl:
+ forall l1 l2,
+ forallb (fun r => In_dec mreg_eq r l2) l1 = true ->
+ incl l1 l2.
Proof.
- intros; red; intros. rewrite <- H; auto.
- apply Locmap.gso. destruct l; simpl; auto.
- red; intro. subst m. elim (loc_result_not_callee_save _ H0).
+ intros; red; intros.
+ rewrite forallb_forall in H. eapply proj_sumbool_true. eauto.
+Qed.
+
+Remark destroyed_by_op_caller_save:
+ forall op, incl (destroyed_by_op op) destroyed_at_call.
+Proof.
+ destruct op; apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark destroyed_by_load_caller_save:
+ forall chunk addr, incl (destroyed_by_load chunk addr) destroyed_at_call.
+Proof.
+ intros. destruct chunk; apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark destroyed_by_store_caller_save:
+ forall chunk addr, incl (destroyed_by_store chunk addr) destroyed_at_call.
+Proof.
+ intros. destruct chunk; apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark destroyed_by_cond_caller_save:
+ forall cond, incl (destroyed_by_cond cond) destroyed_at_call.
+Proof.
+ destruct cond; apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark destroyed_by_jumptable_caller_save:
+ incl destroyed_by_jumptable destroyed_at_call.
+Proof.
+ apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark destroyed_at_function_entry_caller_save:
+ incl destroyed_at_function_entry destroyed_at_call.
+Proof.
+ apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark destroyed_by_move_at_function_entry:
+ incl (destroyed_by_op Omove) destroyed_at_function_entry.
+Proof.
+ apply check_mreg_list_incl; compute; auto.
+Qed.
+
+Remark temp_for_parent_frame_caller_save:
+ In temp_for_parent_frame destroyed_at_call.
+Proof.
+ Transparent temp_for_parent_frame.
+ Transparent destroyed_at_call.
+ unfold temp_for_parent_frame; simpl; tauto.
+Qed.
+
+Hint Resolve destroyed_by_op_caller_save destroyed_by_load_caller_save
+ destroyed_by_store_caller_save
+ destroyed_by_cond_caller_save destroyed_by_jumptable_caller_save
+ destroyed_at_function_entry_caller_save: stacking.
+
+Remark transl_destroyed_by_op:
+ forall op e, destroyed_by_op (transl_op e op) = destroyed_by_op op.
+Proof.
+ intros; destruct op; reflexivity.
+Qed.
+
+Remark transl_destroyed_by_load:
+ forall chunk addr e, destroyed_by_load chunk (transl_addr e addr) = destroyed_by_load chunk addr.
+Proof.
+ intros; destruct chunk; reflexivity.
+Qed.
+
+Remark transl_destroyed_by_store:
+ forall chunk addr e, destroyed_by_store chunk (transl_addr e addr) = destroyed_by_store chunk addr.
+Proof.
+ intros; destruct chunk; reflexivity.
Qed.
(** * Correctness of saving and restoring of callee-save registers *)
@@ -1157,7 +1226,7 @@ Hypothesis csregs_typ:
forall r, In r csregs -> mreg_type r = ty.
Hypothesis ls_temp_undef:
- forall r, In r destroyed_at_move_regs -> ls (R r) = Vundef.
+ forall r, In r (destroyed_by_op Omove) -> ls (R r) = Vundef.
Hypothesis wt_ls: wt_locset ls.
@@ -1200,18 +1269,18 @@ Proof.
(* a store takes place *)
exploit store_index_succeeds. apply (mkindex_valid a); auto with coqlib.
eauto. instantiate (1 := rs a). intros [m1 ST].
- exploit (IHl k (undef_setstack rs) m1). auto with coqlib. auto.
+ exploit (IHl k (undef_regs (destroyed_by_op Omove) rs) m1). auto with coqlib. auto.
red; eauto with mem.
apply agree_regs_exten with ls rs. auto.
- intros. destruct (In_dec mreg_eq r destroyed_at_move_regs).
+ intros. destruct (In_dec mreg_eq r (destroyed_by_op Omove)).
left. apply ls_temp_undef; auto.
- right; split. auto. unfold undef_setstack, undef_move. apply undef_regs_other; auto.
+ right; split. auto. apply undef_regs_other; auto.
intros [rs' [m' [A [B [C [D [E F]]]]]]].
exists rs'; exists m'.
split. eapply star_left; eauto. econstructor.
rewrite <- (mkindex_typ (number a)).
apply store_stack_succeeds; auto with coqlib.
- traceEq.
+ auto. traceEq.
split; intros.
simpl in H3. destruct (mreg_eq a r). subst r.
assert (index_contains_inj j m1 sp (mkindex (number a)) (ls (R a))).
@@ -1240,9 +1309,33 @@ Qed.
End SAVE_CALLEE_SAVE.
+Remark LTL_undef_regs_same:
+ forall r rl ls, In r rl -> LTL.undef_regs rl ls (R r) = Vundef.
+Proof.
+ induction rl; simpl; intros. contradiction.
+ unfold Locmap.set. destruct (Loc.eq (R a) (R r)). auto.
+ destruct (Loc.diff_dec (R a) (R r)); auto.
+ apply IHrl. intuition congruence.
+Qed.
+
+Remark LTL_undef_regs_others:
+ forall r rl ls, ~In r rl -> LTL.undef_regs rl ls (R r) = ls (R r).
+Proof.
+ induction rl; simpl; intros. auto.
+ rewrite Locmap.gso. apply IHrl. intuition. red. intuition.
+Qed.
+
+Remark LTL_undef_regs_slot:
+ forall sl ofs ty rl ls, LTL.undef_regs rl ls (S sl ofs ty) = ls (S sl ofs ty).
+Proof.
+ induction rl; simpl; intros. auto.
+ rewrite Locmap.gso. apply IHrl. red; auto.
+Qed.
+
Lemma save_callee_save_correct:
- forall j ls rs sp cs fb k m,
- agree_regs j (call_regs ls) rs -> wt_locset (call_regs ls) ->
+ forall j ls0 rs sp cs fb k m,
+ let ls := LTL.undef_regs destroyed_at_function_entry ls0 in
+ agree_regs j ls rs -> wt_locset ls ->
frame_perm_freeable m sp ->
exists rs', exists m',
star step tge
@@ -1250,10 +1343,10 @@ Lemma save_callee_save_correct:
E0 (State cs fb (Vptr sp Int.zero) k rs' m')
/\ (forall r,
In r int_callee_save_regs -> index_int_callee_save r < b.(bound_int_callee_save) ->
- index_contains_inj j m' sp (FI_saved_int (index_int_callee_save r)) (call_regs ls (R r)))
+ index_contains_inj j m' sp (FI_saved_int (index_int_callee_save r)) (ls (R r)))
/\ (forall r,
In r float_callee_save_regs -> index_float_callee_save r < b.(bound_float_callee_save) ->
- index_contains_inj j m' sp (FI_saved_float (index_float_callee_save r)) (call_regs ls (R r)))
+ index_contains_inj j m' sp (FI_saved_float (index_float_callee_save r)) (ls (R r)))
/\ (forall idx v,
index_valid idx ->
match idx with FI_saved_int _ => False | FI_saved_float _ => False | _ => True end ->
@@ -1261,18 +1354,16 @@ Lemma save_callee_save_correct:
index_contains m' sp idx v)
/\ stores_in_frame sp m m'
/\ frame_perm_freeable m' sp
- /\ agree_regs j (call_regs ls) rs'.
+ /\ agree_regs j ls rs'.
Proof.
intros.
- assert (ls_temp_undef: forall r, In r destroyed_at_move_regs -> call_regs ls (R r) = Vundef).
- intros; unfold call_regs. apply pred_dec_true.
-Transparent destroyed_at_move_regs.
- simpl in *; intuition congruence.
+ assert (UNDEF: forall r, In r (destroyed_by_op Omove) -> ls (R r) = Vundef).
+ intros. unfold ls. apply LTL_undef_regs_same. apply destroyed_by_move_at_function_entry; auto.
exploit (save_callee_save_regs_correct
fe_num_int_callee_save
index_int_callee_save
FI_saved_int Tint
- j cs fb sp int_callee_save_regs (call_regs ls)).
+ j cs fb sp int_callee_save_regs ls).
intros. apply index_int_callee_save_inj; auto.
intros. simpl. split. apply Zge_le. apply index_int_callee_save_pos; auto. assumption.
auto.
@@ -1290,7 +1381,7 @@ Transparent destroyed_at_move_regs.
fe_num_float_callee_save
index_float_callee_save
FI_saved_float Tfloat
- j cs fb sp float_callee_save_regs (call_regs ls)).
+ j cs fb sp float_callee_save_regs ls).
intros. apply index_float_callee_save_inj; auto.
intros. simpl. split. apply Zge_le. apply index_float_callee_save_pos; auto. assumption.
simpl; auto.
@@ -1366,10 +1457,12 @@ Qed.
saving of the used callee-save registers). *)
Lemma function_prologue_correct:
- forall j ls ls0 rs m1 m1' m2 sp parent ra cs fb k,
+ forall j ls ls0 ls1 rs rs1 m1 m1' m2 sp parent ra cs fb k,
agree_regs j ls rs ->
agree_callee_save ls ls0 ->
wt_locset ls ->
+ ls1 = LTL.undef_regs destroyed_at_function_entry (LTL.call_regs ls) ->
+ rs1 = undef_regs destroyed_at_function_entry rs ->
Mem.inject j m1 m1' ->
Mem.alloc m1 0 f.(Linear.fn_stacksize) = (m2, sp) ->
Val.has_type parent Tint -> Val.has_type ra Tint ->
@@ -1378,16 +1471,16 @@ Lemma function_prologue_correct:
/\ store_stack m2' (Vptr sp' Int.zero) Tint tf.(fn_link_ofs) parent = Some m3'
/\ store_stack m3' (Vptr sp' Int.zero) Tint tf.(fn_retaddr_ofs) ra = Some m4'
/\ star step tge
- (State cs fb (Vptr sp' Int.zero) (save_callee_save fe k) (undef_temps rs) m4')
+ (State cs fb (Vptr sp' Int.zero) (save_callee_save fe k) rs1 m4')
E0 (State cs fb (Vptr sp' Int.zero) k rs' m5')
- /\ agree_regs j' (call_regs ls) rs'
- /\ agree_frame j' (call_regs ls) ls0 m2 sp m5' sp' parent ra
+ /\ agree_regs j' ls1 rs'
+ /\ agree_frame j' ls1 ls0 m2 sp m5' sp' parent ra
/\ inject_incr j j'
/\ inject_separated j j' m1 m1'
/\ Mem.inject j' m2 m5'
/\ stores_in_frame sp' m2' m5'.
Proof.
- intros until k; intros AGREGS AGCS WTREGS INJ1 ALLOC TYPAR TYRA.
+ intros until k; intros AGREGS AGCS WTREGS LS1 RS1 INJ1 ALLOC TYPAR TYRA.
rewrite unfold_transf_function.
unfold fn_stacksize, fn_link_ofs, fn_retaddr_ofs.
(* Allocation step *)
@@ -1424,9 +1517,12 @@ Proof.
assert (PERM4: frame_perm_freeable m4' sp').
red; intros. eauto with mem.
exploit save_callee_save_correct.
+ instantiate (1 := rs1). instantiate (1 := call_regs ls). instantiate (1 := j').
+ subst rs1. apply agree_regs_undef_regs.
apply agree_regs_call_regs. eapply agree_regs_inject_incr; eauto.
- apply wt_call_regs. auto.
+ apply wt_undef_regs. apply wt_call_regs. auto.
eexact PERM4.
+ rewrite <- LS1.
intros [rs' [m5' [STEPS [ICS [FCS [OTHERS [STORES [PERM5 AGREGS']]]]]]]].
(* stores in frames *)
assert (SIF: stores_in_frame sp' m2' m5').
@@ -1460,15 +1556,20 @@ Proof.
(* agree frame *)
split. constructor; intros.
(* unused regs *)
- unfold call_regs. destruct (in_dec Loc.eq (R r) temporaries).
- elim H. apply temporary_within_bounds; auto.
- apply AGCS. apply mreg_not_within_bounds_callee_save; auto.
+ assert (~In r destroyed_at_call).
+ red; intros; elim H; apply caller_save_reg_within_bounds; auto.
+ rewrite LS1. rewrite LTL_undef_regs_others. unfold call_regs.
+ apply AGCS; auto. red; intros; elim H0.
+ apply destroyed_at_function_entry_caller_save; auto.
(* locals *)
- simpl. apply index_contains_inj_undef; auto.
+ rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
+ apply index_contains_inj_undef; auto with stacking.
(* outgoing *)
- simpl. apply index_contains_inj_undef; auto.
+ rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
+ apply index_contains_inj_undef; auto with stacking.
(* incoming *)
- unfold call_regs. apply AGCS. auto.
+ rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
+ apply AGCS; auto.
(* parent *)
apply OTHERS; auto. red; auto.
eapply gso_index_contains; eauto. red; auto.
@@ -1478,17 +1579,17 @@ Proof.
apply OTHERS; auto. red; auto.
eapply gss_index_contains; eauto. red; auto.
(* int callee save *)
- rewrite <- AGCS. replace (ls (R r)) with (call_regs ls (R r)).
- apply ICS; auto.
- unfold call_regs. apply pred_dec_false.
- red; intros; exploit int_callee_save_not_destroyed; eauto.
- auto.
+ assert (~In r destroyed_at_call).
+ red; intros. eapply int_callee_save_not_destroyed; eauto.
+ exploit ICS; eauto. rewrite LS1. rewrite LTL_undef_regs_others. unfold call_regs.
+ rewrite AGCS; auto.
+ red; intros; elim H1. apply destroyed_at_function_entry_caller_save; auto.
(* float callee save *)
- rewrite <- AGCS. replace (ls (R r)) with (call_regs ls (R r)).
- apply FCS; auto.
- unfold call_regs. apply pred_dec_false.
- red; intros; exploit float_callee_save_not_destroyed; eauto.
- auto.
+ assert (~In r destroyed_at_call).
+ red; intros. eapply float_callee_save_not_destroyed; eauto.
+ exploit FCS; eauto. rewrite LS1. rewrite LTL_undef_regs_others. unfold call_regs.
+ rewrite AGCS; auto.
+ red; intros; elim H1. apply destroyed_at_function_entry_caller_save; auto.
(* inj *)
auto.
(* inj_unique *)
@@ -1502,7 +1603,7 @@ Proof.
(* perms *)
auto.
(* wt *)
- apply wt_call_regs; auto.
+ rewrite LS1. apply wt_undef_regs. apply wt_call_regs; auto.
(* incr *)
split. auto.
(* separated *)
@@ -1625,7 +1726,12 @@ Proof.
Tint
int_callee_save_regs
j cs fb sp' ls0 m'); auto.
- intros. unfold mreg_within_bounds. rewrite (int_callee_save_type r H1). tauto.
+ intros. unfold mreg_within_bounds. split; intros.
+ split; auto. destruct (zlt (index_float_callee_save r) 0).
+ generalize (bound_float_callee_save_pos b). omega.
+ eelim int_float_callee_save_disjoint. eauto.
+ eapply index_float_callee_save_pos2. eauto. auto.
+ destruct H2; auto.
eapply agree_saved_int; eauto.
apply incl_refl.
apply int_callee_save_norepet.
@@ -1638,7 +1744,12 @@ Proof.
Tfloat
float_callee_save_regs
j cs fb sp' ls0 m'); auto.
- intros. unfold mreg_within_bounds. rewrite (float_callee_save_type r H1). tauto.
+ intros. unfold mreg_within_bounds. split; intros.
+ split; auto. destruct (zlt (index_int_callee_save r) 0).
+ generalize (bound_int_callee_save_pos b). omega.
+ eelim int_float_callee_save_disjoint.
+ eapply index_int_callee_save_pos2. eauto. eauto. auto.
+ destruct H2; auto.
eapply agree_saved_float; eauto.
apply incl_refl.
apply float_callee_save_norepet.
@@ -1672,7 +1783,6 @@ Lemma function_epilogue_correct:
E0 (State cs fb (Vptr sp' Int.zero) k rs1 m')
/\ agree_regs j (return_regs ls0 ls) rs1
/\ agree_callee_save (return_regs ls0 ls) ls0
- /\ rs1 IT1 = rs IT1
/\ Mem.inject j m1 m1'.
Proof.
intros.
@@ -1707,16 +1817,12 @@ Proof.
eapply index_contains_load_stack with (idx := FI_retaddr); eauto with stacking.
split. auto.
split. eexact A.
- split. red;intros. unfold return_regs.
+ split. red; intros. unfold return_regs.
generalize (register_classification r) (int_callee_save_not_destroyed r) (float_callee_save_not_destroyed r); intros.
- destruct (in_dec Loc.eq (R r) temporaries).
+ destruct (in_dec mreg_eq r destroyed_at_call).
rewrite C; auto.
- destruct (in_dec Loc.eq (R r) destroyed_at_call).
- rewrite C; auto.
- intuition.
+ apply B. intuition.
split. apply agree_callee_save_return_regs.
- split. apply C. apply int_callee_save_not_destroyed. left; simpl; auto.
- apply float_callee_save_not_destroyed. left; simpl; auto.
auto.
Qed.
@@ -1741,15 +1847,15 @@ Inductive match_stacks (j: meminj) (m m': mem):
match_stacks j m m' nil nil sg bound bound'
| match_stacks_cons: forall f sp ls c cs fb sp' ra c' cs' sg bound bound' trf
(TAIL: is_tail c (Linear.fn_code f))
- (WTF: wt_function f)
+ (WTF: wt_function f = true)
(FINDF: Genv.find_funct_ptr tge fb = Some (Internal trf))
(TRF: transf_function f = OK trf)
(TRC: transl_code (make_env (function_bounds f)) c = c')
(TY_RA: Val.has_type ra Tint)
(FRM: agree_frame f j ls (parent_locset cs) m sp m' sp' (parent_sp cs') (parent_ra cs'))
(ARGS: forall ofs ty,
- In (S (Outgoing ofs ty)) (loc_arguments sg) ->
- slot_within_bounds f (function_bounds f) (Outgoing ofs ty))
+ In (S Outgoing ofs ty) (loc_arguments sg) ->
+ slot_within_bounds (function_bounds f) Outgoing ofs ty)
(STK: match_stacks j m m' cs cs' (Linear.fn_sig f) sp sp')
(BELOW: sp < bound)
(BELOW': sp' < bound'),
@@ -1903,8 +2009,9 @@ Lemma match_stacks_change_sig:
tailcall_possible sg1 ->
match_stacks j m m' cs cs' sg1 bound bound'.
Proof.
- induction 1; intros. econstructor; eauto. econstructor; eauto.
- intros. elim (H0 _ H1).
+ induction 1; intros.
+ econstructor; eauto.
+ econstructor; eauto. intros. elim (H0 _ H1).
Qed.
(** [match_stacks] implies [match_globalenvs], which implies [meminj_preserves_globals]. *)
@@ -2182,18 +2289,6 @@ Proof.
rewrite symbols_preserved. auto.
Qed.
-Hypothesis wt_prog: wt_program prog.
-
-Lemma find_function_well_typed:
- forall ros ls f,
- Linear.find_function ge ros ls = Some f -> wt_fundef f.
-Proof.
- intros until f; destruct ros; simpl; unfold ge.
- intro. eapply Genv.find_funct_prop; eauto.
- destruct (Genv.find_symbol (Genv.globalenv prog) i); try congruence.
- intro. eapply Genv.find_funct_ptr_prop; eauto.
-Qed.
-
(** Preservation of the arguments to an external call. *)
Section EXTERNAL_ARGUMENTS.
@@ -2218,15 +2313,21 @@ Proof.
intros.
assert (loc_argument_acceptable l). apply loc_arguments_acceptable with sg; auto.
destruct l; red in H0.
- exists (rs m0); split. constructor. auto.
- destruct s; try contradiction.
- inv MS.
+ exists (rs r); split. constructor. auto.
+ destruct sl; try contradiction.
+ inv MS.
elim (H4 _ H).
- unfold parent_sp.
+ unfold parent_sp.
+ assert (slot_valid f Outgoing pos ty = true).
+ exploit loc_arguments_acceptable; eauto. intros [A B].
+ unfold slot_valid. unfold proj_sumbool. rewrite zle_true.
+ destruct ty; reflexivity || congruence. omega.
+ assert (slot_within_bounds (function_bounds f) Outgoing pos ty).
+ eauto.
exploit agree_outgoing; eauto. intros [v [A B]].
exists v; split.
constructor.
- eapply index_contains_load_stack with (idx := FI_arg z t); eauto.
+ eapply index_contains_load_stack with (idx := FI_arg pos ty); eauto.
red in AGCS. rewrite AGCS; auto.
Qed.
@@ -2273,37 +2374,34 @@ Hypothesis AGF: agree_frame f j ls ls0 m sp m' sp' parent retaddr.
Lemma transl_annot_param_correct:
forall l,
- loc_acceptable l ->
- match l with S s => slot_within_bounds f b s | _ => True end ->
+ loc_valid f l = true ->
+ match l with S sl ofs ty => slot_within_bounds b sl ofs ty | _ => True end ->
exists v, annot_arg rs m' (Vptr sp' Int.zero) (transl_annot_param fe l) v
/\ val_inject j (ls l) v.
Proof.
- intros. destruct l; red in H.
+ intros. destruct l; simpl in H.
(* reg *)
- exists (rs m0); split. constructor. auto.
+ exists (rs r); split. constructor. auto.
(* stack *)
- destruct s; try contradiction.
+ destruct sl; try discriminate.
exploit agree_locals; eauto. intros [v [A B]]. inv A.
exists v; split. constructor. rewrite Zplus_0_l. auto. auto.
Qed.
Lemma transl_annot_params_correct:
forall ll,
- locs_acceptable ll ->
- (forall s, In (S s) ll -> slot_within_bounds f b s) ->
+ forallb (loc_valid f) ll = true ->
+ (forall sl ofs ty, In (S sl ofs ty) ll -> slot_within_bounds b sl ofs ty) ->
exists vl,
annot_arguments rs m' (Vptr sp' Int.zero) (map (transl_annot_param fe) ll) vl
/\ val_list_inject j (map ls ll) vl.
Proof.
- induction ll; intros.
+ induction ll; simpl; intros.
exists (@nil val); split; constructor.
- exploit (transl_annot_param_correct a).
- apply H; auto with coqlib.
- destruct a; auto with coqlib.
+ InvBooleans.
+ exploit (transl_annot_param_correct a). auto. destruct a; auto.
intros [v1 [A B]].
- exploit IHll.
- red; intros; apply H; auto with coqlib.
- intros; apply H0; auto with coqlib.
+ exploit IHll. auto. auto.
intros [vl [C D]].
exists (v1 :: vl); split; constructor; auto.
Qed.
@@ -2339,7 +2437,7 @@ Inductive match_states: Linear.state -> Mach.state -> Prop :=
(STACKS: match_stacks j m m' cs cs' f.(Linear.fn_sig) sp sp')
(TRANSL: transf_function f = OK tf)
(FIND: Genv.find_funct_ptr tge fb = Some (Internal tf))
- (WTF: wt_function f)
+ (WTF: wt_function f = true)
(AGREGS: agree_regs j ls rs)
(AGFRAME: agree_frame f j ls (parent_locset cs) m sp m' sp' (parent_sp cs') (parent_ra cs'))
(TAIL: is_tail c (Linear.fn_code f)),
@@ -2351,7 +2449,6 @@ Inductive match_states: Linear.state -> Mach.state -> Prop :=
(STACKS: match_stacks j m m' cs cs' (Linear.funsig f) (Mem.nextblock m) (Mem.nextblock m'))
(TRANSL: transf_fundef f = OK tf)
(FIND: Genv.find_funct_ptr tge fb = Some tf)
- (WTF: wt_fundef f)
(WTLS: wt_locset ls)
(AGREGS: agree_regs j ls rs)
(AGLOCS: agree_callee_save ls (parent_locset cs)),
@@ -2372,16 +2469,21 @@ Theorem transf_step_correct:
forall s1' (MS: match_states s1 s1'),
exists s2', plus step tge s1' t s2' /\ match_states s2 s2'.
Proof.
+ assert (USEWTF: forall f i c,
+ wt_function f = true -> is_tail (i :: c) (Linear.fn_code f) ->
+ wt_instr f i = true).
+ intros. unfold wt_function, wt_code in H. rewrite forallb_forall in H.
+ apply H. eapply is_tail_in; eauto.
induction 1; intros;
try inv MS;
try rewrite transl_code_eq;
- try (generalize (WTF _ (is_tail_in TAIL)); intro WTI);
- try (generalize (function_is_within_bounds f WTF _ (is_tail_in TAIL));
+ try (generalize (USEWTF _ _ _ WTF TAIL); intro WTI; simpl in WTI; InvBooleans);
+ try (generalize (function_is_within_bounds f _ (is_tail_in TAIL));
intro BOUND; simpl in BOUND);
unfold transl_instr.
(* Lgetstack *)
- inv WTI. destruct BOUND. unfold undef_getstack; destruct sl.
+ destruct BOUND. unfold destroyed_by_getstack; destruct sl.
(* Lgetstack, local *)
exploit agree_locals; eauto. intros [v [A B]].
econstructor; split.
@@ -2389,26 +2491,33 @@ Proof.
eapply index_contains_load_stack; eauto.
econstructor; eauto with coqlib.
apply agree_regs_set_reg; auto.
- apply agree_frame_set_reg; auto. simpl; rewrite <- H1. eapply agree_wt_ls; eauto.
+ apply agree_frame_set_reg; auto. simpl. rewrite <- H. eapply agree_wt_ls; eauto.
(* Lgetstack, incoming *)
- red in H2. exploit incoming_slot_in_parameters; eauto. intros IN_ARGS.
- inv STACKS. elim (H6 _ IN_ARGS).
+ unfold slot_valid in H0. InvBooleans.
+ exploit incoming_slot_in_parameters; eauto. intros IN_ARGS.
+ inversion STACKS; clear STACKS.
+ elim (H7 _ IN_ARGS).
+ subst bound bound' s cs'.
exploit agree_outgoing. eexact FRM. eapply ARGS; eauto.
+ exploit loc_arguments_acceptable; eauto. intros [A B].
+ unfold slot_valid, proj_sumbool. rewrite zle_true.
+ destruct ty; reflexivity || congruence. omega.
intros [v [A B]].
econstructor; split.
apply plus_one. eapply exec_Mgetparam; eauto.
rewrite (unfold_transf_function _ _ TRANSL). unfold fn_link_ofs.
eapply index_contains_load_stack with (idx := FI_link). eapply TRANSL. eapply agree_link; eauto.
simpl parent_sp.
- change (offset_of_index (make_env (function_bounds f)) (FI_arg z t))
- with (offset_of_index (make_env (function_bounds f0)) (FI_arg z t)).
- eapply index_contains_load_stack with (idx := FI_arg z t). eauto. eauto.
+ change (offset_of_index (make_env (function_bounds f)) (FI_arg ofs ty))
+ with (offset_of_index (make_env (function_bounds f0)) (FI_arg ofs ty)).
+ eapply index_contains_load_stack with (idx := FI_arg ofs ty). eauto. eauto.
exploit agree_incoming; eauto. intros EQ; simpl in EQ.
econstructor; eauto with coqlib. econstructor; eauto.
apply agree_regs_set_reg. apply agree_regs_set_reg. auto. auto. congruence.
eapply agree_frame_set_reg; eauto. eapply agree_frame_set_reg; eauto.
- apply temporary_within_bounds. simpl; auto.
- simpl; auto. simpl; rewrite <- H1. eapply agree_wt_ls; eauto.
+ apply caller_save_reg_within_bounds.
+ apply temp_for_parent_frame_caller_save.
+ subst ty. simpl. eapply agree_wt_ls; eauto.
(* Lgetstack, outgoing *)
exploit agree_outgoing; eauto. intros [v [A B]].
econstructor; split.
@@ -2416,14 +2525,13 @@ Proof.
eapply index_contains_load_stack; eauto.
econstructor; eauto with coqlib.
apply agree_regs_set_reg; auto.
- apply agree_frame_set_reg; auto. simpl; rewrite <- H1; eapply agree_wt_ls; eauto.
+ apply agree_frame_set_reg; auto. simpl; rewrite <- H; eapply agree_wt_ls; eauto.
(* Lsetstack *)
- inv WTI.
set (idx := match sl with
- | Local ofs ty => FI_local ofs ty
- | Incoming ofs ty => FI_link (*dummy*)
- | Outgoing ofs ty => FI_arg ofs ty
+ | Local => FI_local ofs ty
+ | Incoming => FI_link (*dummy*)
+ | Outgoing => FI_arg ofs ty
end).
assert (index_valid f idx).
unfold idx; destruct sl.
@@ -2431,13 +2539,13 @@ Proof.
red; auto.
apply index_arg_valid; auto.
exploit store_index_succeeds; eauto. eapply agree_perm; eauto.
- instantiate (1 := rs0 r). intros [m1' STORE].
+ instantiate (1 := rs0 src). intros [m1' STORE].
econstructor; split.
- apply plus_one. destruct sl; simpl in H3.
- econstructor. eapply store_stack_succeeds with (idx := idx); eauto.
- contradiction.
- econstructor. eapply store_stack_succeeds with (idx := idx); eauto.
- econstructor; eauto with coqlib.
+ apply plus_one. destruct sl; simpl in H0.
+ econstructor. eapply store_stack_succeeds with (idx := idx); eauto. eauto.
+ discriminate.
+ econstructor. eapply store_stack_succeeds with (idx := idx); eauto. auto.
+ econstructor.
eapply Mem.store_outside_inject; eauto.
intros. exploit agree_inj_unique; eauto. intros [EQ1 EQ2]; subst b' delta.
rewrite size_type_chunk in H5.
@@ -2446,20 +2554,30 @@ Proof.
omega.
apply match_stacks_change_mach_mem with m'; auto.
eauto with mem. eauto with mem. intros. rewrite <- H4; eapply Mem.load_store_other; eauto. left; unfold block; omega.
- apply agree_regs_set_slot. apply agree_regs_undef_setstack; auto.
+ eauto. eauto. auto.
+ apply agree_regs_set_slot. apply agree_regs_undef_regs; auto.
destruct sl.
- eapply agree_frame_set_local. eapply agree_frame_undef_setstack; eauto. auto. auto.
- simpl in H1; rewrite H1; eapply agree_wt_ls; eauto. auto.
- simpl in H3; contradiction.
- eapply agree_frame_set_outgoing. apply agree_frame_undef_setstack; eauto. auto. auto.
- simpl in H1; rewrite H1; eapply agree_wt_ls; eauto. auto.
+ eapply agree_frame_set_local. eapply agree_frame_undef_locs; eauto.
+ apply destroyed_by_op_caller_save. auto. auto. apply AGREGS.
+ rewrite H; eapply agree_wt_ls; eauto.
+ assumption.
+ simpl in H0; discriminate.
+ eapply agree_frame_set_outgoing. eapply agree_frame_undef_locs; eauto.
+ apply destroyed_by_op_caller_save. auto. auto. apply AGREGS.
+ rewrite H; eapply agree_wt_ls; eauto.
+ assumption.
+ eauto with coqlib.
(* Lop *)
assert (Val.has_type v (mreg_type res)).
- inv WTI. simpl in H. inv H. rewrite <- H1. eapply agree_wt_ls; eauto.
+ destruct (is_move_operation op args) as [arg|] eqn:?.
+ exploit is_move_operation_correct; eauto. intros [EQ1 EQ2]; subst op args.
+ InvBooleans. simpl in H. inv H. rewrite <- H0. eapply agree_wt_ls; eauto.
replace (mreg_type res) with (snd (type_of_operation op)).
- eapply type_of_operation_sound; eauto.
- rewrite <- H4; auto.
+ eapply type_of_operation_sound; eauto.
+ red; intros. subst op. simpl in Heqo.
+ destruct args; simpl in H. discriminate. destruct args. discriminate. simpl in H. discriminate.
+ destruct (type_of_operation op) as [targs tres]. InvBooleans. auto.
assert (exists v',
eval_operation ge (Vptr sp' Int.zero) (transl_op (make_env (function_bounds f)) op) rs0##args m' = Some v'
/\ val_inject j v v').
@@ -2468,12 +2586,14 @@ Proof.
eapply agree_inj; eauto. eapply agree_reglist; eauto.
destruct H1 as [v' [A B]].
econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. econstructor.
instantiate (1 := v'). rewrite <- A. apply eval_operation_preserved.
- exact symbols_preserved.
+ exact symbols_preserved. eauto.
econstructor; eauto with coqlib.
- apply agree_regs_set_reg; auto. apply agree_regs_undef_op; auto.
- apply agree_frame_set_reg; auto. apply agree_frame_undef_op; auto.
+ apply agree_regs_set_reg; auto.
+ rewrite transl_destroyed_by_op. apply agree_regs_undef_regs; auto.
+ apply agree_frame_set_reg; auto. apply agree_frame_undef_locs; auto.
+ apply destroyed_by_op_caller_save.
(* Lload *)
assert (exists a',
@@ -2482,17 +2602,17 @@ Proof.
eapply eval_addressing_inject; eauto.
eapply match_stacks_preserves_globals; eauto.
eapply agree_inj; eauto. eapply agree_reglist; eauto.
- destruct H1 as [a' [A B]].
+ destruct H2 as [a' [A B]].
exploit Mem.loadv_inject; eauto. intros [v' [C D]].
econstructor; split.
apply plus_one. econstructor.
instantiate (1 := a'). rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved.
- eexact C.
+ eexact C. eauto.
econstructor; eauto with coqlib.
- apply agree_regs_set_reg; auto. apply agree_regs_undef_temps; auto.
- apply agree_frame_set_reg; auto. apply agree_frame_undef_temps; auto.
- simpl. inv WTI. rewrite H6.
- inv B; simpl in H0; try discriminate. eapply Mem.load_type; eauto.
+ apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto.
+ apply agree_frame_set_reg. apply agree_frame_undef_locs; auto.
+ apply destroyed_by_load_caller_save. auto.
+ simpl. rewrite H1. destruct a; simpl in H0; try discriminate. eapply Mem.load_type; eauto.
(* Lstore *)
assert (exists a',
@@ -2506,12 +2626,14 @@ Proof.
econstructor; split.
apply plus_one. econstructor.
instantiate (1 := a'). rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved.
- eexact C.
+ eexact C. eauto.
econstructor; eauto with coqlib.
eapply match_stacks_parallel_stores. eexact MINJ. eexact B. eauto. eauto. auto.
- apply agree_regs_undef_temps; auto.
- apply agree_frame_undef_temps; auto.
+ rewrite transl_destroyed_by_store.
+ apply agree_regs_undef_regs; auto.
+ apply agree_frame_undef_locs; auto.
eapply agree_frame_parallel_stores; eauto.
+ apply destroyed_by_store_caller_save.
(* Lcall *)
exploit find_function_translated; eauto. intros [bf [tf' [A [B C]]]].
@@ -2524,84 +2646,80 @@ Proof.
econstructor; eauto.
econstructor; eauto with coqlib.
simpl; auto.
- intros; red. split.
- generalize (loc_arguments_acceptable _ _ H0). simpl. omega.
+ intros; red.
apply Zle_trans with (size_arguments (Linear.funsig f')); auto.
apply loc_arguments_bounded; auto.
eapply agree_valid_linear; eauto.
eapply agree_valid_mach; eauto.
- eapply find_function_well_typed; eauto.
eapply agree_wt_ls; eauto.
simpl; red; auto.
(* Ltailcall *)
- exploit find_function_translated; eauto. intros [bf [tf' [A [B C]]]].
exploit function_epilogue_correct; eauto.
- intros [rs1 [m1' [P [Q [R [S [T [U [V W]]]]]]]]].
+ intros [rs1 [m1' [P [Q [R [S [T [U V]]]]]]]].
+ exploit find_function_translated; eauto. intros [bf [tf' [A [B C]]]].
econstructor; split.
- eapply plus_right. eexact S. econstructor; eauto.
- replace (find_function_ptr tge ros rs1)
- with (find_function_ptr tge ros rs0). eauto.
- destruct ros; simpl; auto. inv WTI. rewrite V; auto.
- traceEq.
+ eapply plus_right. eexact S. econstructor; eauto. traceEq.
econstructor; eauto.
- inv WTI. apply match_stacks_change_sig with (Linear.fn_sig f); auto.
+ apply match_stacks_change_sig with (Linear.fn_sig f); auto.
apply match_stacks_change_bounds with stk sp'.
apply match_stacks_change_linear_mem with m.
apply match_stacks_change_mach_mem with m'0.
auto.
eauto with mem. intros. eapply Mem.perm_free_1; eauto. left; unfold block; omega.
- intros. rewrite <- H2. eapply Mem.load_free; eauto. left; unfold block; omega.
+ intros. rewrite <- H3. eapply Mem.load_free; eauto. left; unfold block; omega.
eauto with mem. intros. eapply Mem.perm_free_3; eauto.
apply Zlt_le_weak. change (Mem.valid_block m' stk). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_linear; eauto.
apply Zlt_le_weak. change (Mem.valid_block m1' sp'). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_mach; eauto.
- eapply find_function_well_typed; eauto.
+ apply zero_size_arguments_tailcall_possible; auto.
apply wt_return_regs; auto. eapply match_stacks_wt_locset; eauto. eapply agree_wt_ls; eauto.
(* Lbuiltin *)
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject'; eauto.
eapply match_stacks_preserves_globals; eauto.
eapply agree_reglist; eauto.
intros [j' [res' [m1' [A [B [C [D [E [F G]]]]]]]]].
econstructor; split.
apply plus_one. econstructor; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto with coqlib.
+ inversion H; inversion A; subst.
eapply match_stack_change_extcall; eauto.
apply Zlt_le_weak. change (Mem.valid_block m sp0). eapply agree_valid_linear; eauto.
apply Zlt_le_weak. change (Mem.valid_block m'0 sp'). eapply agree_valid_mach; eauto.
- apply agree_regs_set_reg; auto. apply agree_regs_undef_temps; auto. eapply agree_regs_inject_incr; eauto.
- apply agree_frame_set_reg; auto. apply agree_frame_undef_temps; auto.
+ apply agree_regs_set_regs; auto. apply agree_regs_undef_regs; auto. eapply agree_regs_inject_incr; eauto.
+ apply agree_frame_set_regs; auto. apply agree_frame_undef_regs; auto.
eapply agree_frame_inject_incr; eauto.
apply agree_frame_extcall_invariant with m m'0; auto.
- eapply external_call_valid_block; eauto.
- intros. eapply external_call_max_perm; eauto. eapply agree_valid_linear; eauto.
- eapply external_call_valid_block; eauto.
+ eapply external_call_valid_block'; eauto.
+ intros. inv H; eapply external_call_max_perm; eauto. eapply agree_valid_linear; eauto.
+ eapply external_call_valid_block'; eauto.
eapply agree_valid_mach; eauto.
- inv WTI. simpl; rewrite H4. eapply external_call_well_typed; eauto.
+ simpl. rewrite list_map_compose.
+ change (fun x => Loc.type (R x)) with mreg_type.
+ rewrite H0. eapply external_call_well_typed'; eauto.
(* Lannot *)
- inv WTI.
exploit transl_annot_params_correct; eauto.
intros [vargs' [P Q]].
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject'; eauto.
eapply match_stacks_preserves_globals; eauto.
intros [j' [res' [m1' [A [B [C [D [E [F G]]]]]]]]].
econstructor; split.
apply plus_one. econstructor; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto with coqlib.
- eapply match_stack_change_extcall; eauto.
+ inv H; inv A. eapply match_stack_change_extcall; eauto.
apply Zlt_le_weak. change (Mem.valid_block m sp0). eapply agree_valid_linear; eauto.
apply Zlt_le_weak. change (Mem.valid_block m'0 sp'). eapply agree_valid_mach; eauto.
eapply agree_regs_inject_incr; eauto.
eapply agree_frame_inject_incr; eauto.
apply agree_frame_extcall_invariant with m m'0; auto.
- eapply external_call_valid_block; eauto.
- intros. eapply external_call_max_perm; eauto. eapply agree_valid_linear; eauto.
- eapply external_call_valid_block; eauto.
+ eapply external_call_valid_block'; eauto.
+ intros. inv H; eapply external_call_max_perm; eauto. eapply agree_valid_linear; eauto.
+ eapply external_call_valid_block'; eauto.
eapply agree_valid_mach; eauto.
(* Llabel *)
@@ -2620,19 +2738,20 @@ Proof.
econstructor; split.
apply plus_one. eapply exec_Mcond_true; eauto.
eapply eval_condition_inject; eauto. eapply agree_reglist; eauto.
- eapply transl_find_label; eauto.
- econstructor; eauto with coqlib.
- apply agree_regs_undef_temps; auto.
- apply agree_frame_undef_temps; auto.
+ eapply transl_find_label; eauto.
+ econstructor. eauto. eauto. eauto. eauto. eauto.
+ apply agree_regs_undef_regs; auto.
+ apply agree_frame_undef_locs; auto. apply destroyed_by_cond_caller_save.
eapply find_label_tail; eauto.
(* Lcond, false *)
econstructor; split.
apply plus_one. eapply exec_Mcond_false; eauto.
eapply eval_condition_inject; eauto. eapply agree_reglist; eauto.
- econstructor; eauto with coqlib.
- apply agree_regs_undef_temps; auto.
- apply agree_frame_undef_temps; auto.
+ econstructor. eauto. eauto. eauto. eauto. eauto.
+ apply agree_regs_undef_regs; auto.
+ apply agree_frame_undef_locs; auto. apply destroyed_by_cond_caller_save.
+ eauto with coqlib.
(* Ljumptable *)
assert (rs0 arg = Vint n).
@@ -2640,14 +2759,14 @@ Proof.
econstructor; split.
apply plus_one; eapply exec_Mjumptable; eauto.
apply transl_find_label; eauto.
- econstructor; eauto.
- apply agree_regs_undef_temps; auto.
- apply agree_frame_undef_temps; auto.
+ econstructor. eauto. eauto. eauto. eauto. eauto.
+ apply agree_regs_undef_regs; auto.
+ apply agree_frame_undef_locs; auto. apply destroyed_by_jumptable_caller_save.
eapply find_label_tail; eauto.
(* Lreturn *)
exploit function_epilogue_correct; eauto.
- intros [rs1 [m1' [P [Q [R [S [T [U [V W]]]]]]]]].
+ intros [rs1 [m1' [P [Q [R [S [T [U V]]]]]]]].
econstructor; split.
eapply plus_right. eexact S. econstructor; eauto.
traceEq.
@@ -2667,7 +2786,6 @@ Proof.
revert TRANSL. unfold transf_fundef, transf_partial_fundef.
caseEq (transf_function f); simpl; try congruence.
intros tfn TRANSL EQ. inversion EQ; clear EQ; subst tf.
- inversion WTF as [|f' WTFN]. subst f'.
exploit function_prologue_correct; eauto.
eapply match_stacks_type_sp; eauto.
eapply match_stacks_type_retaddr; eauto.
@@ -2689,30 +2807,28 @@ Proof.
intros. eapply stores_in_frame_perm; eauto with mem.
intros. rewrite <- H1. transitivity (Mem.load chunk m2' b ofs). eapply stores_in_frame_contents; eauto.
eapply Mem.load_alloc_unchanged; eauto. red. congruence.
+ eapply transf_function_well_typed; eauto.
auto with coqlib.
(* external function *)
simpl in TRANSL. inversion TRANSL; subst tf.
- inversion WTF. subst ef0.
exploit transl_external_arguments; eauto. intros [vl [ARGS VINJ]].
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject'; eauto.
eapply match_stacks_preserves_globals; eauto.
intros [j' [res' [m1' [A [B [C [D [E [F G]]]]]]]]].
econstructor; split.
apply plus_one. eapply exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto.
apply match_stacks_change_bounds with (Mem.nextblock m) (Mem.nextblock m'0).
- eapply match_stack_change_extcall; eauto. omega. omega.
- exploit external_call_valid_block. eexact H.
+ inv H0; inv A. eapply match_stack_change_extcall; eauto. omega. omega.
+ exploit external_call_valid_block'. eexact H0.
instantiate (1 := (Mem.nextblock m - 1)). red; omega. unfold Mem.valid_block; omega.
- exploit external_call_valid_block. eexact A.
+ exploit external_call_valid_block'. eexact A.
instantiate (1 := (Mem.nextblock m'0 - 1)). red; omega. unfold Mem.valid_block; omega.
- apply wt_setloc; auto. simpl. rewrite loc_result_type.
- change (Val.has_type res (proj_sig_res (ef_sig ef))).
- eapply external_call_well_typed; eauto.
- apply agree_regs_set_reg; auto. apply agree_regs_inject_incr with j; auto.
+ inv H0. apply wt_setlist_result. eapply external_call_well_typed; eauto. auto.
+ apply agree_regs_set_regs; auto. apply agree_regs_inject_incr with j; auto.
apply agree_callee_save_set_result; auto.
(* return *)
@@ -2745,8 +2861,6 @@ Proof.
intros. change (Mem.valid_block m0 b0). eapply Genv.find_funct_ptr_not_fresh; eauto.
intros. change (Mem.valid_block m0 b0). eapply Genv.find_var_info_not_fresh; eauto.
rewrite H3. red; intros. contradiction.
- eapply Genv.find_funct_ptr_prop. eexact wt_prog.
- fold ge0; eauto.
apply wt_init.
unfold Locmap.init. red; intros; auto.
unfold parent_locset. red; auto.
@@ -2757,9 +2871,8 @@ Lemma transf_final_states:
match_states st1 st2 -> Linear.final_state st1 r -> Mach.final_state st2 r.
Proof.
intros. inv H0. inv H. inv STACKS.
- constructor.
- set (rres := loc_result {| sig_args := nil; sig_res := Some Tint |}) in *.
- generalize (AGREGS rres). rewrite H1. intros IJ; inv IJ. auto.
+ generalize (AGREGS r0). rewrite H2. intros A; inv A.
+ econstructor; eauto.
Qed.
Theorem transf_program_correct:
diff --git a/backend/Tunneling.v b/backend/Tunneling.v
index 18414a8..bdc8117 100644
--- a/backend/Tunneling.v
+++ b/backend/Tunneling.v
@@ -64,9 +64,9 @@ Require Import LTL.
Module U := UnionFind.UF(PTree).
-Definition record_goto (uf: U.t) (pc: node) (i: instruction) : U.t :=
- match i with
- | Lnop s => U.union uf pc s
+Definition record_goto (uf: U.t) (pc: node) (b: bblock) : U.t :=
+ match b with
+ | Lbranch s :: _ => U.union uf pc s
| _ => uf
end.
@@ -77,37 +77,23 @@ Definition record_gotos (f: LTL.function) : U.t :=
successor [s] of every instruction by the canonical representative
of its equivalence class in the union-find data structure. *)
-Definition tunnel_instr (uf: U.t) (b: instruction) : instruction :=
- match b with
- | Lnop s =>
- Lnop (U.repr uf s)
- | Lop op args res s =>
- Lop op args res (U.repr uf s)
- | Lload chunk addr args dst s =>
- Lload chunk addr args dst (U.repr uf s)
- | Lstore chunk addr args src s =>
- Lstore chunk addr args src (U.repr uf s)
- | Lcall sig ros args res s =>
- Lcall sig ros args res (U.repr uf s)
- | Ltailcall sig ros args =>
- Ltailcall sig ros args
- | Lbuiltin ef args res s =>
- Lbuiltin ef args res (U.repr uf s)
- | Lcond cond args s1 s2 =>
- Lcond cond args (U.repr uf s1) (U.repr uf s2)
- | Ljumptable arg tbl =>
- Ljumptable arg (List.map (U.repr uf) tbl)
- | Lreturn or =>
- Lreturn or
+Definition tunnel_instr (uf: U.t) (i: instruction) : instruction :=
+ match i with
+ | Lbranch s => Lbranch (U.repr uf s)
+ | Lcond cond args s1 s2 => Lcond cond args (U.repr uf s1) (U.repr uf s2)
+ | Ljumptable arg tbl => Ljumptable arg (List.map (U.repr uf) tbl)
+ | _ => i
end.
+Definition tunnel_block (uf: U.t) (b: bblock) : bblock :=
+ List.map (tunnel_instr uf) b.
+
Definition tunnel_function (f: LTL.function) : LTL.function :=
let uf := record_gotos f in
mkfunction
(fn_sig f)
- (fn_params f)
(fn_stacksize f)
- (PTree.map (fun pc b => tunnel_instr uf b) (fn_code f))
+ (PTree.map1 (tunnel_block uf) (fn_code f))
(U.repr uf (fn_entrypoint f)).
Definition tunnel_fundef (f: LTL.fundef) : LTL.fundef :=
diff --git a/backend/Tunnelingproof.v b/backend/Tunnelingproof.v
index d589260..d02cb2e 100644
--- a/backend/Tunnelingproof.v
+++ b/backend/Tunnelingproof.v
@@ -36,16 +36,16 @@ Definition measure_edge (u: U.t) (pc s: node) (f: node -> nat) : node -> nat :=
else if peq (U.repr u x) pc then (f x + f s + 1)%nat
else f x.
-Definition record_goto' (uf: U.t * (node -> nat)) (pc: node) (i: instruction) : U.t * (node -> nat) :=
- match i with
- | Lnop s => let (u, f) := uf in (U.union u pc s, measure_edge u pc s f)
+Definition record_goto' (uf: U.t * (node -> nat)) (pc: node) (b: bblock) : U.t * (node -> nat) :=
+ match b with
+ | Lbranch s :: b' => let (u, f) := uf in (U.union u pc s, measure_edge u pc s f)
| _ => uf
end.
Definition branch_map_correct (c: code) (uf: U.t * (node -> nat)): Prop :=
forall pc,
match c!pc with
- | Some(Lnop s) =>
+ | Some(Lbranch s :: b) =>
U.repr (fst uf) pc = pc \/ (U.repr (fst uf) pc = U.repr (fst uf) s /\ snd uf s < snd uf pc)%nat
| _ =>
U.repr (fst uf) pc = pc
@@ -65,21 +65,21 @@ Proof.
red; intros; simpl. rewrite PTree.gempty. apply U.repr_empty.
(* inductive case *)
- intros m uf pc i; intros. destruct uf as [u f].
+ intros m uf pc bb; intros. destruct uf as [u f].
assert (PC: U.repr u pc = pc).
generalize (H1 pc). rewrite H. auto.
- assert (record_goto' (u, f) pc i = (u, f)
- \/ exists s, i = Lnop s /\ record_goto' (u, f) pc i = (U.union u pc s, measure_edge u pc s f)).
- unfold record_goto'; simpl. destruct i; auto. right. exists n; auto.
- destruct H2 as [B | [s [EQ B]]].
+ assert (record_goto' (u, f) pc bb = (u, f)
+ \/ exists s, exists bb', bb = Lbranch s :: bb' /\ record_goto' (u, f) pc bb = (U.union u pc s, measure_edge u pc s f)).
+ unfold record_goto'; simpl. destruct bb; auto. destruct i; auto. right. exists s; exists bb; auto.
+ destruct H2 as [B | [s [bb' [EQ B]]]].
(* u and f are unchanged *)
rewrite B.
red. intro pc'. simpl. rewrite PTree.gsspec. destruct (peq pc' pc). subst pc'.
- destruct i; auto.
+ destruct bb; auto. destruct i; auto.
apply H1.
-(* i is Lnop s, u becomes union u pc s, f becomes measure_edge u pc s f *)
+(* b is Lbranch s, u becomes union u pc s, f becomes measure_edge u pc s f *)
rewrite B.
red. intro pc'. simpl. rewrite PTree.gsspec. destruct (peq pc' pc). subst pc'. rewrite EQ.
@@ -91,11 +91,11 @@ Proof.
(* An old instruction *)
assert (U.repr u pc' = pc' -> U.repr (U.union u pc s) pc' = pc').
intro. rewrite <- H2 at 2. apply U.repr_union_1. congruence.
- generalize (H1 pc'). simpl. destruct (m!pc'); auto. destruct i0; auto.
+ generalize (H1 pc'). simpl. destruct (m!pc'); auto. destruct b; auto. destruct i; auto.
intros [P | [P Q]]. left; auto. right.
split. apply U.sameclass_union_2. auto.
unfold measure_edge. destruct (peq (U.repr u s) pc). auto.
- rewrite P. destruct (peq (U.repr u n0) pc). omega. auto.
+ rewrite P. destruct (peq (U.repr u s0) pc). omega. auto.
Qed.
Definition record_gotos' (f: function) :=
@@ -110,7 +110,7 @@ Proof.
induction l; intros; simpl.
auto.
unfold record_goto' at 2. unfold record_goto at 2.
- destruct (snd a); apply IHl.
+ destruct (snd a). apply IHl. destruct i; apply IHl.
Qed.
Definition branch_target (f: function) (pc: node) : node :=
@@ -122,7 +122,7 @@ Definition count_gotos (f: function) (pc: node) : nat :=
Theorem record_gotos_correct:
forall f pc,
match f.(fn_code)!pc with
- | Some(Lnop s) =>
+ | Some(Lbranch s :: b) =>
branch_target f pc = pc \/
(branch_target f pc = branch_target f s /\ count_gotos f s < count_gotos f pc)%nat
| _ => branch_target f pc = pc
@@ -204,15 +204,18 @@ Qed.
the values of locations and the memory states are identical in the
original and transformed codes. *)
+Definition tunneled_block (f: function) (b: bblock) :=
+ tunnel_block (record_gotos f) b.
+
Definition tunneled_code (f: function) :=
- PTree.map (fun pc b => tunnel_instr (record_gotos f) b) (fn_code f).
+ PTree.map1 (tunneled_block f) (fn_code f).
Inductive match_stackframes: stackframe -> stackframe -> Prop :=
| match_stackframes_intro:
- forall res f sp ls0 pc,
+ forall f sp ls0 bb,
match_stackframes
- (Stackframe res f sp ls0 pc)
- (Stackframe res (tunnel_function f) sp ls0 (branch_target f pc)).
+ (Stackframe f sp ls0 bb)
+ (Stackframe (tunnel_function f) sp ls0 (tunneled_block f bb)).
Inductive match_states: state -> state -> Prop :=
| match_states_intro:
@@ -220,6 +223,16 @@ Inductive match_states: state -> state -> Prop :=
list_forall2 match_stackframes s ts ->
match_states (State s f sp pc ls m)
(State ts (tunnel_function f) sp (branch_target f pc) ls m)
+ | match_states_block:
+ forall s f sp bb ls m ts,
+ list_forall2 match_stackframes s ts ->
+ match_states (Block s f sp bb ls m)
+ (Block ts (tunnel_function f) sp (tunneled_block f bb) ls m)
+ | match_states_interm:
+ forall s f sp pc bb ls m ts,
+ list_forall2 match_stackframes s ts ->
+ match_states (Block s f sp (Lbranch pc :: bb) ls m)
+ (State ts (tunnel_function f) sp (branch_target f pc) ls m)
| match_states_call:
forall s f ls m ts,
list_forall2 match_stackframes s ts ->
@@ -238,17 +251,19 @@ Inductive match_states: state -> state -> Prop :=
Definition measure (st: state) : nat :=
match st with
- | State s f sp pc ls m => count_gotos f pc
+ | State s f sp pc ls m => (count_gotos f pc * 2)%nat
+ | Block s f sp (Lbranch pc :: _) ls m => (count_gotos f pc * 2 + 1)%nat
+ | Block s f sp bb ls m => 0%nat
| Callstate s f ls m => 0%nat
| Returnstate s ls m => 0%nat
end.
-Lemma tunnel_function_lookup:
- forall f pc i,
- f.(fn_code)!pc = Some i ->
- (tunnel_function f).(fn_code)!pc = Some (tunnel_instr (record_gotos f) i).
+Lemma match_parent_locset:
+ forall s ts,
+ list_forall2 match_stackframes s ts ->
+ parent_locset ts = parent_locset s.
Proof.
- intros. simpl. rewrite PTree.gmap. rewrite H. auto.
+ induction 1; simpl. auto. inv H; auto.
Qed.
Lemma tunnel_step_correct:
@@ -258,98 +273,113 @@ Lemma tunnel_step_correct:
\/ (measure st2 < measure st1 /\ t = E0 /\ match_states st2 st1')%nat.
Proof.
induction 1; intros; try inv MS.
- (* Lnop *)
- generalize (record_gotos_correct f pc); rewrite H.
- intros [A | [B C]].
- left; econstructor; split.
- eapply exec_Lnop. rewrite A.
- rewrite (tunnel_function_lookup _ _ _ H); simpl; auto.
- econstructor; eauto.
- right. split. simpl. auto. split. auto.
- rewrite B. econstructor; eauto.
+
+ (* entering a block *)
+ assert (DEFAULT: branch_target f pc = pc ->
+ (exists st2' : state,
+ step tge (State ts (tunnel_function f) sp (branch_target f pc) rs m) E0 st2'
+ /\ match_states (Block s f sp bb rs m) st2')).
+ intros. rewrite H0. econstructor; split.
+ econstructor. simpl. rewrite PTree.gmap1. rewrite H. simpl. eauto.
+ econstructor; eauto.
+
+ generalize (record_gotos_correct f pc). rewrite H.
+ destruct bb; auto. destruct i; auto.
+ intros [A | [B C]]. auto.
+ right. split. simpl. omega.
+ split. auto.
+ rewrite B. econstructor; eauto.
+
(* Lop *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_Lop with (v := v); eauto.
- rewrite (tunnel_function_lookup _ _ _ H); simpl; auto.
- rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved.
+ rewrite <- H. apply eval_operation_preserved. exact symbols_preserved.
econstructor; eauto.
(* Lload *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_Lload with (a := a).
- rewrite (tunnel_function_lookup _ _ _ H); simpl; auto.
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
- eauto.
+ rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
+ eauto. eauto.
+ econstructor; eauto.
+ (* Lgetstack *)
+ left; simpl; econstructor; split.
+ econstructor; eauto.
+ econstructor; eauto.
+ (* Lsetstack *)
+ left; simpl; econstructor; split.
+ econstructor; eauto.
econstructor; eauto.
(* Lstore *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_Lstore with (a := a).
- rewrite (tunnel_function_lookup _ _ _ H); simpl; auto.
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
- eauto.
+ rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
+ eauto. eauto.
econstructor; eauto.
(* Lcall *)
- generalize (record_gotos_correct f pc); rewrite H; intro A.
- left; econstructor; split.
- eapply exec_Lcall with (f' := tunnel_fundef f'); eauto.
- rewrite A. rewrite (tunnel_function_lookup _ _ _ H); simpl.
- rewrite sig_preserved. auto.
+ left; simpl; econstructor; split.
+ eapply exec_Lcall with (fd := tunnel_fundef fd); eauto.
apply find_function_translated; auto.
+ rewrite sig_preserved. auto.
econstructor; eauto.
constructor; auto.
constructor; auto.
(* Ltailcall *)
- generalize (record_gotos_correct f pc); rewrite H; intro A.
- left; econstructor; split.
- eapply exec_Ltailcall with (f' := tunnel_fundef f'); eauto.
- rewrite A. rewrite (tunnel_function_lookup _ _ _ H); simpl.
- rewrite sig_preserved. auto.
+ left; simpl; econstructor; split.
+ eapply exec_Ltailcall with (fd := tunnel_fundef fd); eauto.
+ erewrite match_parent_locset; eauto.
apply find_function_translated; auto.
+ apply sig_preserved.
+ erewrite <- match_parent_locset; eauto.
econstructor; eauto.
(* Lbuiltin *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_Lbuiltin; eauto.
- rewrite (tunnel_function_lookup _ _ _ H); simpl; auto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
econstructor; eauto.
- (* cond *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ (* Lannot *)
+ left; simpl; econstructor; split.
+ eapply exec_Lannot; eauto.
+ eapply external_call_symbols_preserved'; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ econstructor; eauto.
+
+ (* Lbranch (preserved) *)
+ left; simpl; econstructor; split.
+ eapply exec_Lbranch; eauto.
+ fold (branch_target f pc). econstructor; eauto.
+ (* Lbranch (eliminated) *)
+ right; split. simpl. omega. split. auto. constructor; auto.
+
+ (* Lcond *)
+ left; simpl; econstructor; split.
eapply exec_Lcond; eauto.
- rewrite (tunnel_function_lookup _ _ _ H); simpl; eauto.
destruct b; econstructor; eauto.
- (* jumptable *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ (* Ljumptable *)
+ left; simpl; econstructor; split.
eapply exec_Ljumptable.
- rewrite (tunnel_function_lookup _ _ _ H); simpl; eauto.
- eauto. rewrite list_nth_z_map. change U.elt with node. rewrite H1. reflexivity.
+ eauto. rewrite list_nth_z_map. change U.elt with node. rewrite H0. reflexivity. eauto.
econstructor; eauto.
- (* return *)
- generalize (record_gotos_correct f pc); rewrite H; intro A; rewrite A.
- left; econstructor; split.
+ (* Lreturn *)
+ left; simpl; econstructor; split.
eapply exec_Lreturn; eauto.
- rewrite (tunnel_function_lookup _ _ _ H); simpl; eauto.
- simpl. constructor; auto.
+ erewrite <- match_parent_locset; eauto.
+ constructor; auto.
(* internal function *)
- simpl. left; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_function_internal; eauto.
simpl. econstructor; eauto.
(* external function *)
- simpl. left; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact varinfo_preserved.
simpl. econstructor; eauto.
(* return *)
inv H3. inv H1.
left; econstructor; split.
eapply exec_return; eauto.
- constructor. auto.
+ constructor; auto.
Qed.
Lemma transf_initial_states:
@@ -357,7 +387,7 @@ Lemma transf_initial_states:
exists st2, initial_state tprog st2 /\ match_states st1 st2.
Proof.
intros. inversion H.
- exists (Callstate nil (tunnel_fundef f) nil m0); split.
+ exists (Callstate nil (tunnel_fundef f) (Locmap.init Vundef) m0); split.
econstructor; eauto.
apply Genv.init_mem_transf; auto.
change (prog_main tprog) with (prog_main prog).
@@ -371,7 +401,7 @@ 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 H6. econstructor; eauto.
Qed.
Theorem transf_program_correct:
diff --git a/backend/Tunnelingtyping.v b/backend/Tunnelingtyping.v
deleted file mode 100644
index dfc36b6..0000000
--- a/backend/Tunnelingtyping.v
+++ /dev/null
@@ -1,103 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *********************************************************************)
-
-(** Type preservation for the Tunneling pass *)
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import LTL.
-Require Import LTLtyping.
-Require Import Tunneling.
-Require Import Tunnelingproof.
-
-(** Tunneling preserves typing. *)
-
-Lemma branch_target_valid_1:
- forall f pc, wt_function f ->
- valid_successor f pc ->
- valid_successor f (branch_target f pc).
-Proof.
- intros.
- assert (forall n p,
- (count_gotos f p < n)%nat ->
- valid_successor f p ->
- valid_successor f (branch_target f p)).
- induction n; intros.
- omegaContradiction.
- elim H2; intros i EQ.
- generalize (record_gotos_correct f p). rewrite EQ.
- destruct i; try congruence.
- intros [A | [B C]]. congruence.
- generalize (wt_instrs _ H _ _ EQ); intro WTI; inv WTI.
- rewrite B. apply IHn. omega. auto.
-
- apply H1 with (Datatypes.S (count_gotos f pc)); auto.
-Qed.
-
-Lemma tunnel_valid_successor:
- forall f pc,
- valid_successor f pc -> valid_successor (tunnel_function f) pc.
-Proof.
- intros. destruct H as [i AT].
- unfold valid_successor; simpl. rewrite PTree.gmap. rewrite AT.
- simpl. exists (tunnel_instr (record_gotos f) i); auto.
-Qed.
-
-Lemma branch_target_valid:
- forall f pc,
- wt_function f ->
- valid_successor f pc ->
- valid_successor (tunnel_function f) (branch_target f pc).
-Proof.
- intros. apply tunnel_valid_successor. apply branch_target_valid_1; auto.
-Qed.
-
-Lemma wt_tunnel_instr:
- forall f i,
- wt_function f ->
- wt_instr f i -> wt_instr (tunnel_function f) (tunnel_instr (record_gotos f) i).
-Proof.
- intros; inv H0; simpl; econstructor; eauto;
- try (eapply branch_target_valid; eauto).
- intros. exploit list_in_map_inv; eauto. intros [x [A B]]. subst lbl.
- eapply branch_target_valid; eauto.
- rewrite list_length_z_map. auto.
-Qed.
-
-Lemma wt_tunnel_function:
- forall f, wt_function f -> wt_function (tunnel_function f).
-Proof.
- intros. inversion H. constructor; simpl; auto.
- intros until instr. rewrite PTree.gmap. unfold option_map.
- caseEq (fn_code f)!pc. intros b0 AT EQ. inv EQ.
- apply wt_tunnel_instr; eauto.
- congruence.
- eapply branch_target_valid; eauto.
-Qed.
-
-Lemma wt_tunnel_fundef:
- forall f, wt_fundef f -> wt_fundef (tunnel_fundef f).
-Proof.
- intros. inversion H; simpl. constructor; auto.
- constructor. apply wt_tunnel_function; auto.
-Qed.
-
-Lemma program_typing_preserved:
- forall (p: LTL.program),
- wt_program p -> wt_program (tunnel_program p).
-Proof.
- intros; red; intros.
- generalize (transform_program_function tunnel_fundef p i f H0).
- intros [f0 [IN TRANSF]].
- subst f. apply wt_tunnel_fundef. eauto.
-Qed.
diff --git a/backend/XTL.ml b/backend/XTL.ml
new file mode 100644
index 0000000..93cab36
--- /dev/null
+++ b/backend/XTL.ml
@@ -0,0 +1,213 @@
+(* *********************************************************************)
+(* *)
+(* 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 XTL intermediate language for register allocation *)
+
+open Datatypes
+open Camlcoq
+open Maps
+open AST
+open Registers
+open Op
+open Locations
+
+type var = V of reg * typ | L of loc
+
+type node = P.t
+
+type instruction =
+ | Xmove of var * var
+ | Xreload of var * var
+ | Xspill of var * var
+ | Xparmove of var list * var list * var * var
+ | Xop of operation * var list * var
+ | Xload of memory_chunk * addressing * var list * var
+ | Xstore of memory_chunk * addressing * var list * var
+ | Xcall of signature * (var, ident) sum * var list * var list
+ | Xtailcall of signature * (var, ident) sum * var list
+ | Xbuiltin of external_function * var list * var list
+ | Xbranch of node
+ | Xcond of condition * var list * node * node
+ | Xjumptable of var * node list
+ | Xreturn of var list
+
+type block = instruction list
+ (* terminated by one of Xbranch, Xcond, Xjumptable, Xtailcall or Xreturn *)
+
+type code = block PTree.t
+ (* mapping node -> block *)
+
+type xfunction = {
+ fn_sig: signature;
+ fn_stacksize: Z.t;
+ fn_code: code;
+ fn_entrypoint: node
+}
+
+(* Type of a variable *)
+
+let typeof = function V(_, ty) -> ty | L l -> Loc.coq_type l
+
+(* Constructors for type [var] *)
+
+let vloc l = L l
+let vlocs ll = List.map vloc ll
+let vmreg mr = L(R mr)
+let vmregs mrl = List.map vmreg mrl
+
+(* Sets of variables *)
+
+module VSet = Set.Make(struct type t = var let compare = compare end)
+
+(*** Generation of fresh registers and fresh register variables *)
+
+let next_temp = ref P.one
+
+let twin_table : (int32, P.t) Hashtbl.t = Hashtbl.create 27
+
+let reset_temps () =
+ next_temp := P.one; Hashtbl.clear twin_table
+
+let new_reg() =
+ let r = !next_temp in next_temp := P.succ !next_temp; r
+
+let new_temp ty = V (new_reg(), ty)
+
+let twin_reg r =
+ let r = P.to_int32 r in
+ try
+ Hashtbl.find twin_table r
+ with Not_found ->
+ let t = new_reg() in Hashtbl.add twin_table r t; t
+
+(*** Successors (for dataflow analysis) *)
+
+let rec successors_block = function
+ | Xbranch s :: _ -> [s]
+ | Xtailcall(sg, vos, args) :: _ -> []
+ | Xcond(cond, args, s1, s2) :: _ -> [s1; s2]
+ | Xjumptable(arg, tbl) :: _ -> tbl
+ | Xreturn _:: _ -> []
+ | instr :: blk -> successors_block blk
+ | [] -> assert false
+
+let successors fn =
+ PTree.map1 successors_block fn.fn_code
+
+(**** Type checking for XTL *)
+
+exception Type_error
+
+exception Type_error_at of node
+
+let set_var_type v ty =
+ if typeof v <> ty then raise Type_error
+
+let rec set_vars_type vl tyl =
+ match vl, tyl with
+ | [], [] -> ()
+ | v1 :: vl, ty1 :: tyl -> set_var_type v1 ty1; set_vars_type vl tyl
+ | _, _ -> raise Type_error
+
+let unify_var_type v1 v2 =
+ if typeof v1 <> typeof v2 then raise Type_error
+
+let type_instr = function
+ | Xmove(src, dst) | Xspill(src, dst) | Xreload(src, dst) ->
+ unify_var_type src dst
+ | Xparmove(srcs, dsts, itmp, ftmp) ->
+ List.iter2 unify_var_type srcs dsts;
+ set_var_type itmp Tint;
+ set_var_type ftmp Tfloat
+ | Xop(op, args, res) ->
+ let (targs, tres) = type_of_operation op in
+ set_vars_type args targs;
+ set_var_type res tres
+ | Xload(chunk, addr, args, dst) ->
+ set_vars_type args (type_of_addressing addr);
+ set_var_type dst (type_of_chunk chunk)
+ | Xstore(chunk, addr, args, src) ->
+ set_vars_type args (type_of_addressing addr);
+ set_var_type src (type_of_chunk chunk)
+ | Xcall(sg, Coq_inl v, args, res) ->
+ set_var_type v Tint
+ | Xcall(sg, Coq_inr id, args, res) ->
+ ()
+ | Xtailcall(sg, Coq_inl v, args) ->
+ set_var_type v Tint
+ | Xtailcall(sg, Coq_inr id, args) ->
+ ()
+ | Xbuiltin(ef, args, res) ->
+ let sg = ef_sig ef in
+ set_vars_type args sg.sig_args;
+ set_vars_type res (Events.proj_sig_res' sg)
+ | Xbranch s ->
+ ()
+ | Xcond(cond, args, s1, s2) ->
+ set_vars_type args (type_of_condition cond)
+ | Xjumptable(arg, tbl) ->
+ set_var_type arg Tint
+ | Xreturn args ->
+ ()
+
+let type_block blk =
+ List.iter type_instr blk
+
+let type_function f =
+ PTree.fold
+ (fun () pc blk ->
+ try
+ type_block blk
+ with Type_error ->
+ raise (Type_error_at pc))
+ f.fn_code ()
+
+(*** A generic framework for transforming extended basic blocks *)
+
+(* Determine instructions that start an extended basic block.
+ These are instructions that have >= 2 predecessors. *)
+
+let basic_blocks_map f = (* return mapping pc -> number of predecessors *)
+ let add_successor map s =
+ PMap.set s (1 + PMap.get s map) map in
+ let add_successors_block map pc blk =
+ List.fold_left add_successor map (successors_block blk) in
+ PTree.fold add_successors_block f.fn_code
+ (PMap.set f.fn_entrypoint 2 (PMap.init 0))
+
+let transform_basic_blocks
+ (transf: node -> block -> 'state -> block * 'state)
+ (top: 'state)
+ f =
+ let bbmap = basic_blocks_map f in
+ let rec transform_block st newcode pc bb =
+ assert (PTree.get pc newcode = None);
+ let (bb', st') = transf pc bb st in
+ (* Record new code after transformation *)
+ let newcode' = PTree.set pc bb' newcode in
+ (* Propagate outgoing state to all successors *)
+ List.fold_left (transform_successor st') newcode' (successors_block bb)
+ and transform_successor st newcode pc =
+ if PMap.get pc bbmap <> 1 then newcode else begin
+ match PTree.get pc f.fn_code with
+ | None -> newcode
+ | Some bb -> transform_block st newcode pc bb
+ end in
+ (* Iterate over all extended basic block heads *)
+ let newcode =
+ PTree.fold
+ (fun newcode pc bb ->
+ if PMap.get pc bbmap >= 2
+ then transform_block top newcode pc bb
+ else newcode)
+ f.fn_code PTree.empty
+ in {f with fn_code = newcode}
diff --git a/backend/XTL.mli b/backend/XTL.mli
new file mode 100644
index 0000000..f2c2715
--- /dev/null
+++ b/backend/XTL.mli
@@ -0,0 +1,100 @@
+(* *********************************************************************)
+(* *)
+(* 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 XTL intermediate language for register allocation *)
+
+open Datatypes
+open Camlcoq
+open Maps
+open AST
+open Registers
+open Machregs
+open Locations
+open Op
+
+type var = V of reg * typ | L of loc
+
+type node = P.t
+
+type instruction =
+ | Xmove of var * var
+ | Xreload of var * var
+ | Xspill of var * var
+ | Xparmove of var list * var list * var * var
+ | Xop of operation * var list * var
+ | Xload of memory_chunk * addressing * var list * var
+ | Xstore of memory_chunk * addressing * var list * var
+ | Xcall of signature * (var, ident) sum * var list * var list
+ | Xtailcall of signature * (var, ident) sum * var list
+ | Xbuiltin of external_function * var list * var list
+ | Xbranch of node
+ | Xcond of condition * var list * node * node
+ | Xjumptable of var * node list
+ | Xreturn of var list
+
+type block = instruction list
+ (* terminated by one of Xbranch, Xcond, Xjumptable, Xtailcall or Xreturn *)
+
+type code = block PTree.t
+ (* mapping node -> block *)
+
+type xfunction = {
+ fn_sig: signature;
+ fn_stacksize: Z.t;
+ fn_code: code;
+ fn_entrypoint: node
+}
+
+(* Type of a variable *)
+
+val typeof: var -> typ
+
+(* Constructors for type [var] *)
+
+val vloc: loc -> var
+val vlocs: loc list -> var list
+val vmreg: mreg -> var
+val vmregs: mreg list -> var list
+
+(* Sets of variables *)
+
+module VSet: Set.S with type elt = var
+
+(* Generation of fresh registers and fresh register variables *)
+
+val reset_temps: unit -> unit
+val new_reg: unit -> reg
+val new_temp: typ -> var
+val twin_reg: reg -> reg
+
+(* Type checking *)
+
+val type_function: xfunction -> unit
+exception Type_error_at of node
+
+(* Successors for dataflow analysis *)
+
+val successors_block: block -> node list
+val successors: xfunction -> node list PTree.t
+
+(* A generic framework for transforming extended basic blocks *)
+
+val transform_basic_blocks:
+ (node -> block -> 'state -> block * 'state) ->
+ 'state ->
+ xfunction -> xfunction
+(* First arg is the transfer function
+ (node, block, state before) -> (transformed block, state after).
+ Second arg is "top" state, to be used as initial state for
+ extended basic block heads. *)
+
+