aboutsummaryrefslogtreecommitdiffstats
path: root/backend
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2015-10-20 13:32:18 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2015-10-20 13:32:18 +0200
commit4d542bc7eafadb16b845cf05d1eb4988eb55ed0f (patch)
tree1961b41815fc6e392cc0bd2beeb0fb504bc160ce /backend
parent7a6bb90048db7a254e959b1e3c308bac5fe6c418 (diff)
downloadcompcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.tar.gz
compcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.zip
Updated PR by removing whitespaces. Bug 17450.
Diffstat (limited to 'backend')
-rw-r--r--backend/Allocation.v92
-rw-r--r--backend/Allocproof.v756
-rw-r--r--backend/Asmexpandaux.ml20
-rw-r--r--backend/Asmgenproof0.v120
-rw-r--r--backend/Bounds.v34
-rw-r--r--backend/CMlexer.mll6
-rw-r--r--backend/CMparser.mly20
-rw-r--r--backend/CMtypecheck.ml2
-rw-r--r--backend/CSE.v14
-rw-r--r--backend/CSEdomain.v2
-rw-r--r--backend/CSEproof.v276
-rw-r--r--backend/CleanupLabels.v6
-rw-r--r--backend/CleanupLabelsproof.v62
-rw-r--r--backend/Cminor.v74
-rw-r--r--backend/CminorSel.v16
-rw-r--r--backend/Constprop.v4
-rw-r--r--backend/Constpropproof.v128
-rw-r--r--backend/Conventions.v18
-rw-r--r--backend/Deadcode.v4
-rw-r--r--backend/Deadcodeproof.v378
-rw-r--r--backend/Debugvar.v8
-rw-r--r--backend/Debugvarproof.v124
-rw-r--r--backend/IRC.ml14
-rw-r--r--backend/Inlining.v22
-rw-r--r--backend/Inliningproof.v422
-rw-r--r--backend/Inliningspec.v172
-rw-r--r--backend/Kildall.v324
-rw-r--r--backend/LTL.v6
-rw-r--r--backend/Linearize.v6
-rw-r--r--backend/Linearizeaux.ml4
-rw-r--r--backend/Linearizeproof.v134
-rw-r--r--backend/Lineartyping.v70
-rw-r--r--backend/Liveness.v10
-rw-r--r--backend/Locations.v112
-rw-r--r--backend/Mach.v10
-rw-r--r--backend/NeedDomain.v320
-rw-r--r--backend/PrintAsm.ml10
-rw-r--r--backend/PrintAsmaux.ml14
-rw-r--r--backend/PrintCminor.ml12
-rw-r--r--backend/PrintRTL.ml2
-rw-r--r--backend/PrintXTL.ml4
-rw-r--r--backend/RTL.v48
-rw-r--r--backend/RTLgen.v12
-rw-r--r--backend/RTLgenaux.ml4
-rw-r--r--backend/RTLgenproof.v232
-rw-r--r--backend/RTLgenspec.v188
-rw-r--r--backend/RTLtyping.v160
-rw-r--r--backend/Regalloc.ml28
-rw-r--r--backend/Registers.v2
-rw-r--r--backend/Renumber.v2
-rw-r--r--backend/Renumberproof.v68
-rw-r--r--backend/SelectDivproof.v196
-rw-r--r--backend/SelectLongproof.v378
-rw-r--r--backend/Selection.v12
-rw-r--r--backend/Selectionproof.v236
-rw-r--r--backend/Splitting.ml6
-rw-r--r--backend/Stacking.v10
-rw-r--r--backend/Stackingproof.v644
-rw-r--r--backend/Tailcall.v4
-rw-r--r--backend/Tailcallproof.v124
-rw-r--r--backend/Tunneling.v2
-rw-r--r--backend/Tunnelingproof.v66
-rw-r--r--backend/Unusedglobproof.v414
-rw-r--r--backend/ValueAnalysis.v628
-rw-r--r--backend/ValueDomain.v778
-rw-r--r--backend/XTL.ml2
66 files changed, 4038 insertions, 4038 deletions
diff --git a/backend/Allocation.v b/backend/Allocation.v
index 196a4075..7534e23f 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -34,7 +34,7 @@ Require Import RTLtyping.
Require Import LTL.
(** The validation algorithm used here is described in
- "Validating register allocation and spilling",
+ "Validating register allocation and spilling",
by Silvain Rideau and Xavier Leroy,
in Compiler Construction (CC 2010), LNCS 6011, Springer, 2010. *)
@@ -157,7 +157,7 @@ Definition classify_operation (op: operation) (args: list reg) : operation_kind
| op, args => operation_other op args
end.
-(** Check RTL instruction [i] against LTL basic block [b].
+(** Check RTL instruction [i] against LTL basic block [b].
On success, return [Some] with a [block_shape] describing the correspondence.
On error, return [None]. *)
@@ -372,7 +372,7 @@ Module OrderedEquation <: OrderedType.
(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).
+ 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.
@@ -380,13 +380,13 @@ Module OrderedEquation <: OrderedType.
Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
Proof.
unfold lt; intros.
- destruct H.
+ destruct H.
destruct H0. left; eapply Plt_trans; eauto.
destruct H0. rewrite <- H0. auto.
- destruct H. rewrite H.
- destruct H0. auto.
+ destruct H. rewrite H.
+ destruct H0. auto.
destruct H0. right; split; auto.
- intuition.
+ intuition.
left; eapply OrderedLoc.lt_trans; eauto.
left; congruence.
left; congruence.
@@ -405,10 +405,10 @@ Module OrderedEquation <: OrderedType.
destruct (OrderedPositive.compare (ereg x) (ereg y)).
- apply LT. red; auto.
- destruct (OrderedLoc.compare (eloc x) (eloc y)).
- + apply LT. red; auto.
+ + 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.
+ * 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.
@@ -416,7 +416,7 @@ Module OrderedEquation <: OrderedType.
Defined.
Definition eq_dec (x y: t) : {x = y} + {x <> y}.
Proof.
- intros. decide equality.
+ intros. decide equality.
apply Loc.eq.
apply peq.
apply IndexedEqKind.eq.
@@ -434,7 +434,7 @@ Module OrderedEquation' <: OrderedType.
(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).
+ 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.
@@ -442,14 +442,14 @@ Module OrderedEquation' <: OrderedType.
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 H.
+ destruct H0. left; eapply OrderedLoc.lt_trans; eauto.
destruct H0. rewrite <- H0. auto.
- destruct H. rewrite H.
- destruct H0. auto.
+ destruct H. rewrite H.
+ destruct H0. auto.
destruct H0. right; split; auto.
- intuition.
- left; eapply Plt_trans; eauto.
+ intuition.
+ left; eapply Plt_trans; eauto.
left; congruence.
left; congruence.
right; split. congruence. eapply OrderedEqKind.lt_trans; eauto.
@@ -467,10 +467,10 @@ Module OrderedEquation' <: OrderedType.
destruct (OrderedLoc.compare (eloc x) (eloc y)).
- apply LT. red; auto.
- destruct (OrderedPositive.compare (ereg x) (ereg y)).
- + apply LT. red; auto.
+ + 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.
+ * 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.
@@ -510,10 +510,10 @@ 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).
+ 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).
+ 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.
@@ -522,10 +522,10 @@ 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).
+ 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).
+ 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.
@@ -585,7 +585,7 @@ Definition subst_reg_kind (r1: reg) (k1: equation_kind) (r2: reg) (k2: equation_
(** [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].
+ partially overlapping [l1].
*)
Definition subst_loc (l1 l2: loc) (e: eqs) : option eqs :=
@@ -784,7 +784,7 @@ Fixpoint can_undef (ml: list mreg) (e: eqs) : bool :=
Fixpoint can_undef_except (l: loc) (ml: list mreg) (e: eqs) : bool :=
match ml with
| nil => true
- | m1 :: ml =>
+ | m1 :: ml =>
(Loc.eq l (R m1) || loc_unconstrained (R m1) e) && can_undef_except l ml e
end.
@@ -967,11 +967,11 @@ Definition transfer_aux (f: RTL.function) (env: regenv)
track_moves env 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'
+ do e1 <- add_equations args args2'
(add_equation (Eq kind_second_word src (R src2')) e);
do e2 <- track_moves env mv2 e1;
assertion (can_undef (destroyed_by_store Mint32 addr) e2);
- do e3 <- add_equations args args1'
+ do e3 <- add_equations args args1'
(add_equation (Eq kind_first_word src (R src1')) e2);
track_moves env mv1 e3
| BScall sg ros args res mv1 ros' mv2 s =>
@@ -1059,22 +1059,22 @@ Module LEq <: SEMILATTICE.
Lemma eq_refl: forall x, eq x x.
Proof.
- intros; destruct x; simpl; auto. red; tauto.
+ 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.
+ unfold eq; intros; destruct x; destruct y; auto.
red in H; red; intros. rewrite H; tauto.
- Qed.
+ 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.
+ red in H; red in H0; red; intros. rewrite H. auto.
Qed.
- Definition beq (x y: t) :=
+ Definition beq (x y: t) :=
match x, y with
| OK a, OK b => EqSet.equal a b
| Error _, Error _ => true
@@ -1083,14 +1083,14 @@ Module LEq <: SEMILATTICE.
Lemma beq_correct: forall x y, beq x y = true -> eq x y.
Proof.
- unfold beq, eq; intros. destruct x; destruct y.
+ unfold beq, eq; intros. destruct x; destruct y.
apply EqSet.equal_2. auto.
discriminate.
discriminate.
auto.
Qed.
- Definition ge (x y: t) :=
+ Definition ge (x y: t) :=
match x, y with
| OK a, OK b => EqSet.Subset b a
| Error _, _ => True
@@ -1099,18 +1099,18 @@ Module LEq <: SEMILATTICE.
Lemma ge_refl: forall x y, eq x y -> ge x y.
Proof.
- unfold eq, ge, EqSet.Equal, EqSet.Subset; intros.
+ 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.
+ 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.
@@ -1126,25 +1126,25 @@ Module LEq <: SEMILATTICE.
| Error _, _ => x
end.
Next Obligation.
- split; intros.
- apply EqSet2.union_1 in H. destruct H; rewrite eqs_same in H.
+ 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 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.
+ 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.
+ unfold lub, ge, EqSet.Subset; intros.
+ destruct x; destruct y; auto.
+ intros; apply EqSet.union_3; auto.
Qed.
End LEq.
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index 57adf102..2bcc038c 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -110,7 +110,7 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
(expand_moves mv1
(Lload Mint32 addr args1' dst1' ::
expand_moves mv2
- (Lload Mint32 addr2 args2' dst2' ::
+ (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 ->
@@ -219,7 +219,7 @@ Proof.
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.
+ 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.
@@ -228,30 +228,30 @@ Proof.
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.
+ 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.
+ 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.
+ split; auto. rewrite <- Q. simpl. unfold expand_moves. rewrite map_app.
rewrite app_ass. simpl. auto.
- intros. exploit IND; eauto. red; intros. elim H0.
+ 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.
+ 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.
@@ -273,9 +273,9 @@ Proof.
(* nop *)
econstructor; eauto.
(* op *)
- destruct (classify_operation o l).
+ destruct (classify_operation o l).
(* move *)
- MonadInv; UseParsingLemmas. econstructor; eauto.
+ MonadInv; UseParsingLemmas. econstructor; eauto.
(* makelong *)
MonadInv; UseParsingLemmas. econstructor; eauto.
(* lowlong *)
@@ -285,7 +285,7 @@ Proof.
(* other ops *)
MonadInv. destruct b0.
MonadInv; UseParsingLemmas.
- destruct i; MonadInv; UseParsingLemmas.
+ destruct i; MonadInv; UseParsingLemmas.
eapply ebs_op; eauto.
inv H0. eapply ebs_op_dead; eauto.
(* load *)
@@ -293,8 +293,8 @@ Proof.
MonadInv; UseParsingLemmas.
destruct i; MonadInv; UseParsingLemmas.
destruct (chunk_eq m Mint64).
- MonadInv; UseParsingLemmas.
- destruct b; MonadInv; UseParsingLemmas. destruct i; MonadInv; UseParsingLemmas.
+ 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.
@@ -310,10 +310,10 @@ Proof.
MonadInv; UseParsingLemmas.
eapply ebs_store; eauto.
(* call *)
- destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
+ destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
(* tailcall *)
destruct b0; MonadInv. destruct i; MonadInv; UseParsingLemmas. econstructor; eauto.
-(* builtin *)
+(* builtin *)
destruct b1; MonadInv. destruct i; MonadInv; UseParsingLemmas.
econstructor; eauto.
(* cond *)
@@ -331,8 +331,8 @@ Lemma matching_instr_block:
exists b, (LTL.fn_code f2)!pc = Some b /\ expand_block_shape bsh i b.
Proof.
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.
+ destruct (LTL.fn_code f2)!pc as [b|].
+ exists b; split; auto. apply pair_instr_block_sound; auto.
discriminate.
Qed.
@@ -367,7 +367,7 @@ Lemma satisf_incr:
forall rs ls (e1 e2: eqs),
satisf rs ls e2 -> EqSet.Subset e1 e2 -> satisf rs ls e1.
Proof.
- unfold satisf; intros. apply H. ESD.fsetdec.
+ unfold satisf; intros. apply H. ESD.fsetdec.
Qed.
Lemma satisf_undef_reg:
@@ -383,14 +383,14 @@ 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.
+ 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.
+ intros. eapply satisf_incr; eauto. unfold add_equation. simpl. ESD.fsetdec.
Qed.
Lemma add_equations_satisf:
@@ -400,7 +400,7 @@ Lemma add_equations_satisf:
Proof.
induction rl; destruct ml; simpl; intros; MonadInv.
auto.
- eapply add_equation_satisf; eauto.
+ eapply add_equation_satisf; eauto.
Qed.
Lemma add_equations_lessdef:
@@ -422,8 +422,8 @@ Lemma add_equations_args_satisf:
satisf rs ls e' -> satisf rs ls e.
Proof.
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.
+ inv H; auto.
+ eapply add_equation_satisf. eapply add_equation_satisf. eauto.
eapply add_equation_satisf. eauto.
eapply add_equation_satisf. eauto.
eapply add_equation_satisf. eauto.
@@ -449,17 +449,17 @@ Lemma add_equations_args_lessdef:
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).
+- 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.
+- destruct H1. constructor; auto.
eapply add_equation_lessdef with (q := Eq Full r1 l1). eapply add_equations_args_satisf; eauto.
-- destruct H1. constructor; auto.
+- destruct H1. constructor; auto.
eapply add_equation_lessdef with (q := Eq Full r1 l1). eapply add_equations_args_satisf; eauto.
-- destruct H1. constructor; auto.
+- destruct H1. constructor; auto.
eapply add_equation_lessdef with (q := Eq Full r1 l1). eapply add_equations_args_satisf; eauto.
- discriminate.
Qed.
@@ -478,7 +478,7 @@ 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.
+ intros. eapply satisf_incr; eauto. unfold remove_equation; simpl. ESD.fsetdec.
Qed.
Lemma remove_equation_res_satisf:
@@ -486,7 +486,7 @@ Lemma remove_equation_res_satisf:
remove_equations_res r oty ll e = Some e' ->
satisf rs ls e -> satisf rs ls e'.
Proof.
- intros. functional inversion H.
+ intros. functional inversion H.
apply remove_equation_satisf. apply remove_equation_satisf; auto.
apply remove_equation_satisf; auto.
Qed.
@@ -498,7 +498,7 @@ Remark select_reg_l_monotone:
Proof.
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]].
+ rewrite Pos.leb_le in *. red in H. destruct H as [A | [A B]].
red in A. zify; omega.
rewrite <- A; auto.
Qed.
@@ -510,7 +510,7 @@ Remark select_reg_h_monotone:
Proof.
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]].
+ rewrite Pos.leb_le in *. red in H. destruct H as [A | [A B]].
red in A. zify; omega.
rewrite A; auto.
Qed.
@@ -520,7 +520,7 @@ Remark select_reg_charact:
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.
+ intros. rewrite H. rewrite ! Pos.leb_refl; auto.
Qed.
Lemma reg_unconstrained_sound:
@@ -530,7 +530,7 @@ Lemma reg_unconstrained_sound:
ereg q <> r.
Proof.
unfold reg_unconstrained; intros. red; intros.
- apply select_reg_charact in H1.
+ 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.
@@ -548,7 +548,7 @@ Lemma reg_unconstrained_satisf:
satisf rs ls e ->
satisf (rs#r <- v) ls e.
Proof.
- red; intros. rewrite PMap.gso. auto. eapply reg_unconstrained_sound; eauto.
+ red; intros. rewrite PMap.gso. auto. eapply reg_unconstrained_sound; eauto.
Qed.
Remark select_loc_l_monotone:
@@ -558,13 +558,13 @@ Remark select_loc_l_monotone:
Proof.
unfold select_loc_l; intros. set (lb := OrderedLoc.diff_low_bound l) in *.
destruct H.
- red in H. subst q2; auto.
+ 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).
+ 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.
eelim OrderedLoc.lt_not_eq; eauto.
eelim OrderedLoc.lt_not_eq. eapply OrderedLoc.lt_trans. eexact l1. eexact H2. red; auto.
@@ -577,13 +577,13 @@ Remark select_loc_h_monotone:
Proof.
unfold select_loc_h; intros. set (lb := OrderedLoc.diff_high_bound l) in *.
destruct H.
- red in H. subst q2; auto.
+ 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).
+ 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.
@@ -594,19 +594,19 @@ Remark select_loc_charact:
select_loc_l l q = false \/ select_loc_h l q = false <-> Loc.diff l (eloc q).
Proof.
unfold select_loc_l, select_loc_h; intros; split; intros.
- apply OrderedLoc.outside_interval_diff.
+ 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.
+ 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; 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.
+ eelim OrderedLoc.lt_not_eq; eauto. apply OrderedLoc.eq_sym; auto.
auto.
Qed.
@@ -616,7 +616,7 @@ Lemma loc_unconstrained_sound:
EqSet.In q e ->
Loc.diff l (eloc q).
Proof.
- unfold loc_unconstrained; intros.
+ 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).
@@ -624,7 +624,7 @@ Proof.
apply EqSet2.mem_between_2 with q; auto.
exact (select_loc_l_monotone l).
exact (select_loc_h_monotone l).
- apply eqs_same. auto.
+ apply eqs_same. auto.
}
rewrite H1 in H; discriminate.
apply select_loc_charact; auto.
@@ -639,12 +639,12 @@ Lemma loc_unconstrained_satisf:
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)).
+ intros; red; intros.
+ destruct (OrderedEquation.eq_dec q (Eq k r l)).
subst q; simpl. unfold l; 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.
+ simpl. ESD.fsetdec.
+ rewrite Locmap.gso. apply H; auto. eapply loc_unconstrained_sound; eauto.
Qed.
Lemma reg_loc_unconstrained_sound:
@@ -653,7 +653,7 @@ Lemma reg_loc_unconstrained_sound:
EqSet.In q e ->
ereg q <> r /\ Loc.diff l (eloc q).
Proof.
- intros. destruct (andb_prop _ _ H).
+ intros. destruct (andb_prop _ _ H).
split. eapply reg_unconstrained_sound; eauto. eapply loc_unconstrained_sound; eauto.
Qed.
@@ -671,7 +671,7 @@ Proof.
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.
+ rewrite Regmap.gso; auto. rewrite Locmap.gso; auto.
Qed.
Lemma parallel_assignment_satisf_2:
@@ -689,20 +689,20 @@ Proof.
{ unfold res'; intros. exploit list_in_map_inv; eauto. intros [mr [A B]]. exists mr; auto. }
functional inversion H.
- (* Two 32-bit halves *)
- subst.
+ subst.
set (e' := remove_equation {| ekind := Low; ereg := res; eloc := l2 |}
(remove_equation {| ekind := High; ereg := res; eloc := l1 |} e)) in *.
rewrite <- H5 in H2. simpl in H2. InvBooleans. simpl.
destruct (OrderedEquation.eq_dec q (Eq Low res l2)).
subst q; simpl. rewrite Regmap.gss.
destruct (ISREG l2) as [r2 EQ]. rewrite <- H5; auto with coqlib. rewrite EQ. rewrite Locmap.gss.
- apply Val.loword_lessdef; auto.
+ 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.
destruct (ISREG l1) as [r1 EQ]. rewrite <- H5; auto with coqlib. rewrite EQ. 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.
+ 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.
@@ -710,11 +710,11 @@ Proof.
subst. rewrite <- H5 in H2. 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)).
+ destruct (OrderedEquation.eq_dec q (Eq Full res l1)).
subst q; simpl. rewrite Regmap.gss.
destruct (ISREG l1) as [r1 EQ]. rewrite <- H5; auto with coqlib. rewrite EQ. rewrite Locmap.gss.
auto.
- assert (EqSet.In q e'). unfold e', remove_equation; simpl. ESD.fsetdec.
+ 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.
@@ -733,7 +733,7 @@ Proof.
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.
+ rewrite select_reg_charact. tauto.
exact (select_reg_l_monotone r1).
exact (select_reg_h_monotone r1).
}
@@ -744,11 +744,11 @@ Proof.
{
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.
+ - 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.
+ left. subst x; auto.
}
set (Q := fun e1 e2 =>
~EqSet.In q e1 ->
@@ -759,11 +759,11 @@ Proof.
- auto.
- simpl. red in H2. rewrite H2 in H4.
rewrite ESF.add_iff. rewrite ESF.remove_iff.
- right. split. apply H3. tauto. tauto.
+ 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.
+ right. split. red; intros. elim n. rewrite IN_ELT. auto. apply H0. auto.
Qed.
Lemma subst_reg_satisf:
@@ -792,7 +792,7 @@ Proof.
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.
+ rewrite select_reg_charact. tauto.
exact (select_reg_l_monotone r1).
exact (select_reg_h_monotone r1).
}
@@ -803,14 +803,14 @@ Proof.
{
intros; apply ESP.fold_rec; unfold P; intros.
- ESD.fsetdec.
- - simpl. red in H1. apply H1 in H3. destruct H3.
+ - 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.
+ 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.
+ auto.
}
set (Q := fun e1 e2 =>
~EqSet.In q e1 \/ ekind q <> k1 ->
@@ -822,8 +822,8 @@ Proof.
- 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.
+ right. split. apply H3. tauto. intuition congruence.
+ apply H3. intuition.
}
destruct (ESP.In_dec q elt).
destruct (IndexedEqKind.eq (ekind q) k1).
@@ -845,17 +845,17 @@ Proof.
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.
+ 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.
+ 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.
+ 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.
+ rewrite Regmap.gso. apply H0; auto. eapply reg_unconstrained_sound; eauto.
Qed.
Lemma subst_reg_kind_satisf_lowlong:
@@ -867,8 +867,8 @@ Lemma subst_reg_kind_satisf_lowlong:
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.
+ 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.
@@ -881,8 +881,8 @@ Lemma subst_reg_kind_satisf_highlong:
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.
+ 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.
@@ -897,7 +897,7 @@ Lemma in_subst_loc:
(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.
+ unfold subst_loc.
set (f := fun (q0 : EqSet2.elt) (opte : option eqs) =>
match opte with
| Some e =>
@@ -921,17 +921,17 @@ Proof.
assert (P elt (EqSet2.fold f elt (Some e0))).
{
apply ESP2.fold_rec; unfold P; intros.
- - ESD2.fsetdec.
- - destruct a as [e2|]; simpl; auto.
+ - 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].
+ 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.
+ 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.
+ left. rewrite e1; auto.
+ right; auto.
}
set (Q := fun e1 (opte: option eqs) =>
match opte with
@@ -941,25 +941,25 @@ Proof.
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.
+ - 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).
+ rewrite SUBST in H; rewrite SUBST in H0; simpl in *.
+ destruct (ESP2.In_dec q elt).
left. apply H; auto.
- right. split; 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.
+ 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.
+ split. apply eqs_same; auto. auto.
Qed.
Lemma loc_type_compat_charact:
@@ -968,12 +968,12 @@ Lemma loc_type_compat_charact:
EqSet.In q e ->
subtype (sel_type (ekind q) (env (ereg q))) (Loc.type l) = true \/ Loc.diff l (eloc q).
Proof.
- unfold loc_type_compat; intros.
+ unfold loc_type_compat; intros.
rewrite EqSet2.for_all_between_iff in H.
destruct (select_loc_l l q) eqn: LL.
destruct (select_loc_h l q) eqn: LH.
- left; apply H; auto. apply eqs_same; auto.
- right. apply select_loc_charact. auto.
+ left; apply H; auto. apply eqs_same; auto.
+ right. apply select_loc_charact. auto.
right. apply select_loc_charact. auto.
intros; subst; auto.
exact (select_loc_l_monotone l).
@@ -990,11 +990,11 @@ Lemma well_typed_move_charact:
| S sl ofs ty => Val.has_type (sel_val k rs#r) ty
end.
Proof.
- unfold well_typed_move; intros.
- destruct l as [mr | sl ofs ty].
+ unfold well_typed_move; intros.
+ destruct l as [mr | sl ofs ty].
auto.
exploit loc_type_compat_charact; eauto. intros [A | A].
- simpl in A. eapply Val.has_subtype; eauto.
+ simpl in A. eapply Val.has_subtype; eauto.
generalize (H1 r). destruct k; simpl; intros.
auto.
destruct (rs#r); exact I.
@@ -1007,7 +1007,7 @@ Remark val_lessdef_normalize:
Val.has_type v ty -> Val.lessdef v v' ->
Val.lessdef v (Val.load_result (chunk_of_type ty) v').
Proof.
- intros. inv H0. rewrite Val.load_result_same; auto. auto.
+ intros. inv H0. rewrite Val.load_result_same; auto. auto.
Qed.
Lemma subst_loc_satisf:
@@ -1024,8 +1024,8 @@ Proof.
destruct q as [k r l]; simpl in *.
exploit well_typed_move_charact; eauto.
destruct l as [mr | sl ofs ty]; intros.
- apply (H2 _ B).
- apply val_lessdef_normalize; auto. apply (H2 _ B).
+ apply (H2 _ B).
+ apply val_lessdef_normalize; auto. apply (H2 _ B).
rewrite Locmap.gso; auto.
Qed.
@@ -1036,7 +1036,7 @@ Proof.
induction ml; simpl; intros.
tauto.
InvBooleans. split.
- apply Loc.diff_sym. eapply loc_unconstrained_sound; eauto.
+ apply Loc.diff_sym. eapply loc_unconstrained_sound; eauto.
eauto.
Qed.
@@ -1045,7 +1045,7 @@ Lemma undef_regs_outside:
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.
+ rewrite Locmap.gso. apply IHml. tauto. apply Loc.diff_sym. tauto.
Qed.
Lemma can_undef_satisf:
@@ -1065,9 +1065,9 @@ Proof.
induction ml; simpl; intros.
tauto.
InvBooleans. split.
- destruct (orb_true_elim _ _ H2).
+ destruct (orb_true_elim _ _ H2).
apply proj_sumbool_true in e0. congruence.
- apply Loc.diff_sym. eapply loc_unconstrained_sound; eauto.
+ apply Loc.diff_sym. eapply loc_unconstrained_sound; eauto.
eapply IHml; eauto.
Qed.
@@ -1086,9 +1086,9 @@ Proof.
destruct q as [k r l]; simpl in *.
exploit well_typed_move_charact; eauto.
destruct l as [mr | sl ofs ty]; intros.
- apply (H3 _ B).
- apply val_lessdef_normalize; auto. apply (H3 _ B).
- rewrite Locmap.gso; auto. rewrite undef_regs_outside. eauto.
+ apply (H3 _ B).
+ apply val_lessdef_normalize; auto. apply (H3 _ B).
+ rewrite Locmap.gso; auto. rewrite undef_regs_outside. eauto.
eapply can_undef_except_sound; eauto. apply Loc.diff_sym; auto.
Qed.
@@ -1100,11 +1100,11 @@ Lemma transfer_use_def_satisf:
(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.
+ 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.
+ intros. eapply parallel_assignment_satisf; eauto. assumption.
+ eapply can_undef_satisf; eauto.
+ eapply add_equations_satisf; eauto.
Qed.
Lemma add_equations_res_lessdef:
@@ -1114,9 +1114,9 @@ Lemma add_equations_res_lessdef:
Val.lessdef_list (encode_long oty rs#r) (map ls ll).
Proof.
intros. functional inversion H.
-- subst. simpl. constructor.
+- subst. simpl. constructor.
eapply add_equation_lessdef with (q := Eq High r l1).
- eapply add_equation_satisf. eauto.
+ eapply add_equation_satisf. eauto.
constructor.
eapply add_equation_lessdef with (q := Eq Low r l2). eauto.
constructor.
@@ -1138,10 +1138,10 @@ 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.
+ intros; red; intros. unfold return_regs. red in H.
destruct l.
- rewrite pred_dec_false; auto.
- destruct sl; auto || congruence.
+ rewrite pred_dec_false; auto.
+ destruct sl; auto || congruence.
Qed.
Lemma no_caller_saves_sound:
@@ -1153,7 +1153,7 @@ 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).
+ 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.
@@ -1175,8 +1175,8 @@ Lemma function_return_satisf:
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.
+- 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.
@@ -1184,20 +1184,20 @@ Proof.
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.
+ 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.
+ 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.
+ 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.
@@ -1210,11 +1210,11 @@ Lemma compat_left_sound:
compat_left r l e = true -> EqSet.In q e -> ereg q = r -> ekind q = Full /\ eloc q = l.
Proof.
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.
+ 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.
intros. subst x2. auto.
exact (select_reg_l_monotone r).
@@ -1227,10 +1227,10 @@ Lemma compat_left2_sound:
(ekind q = High /\ eloc q = l1) \/ (ekind q = Low /\ eloc q = l2).
Proof.
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.
+ 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.
@@ -1259,12 +1259,12 @@ Lemma compat_entry_satisf:
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.
+ 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).
+ 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))).
@@ -1274,17 +1274,17 @@ Proof.
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).
+ 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).
+ 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 Tsingle *)
InvBooleans. simpl in H0. inv H0. simpl.
- red; intros. rewrite Regmap.gsspec. destruct (peq (ereg q) r1).
+ 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 *)
@@ -1295,10 +1295,10 @@ 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.
+ 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.
+ exploit loc_arguments_acceptable; eauto. unfold loc_argument_acceptable.
+ destruct x; auto. destruct sl; tauto.
Qed.
Lemma return_regs_arg_values:
@@ -1306,12 +1306,12 @@ Lemma return_regs_arg_values:
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.
+ intros. apply list_map_exten; intros.
exploit loc_arguments_acceptable; eauto.
- exploit tailcall_is_possible_correct; eauto.
+ exploit tailcall_is_possible_correct; eauto.
unfold loc_argument_acceptable, return_regs.
destruct x; intros.
- rewrite pred_dec_true; auto.
+ rewrite pred_dec_true; auto.
contradiction.
Qed.
@@ -1320,7 +1320,7 @@ Lemma find_function_tailcall:
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.
+ 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.
@@ -1336,9 +1336,9 @@ Lemma loadv_int64_split:
/\ Val.lessdef (Val.loword v) v2.
Proof.
intros. exploit Mem.loadv_int64_split; eauto. intros (v1 & v2 & A & B & C).
- exists v1, v2. split; auto. split; auto.
- inv C; auto. destruct v1, v2; simpl; auto.
- rewrite Int64.hi_ofwords, Int64.lo_ofwords; auto.
+ exists v1, v2. split; auto. split; auto.
+ inv C; auto. destruct v1, v2; simpl; auto.
+ rewrite Int64.hi_ofwords, Int64.lo_ofwords; auto.
Qed.
Lemma add_equations_builtin_arg_satisf:
@@ -1363,12 +1363,12 @@ Proof.
induction 1; simpl; intros e e' arg' AE SAT WT; destruct arg'; MonadInv.
- exploit add_equation_lessdef; eauto. simpl; intros.
exists (ls x0); auto with barg.
-- destruct arg'1; MonadInv. destruct arg'2; MonadInv.
+- destruct arg'1; MonadInv. destruct arg'2; MonadInv.
exploit add_equation_lessdef. eauto. simpl; intros LD1.
exploit add_equation_lessdef. eapply add_equation_satisf. eauto. simpl; intros LD2.
- exists (Val.longofwords (ls x0) (ls x1)); split; auto with barg.
+ exists (Val.longofwords (ls x0) (ls x1)); split; auto with barg.
rewrite <- (val_longofwords_eq rs#x). apply Val.longofwords_lessdef; auto.
- rewrite <- e0; apply WT.
+ rewrite <- e0; apply WT.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
@@ -1377,7 +1377,7 @@ Proof.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
-- exploit IHeval_builtin_arg1; eauto. eapply add_equations_builtin_arg_satisf; eauto.
+- exploit IHeval_builtin_arg1; eauto. eapply add_equations_builtin_arg_satisf; eauto.
intros (v1 & A & B).
exploit IHeval_builtin_arg2; eauto. intros (v2 & C & D).
exists (Val.longofwords v1 v2); split; auto with barg. apply Val.longofwords_lessdef; auto.
@@ -1404,10 +1404,10 @@ Proof.
induction 1; simpl; intros; destruct arg'; MonadInv.
- exists (@nil val); split; constructor.
- exploit IHlist_forall2; eauto. intros (vl' & A & B).
- exploit add_equations_builtin_arg_lessdef; eauto.
+ exploit add_equations_builtin_arg_lessdef; eauto.
eapply add_equations_builtin_args_satisf; eauto. intros (v1' & C & D).
exploit (@eval_builtin_arg_lessdef _ ge ls ls); eauto. intros (v1'' & E & F).
- exists (v1'' :: vl'); split; constructor; auto. eapply Val.lessdef_trans; eauto.
+ exists (v1'' :: vl'); split; constructor; auto. eapply Val.lessdef_trans; eauto.
Qed.
Lemma add_equations_debug_args_satisf:
@@ -1435,7 +1435,7 @@ Proof.
- exists (@nil val); constructor.
- destruct (add_equations_builtin_arg env a1 b e) as [e1|] eqn:A.
+ exploit IHlist_forall2; eauto. intros (vl' & B).
- exploit add_equations_builtin_arg_lessdef; eauto.
+ exploit add_equations_builtin_arg_lessdef; eauto.
eapply add_equations_debug_args_satisf; eauto. intros (v1' & C & D).
exploit (@eval_builtin_arg_lessdef _ ge ls ls); eauto. intros (v1'' & E & F).
exists (v1'' :: vl'); constructor; auto.
@@ -1460,7 +1460,7 @@ Lemma add_equations_builtin_eval:
/\ Val.lessdef vres vres'
/\ Mem.extends m2 m2'.
Proof.
- intros.
+ intros.
assert (DEFAULT: add_equations_builtin_args env args args' e1 = Some e2 ->
satisf rs ls e1 /\
exists vargs' vres' m2',
@@ -1469,19 +1469,19 @@ Proof.
/\ Val.lessdef vres vres'
/\ Mem.extends m2 m2').
{
- intros. split. eapply add_equations_builtin_args_satisf; eauto.
+ intros. split. eapply add_equations_builtin_args_satisf; eauto.
exploit add_equations_builtin_args_lessdef; eauto.
intros (vargs' & A & B).
exploit external_call_mem_extends; eauto.
intros (vres' & m2' & C & D & E & F).
exists vargs', vres', m2'; auto.
}
- destruct ef; auto.
- split. eapply add_equations_debug_args_satisf; eauto.
+ destruct ef; auto.
+ split. eapply add_equations_debug_args_satisf; eauto.
exploit add_equations_debug_args_eval; eauto.
intros (vargs' & A).
simpl in H4; inv H4.
- exists vargs', Vundef, m1'. intuition auto. simpl. constructor.
+ exists vargs', Vundef, m1'. intuition auto. simpl. constructor.
Qed.
Lemma parallel_set_builtin_res_satisf:
@@ -1501,7 +1501,7 @@ Proof.
rename x0 into hi; rename x1 into lo. MonadInv. destruct (mreg_eq hi lo); inv H5.
set (e' := remove_equation {| ekind := High; ereg := x; eloc := R hi |} e0) in *.
set (e'' := remove_equation {| ekind := Low; ereg := x; eloc := R lo |} e') in *.
- simpl in *. red; intros.
+ simpl in *. red; intros.
destruct (OrderedEquation.eq_dec q (Eq Low x (R lo))).
subst q; simpl. rewrite Regmap.gss. rewrite Locmap.gss. apply Val.loword_lessdef; auto.
destruct (OrderedEquation.eq_dec q (Eq High x (R hi))).
@@ -1509,7 +1509,7 @@ Proof.
rewrite Locmap.gss. apply Val.hiword_lessdef; auto.
assert (EqSet.In q e'').
{ unfold e'', e', remove_equation; simpl; ESD.fsetdec. }
- rewrite Regmap.gso. rewrite ! Locmap.gso. auto.
+ rewrite Regmap.gso. rewrite ! Locmap.gso. auto.
eapply loc_unconstrained_sound; eauto.
eapply loc_unconstrained_sound; eauto.
eapply reg_unconstrained_sound; eauto.
@@ -1528,7 +1528,7 @@ Lemma analyze_successors:
Proof.
unfold analyze; intros. exploit DS.fixpoint_allnodes_solution; eauto.
rewrite H2. unfold DS.L.ge. destruct (transfer f env bsh s an#s); intros.
- exists e0; auto.
+ exists e0; auto.
contradiction.
Qed.
@@ -1541,14 +1541,14 @@ Lemma satisf_successors:
satisf rs ls e ->
exists e', transfer f env bsh s an!!s = OK e' /\ satisf rs ls e'.
Proof.
- intros. exploit analyze_successors; eauto. intros [e' [A B]].
+ 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:
+ | transf_function_spec_intro:
forall env an mv k e1 e2,
wt_function f env ->
analyze f env (pair_codes f tf) = Some an ->
@@ -1571,14 +1571,14 @@ Proof.
destruct (type_function f) as [env|] eqn:TY; try discriminate.
destruct (regalloc f); try discriminate.
destruct (check_function f f0 env) as [] eqn:?; inv H.
- unfold check_function in Heqr.
+ unfold check_function in Heqr.
destruct (analyze f env (pair_codes f tf)) as [an|] eqn:?; try discriminate.
- monadInv Heqr.
+ monadInv Heqr.
destruct (check_entrypoints_aux f tf env 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.
+ econstructor; eauto. eapply type_function_correct; eauto. congruence.
Qed.
Lemma invert_code:
@@ -1587,7 +1587,7 @@ Lemma invert_code:
(RTL.fn_code f)!pc = Some i ->
transfer f env (pair_codes f tf) pc opte = OK e ->
exists eafter, exists bsh, exists bb,
- opte = OK eafter /\
+ opte = OK eafter /\
(pair_codes f tf)!pc = Some bsh /\
(LTL.fn_code tf)!pc = Some bb /\
expand_block_shape bsh i bb /\
@@ -1595,11 +1595,11 @@ Lemma invert_code:
wt_instr f env i.
Proof.
intros. destruct opte as [eafter|]; simpl in H1; try discriminate. exists eafter.
- destruct (pair_codes f tf)!pc as [bsh|] eqn:?; try discriminate. exists bsh.
+ 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 env bsh eafter) as [e1|] eqn:?; inv H1.
+ destruct (transfer_aux f env bsh eafter) as [e1|] eqn:?; inv H1.
exists bb. exploit wt_instr_at; eauto.
- tauto.
+ tauto.
Qed.
(** * Semantic preservation *)
@@ -1676,7 +1676,7 @@ Proof.
eapply functions_translated; eauto.
rewrite <- H2 in H. simpl in H. congruence.
(* two symbols *)
- rewrite symbols_preserved. rewrite Heqo.
+ rewrite symbols_preserved. rewrite Heqo.
eapply function_ptr_translated; eauto.
Qed.
@@ -1696,22 +1696,22 @@ Opaque destroyed_by_op.
(* 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 a as [src dst]. unfold expand_moves. simpl.
destruct (track_moves env mv e) as [e1|] eqn:?; MonadInv.
- assert (wf_moves mv). red; intros. apply H0; auto with coqlib.
+ 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.
++ 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.
++ 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.
++ 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.
@@ -1783,17 +1783,17 @@ Lemma match_stackframes_change_sig:
sg'.(sig_res) = sg.(sig_res) ->
match_stackframes s ts sg'.
Proof.
- intros. inv H.
+ 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.
+ unfold proj_sig_res in *. rewrite H0; auto.
+ intros. unfold loc_result in H; rewrite H0 in H; eauto.
Qed.
Ltac UseShape :=
match goal with
| [ WT: wt_function _ _, CODE: (RTL.fn_code _)!_ = Some _, EQ: transfer _ _ _ _ _ = OK _ |- _ ] =>
- destruct (invert_code _ _ _ _ _ _ _ WT CODE EQ) as (eafter & bsh & bb & AFTER & BSH & TCODE & EBS & TR & WTI);
+ destruct (invert_code _ _ _ _ _ _ _ WT CODE EQ) as (eafter & bsh & bb & AFTER & BSH & TCODE & EBS & TR & WTI);
inv EBS; unfold transfer_aux in TR; MonadInv
end.
@@ -1802,7 +1802,7 @@ Remark addressing_not_long:
wt_instr f env (Iload Mint64 addr args dst s) ->
In r args -> r <> dst.
Proof.
- intros.
+ intros.
assert (forall ty, In ty (type_of_addressing addr) -> ty = Tint).
{ intros. destruct addr; simpl in H1; intuition. }
inv H.
@@ -1810,9 +1810,9 @@ Proof.
{ generalize args (type_of_addressing addr) H0 H1 H5.
induction args0; simpl; intros.
contradiction.
- destruct l. discriminate. inv H4.
- destruct H2. subst a. apply H3; auto with coqlib.
- eauto with coqlib.
+ destruct l. discriminate. inv H4.
+ destruct H2. subst a. apply H3; auto with coqlib.
+ eauto with coqlib.
}
red; intros; subst r. rewrite H in H8; discriminate.
Qed.
@@ -1828,108 +1828,108 @@ Proof.
induction 1; intros WT S1' MS; inv MS; try UseShape.
(* nop *)
- exploit exec_moves; eauto. intros [ls1 [X Y]].
+ exploit exec_moves; eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
- eapply star_right. eexact X. econstructor; eauto.
+ 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.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
(* op move *)
- generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'.
- simpl in H0. inv H0.
- exploit (exec_moves mv); eauto. intros [ls1 [X Y]].
+ 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.
+ 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]].
+ exploit satisf_successors; eauto. simpl; eauto. eapply subst_reg_satisf; eauto.
+ intros [enext [U V]].
econstructor; eauto.
(* 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]].
+ 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.
+ 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]].
+ 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]].
+ 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.
+ 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]].
+ 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]].
+ 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.
+ 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]].
+ eapply subst_reg_kind_satisf_highlong. eauto. eauto.
+ intros [enext [U V]].
econstructor; eauto.
(* op regular *)
- generalize (wt_exec_Iop _ _ _ _ _ _ _ _ _ _ _ WTI H0 WTRS). intros WTRS'.
- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ 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 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.
+ 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. eapply star_right. eexact A2. constructor.
eauto. eauto. eauto. traceEq.
- exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
econstructor; eauto.
(* op dead *)
-- exploit exec_moves; eauto. intros [ls1 [X Y]].
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
- eapply star_right. eexact X. econstructor; eauto.
+ 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]].
+ exploit satisf_successors. eauto. eauto. simpl; eauto. eauto.
+ eapply reg_unconstrained_satisf; eauto.
+ intros [enext [U V]].
econstructor; eauto.
- eapply wt_exec_Iop; eauto.
+ eapply wt_exec_Iop; eauto.
(* load regular *)
- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ 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 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.
+ 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.
+ eapply star_right. eexact A2. constructor.
eauto. eauto. eauto. traceEq.
- exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
econstructor; eauto.
(* load pair *)
@@ -1937,49 +1937,49 @@ Proof.
exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2).
set (v2' := if Archi.big_endian then v2 else v1) in *.
set (v1' := if Archi.big_endian then v1 else v2) in *.
- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ 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 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.
+ rewrite Regmap.gss. apply Val.lessdef_trans with v1'; auto.
}
- exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
+ 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.
+ { 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.
+ 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 loc_unconstrained_satisf. eapply can_undef_satisf; eauto.
eapply add_equations_satisf; eauto. assumption.
rewrite Regmap.gss. apply Val.lessdef_trans with v2'; auto.
}
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.
+ 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.
+ 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.
+ constructor.
eauto. eauto. eauto. eauto. eauto. traceEq.
- exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
econstructor; eauto.
(* load first word of a pair *)
@@ -1987,7 +1987,7 @@ Proof.
exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2).
set (v2' := if Archi.big_endian then v2 else v1) in *.
set (v1' := if Archi.big_endian then v1 else v2) in *.
- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ 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]].
@@ -1995,20 +1995,20 @@ Proof.
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.
+ apply Val.lessdef_trans with v1'; auto.
eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto.
}
- exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
+ 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.
+ 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.
+ eexact LOAD1'. instantiate (1 := ls2); auto.
eapply star_right. eexact A3.
- constructor.
+ constructor.
eauto. eauto. eauto. traceEq.
- exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
econstructor; eauto.
(* load second word of a pair *)
@@ -2016,7 +2016,7 @@ Proof.
exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2).
set (v2' := if Archi.big_endian then v2 else v1) in *.
set (v1' := if Archi.big_endian then v1 else v2) in *.
- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ 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.
@@ -2025,51 +2025,51 @@ Proof.
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.
+ apply Val.lessdef_trans with v2'; auto.
eapply can_undef_satisf. eauto. eapply add_equations_satisf; eauto.
}
- exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
+ 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.
+ 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.
+ eexact LOAD2'. instantiate (1 := ls2); auto.
eapply star_right. eexact A3.
- constructor.
+ constructor.
eauto. eauto. eauto. traceEq.
- exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [W Z]].
econstructor; eauto.
(* load dead *)
-- exploit exec_moves; eauto. intros [ls1 [X Y]].
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
- eapply star_right. eexact X. econstructor; eauto.
+ 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]].
+ 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 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 plus_left. econstructor; eauto.
eapply star_trans. eexact X.
- eapply star_two. econstructor. instantiate (1 := a'). rewrite <- F.
+ 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]].
+ eapply can_undef_satisf; eauto. eapply add_equations_satisf; eauto. intros [enext [U V]].
econstructor; eauto.
(* store 2 *)
-- exploit Mem.storev_int64_split; eauto.
+- exploit Mem.storev_int64_split; eauto.
replace (if Archi.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 Archi.big_endian; reflexivity).
@@ -2079,16 +2079,16 @@ Proof.
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.
+ 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 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.
+ 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.
@@ -2098,67 +2098,67 @@ Proof.
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.
+ 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 plus_left. econstructor; eauto.
eapply star_trans. eexact X.
- eapply star_left.
+ 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.
+ 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]].
+ 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.
+ 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 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. congruence.
- intros. exploit (exec_moves mv2). eauto. eauto.
+ intros. exploit (exec_moves mv2). eauto. 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.
+ eapply add_equation_ros_satisf; eauto.
+ eapply add_equations_args_satisf; eauto.
congruence.
- apply wt_regset_assign; auto.
+ apply wt_regset_assign; auto.
intros [ls2 [A2 B2]].
- exists ls2; split.
+ exists ls2; split.
eapply star_right. eexact A2. constructor. traceEq.
- apply satisf_incr with eafter; auto.
+ 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.
- inv WTI. rewrite SIG. rewrite <- H7. apply wt_regset_list; auto.
+ inv WTI. rewrite <- H7. apply wt_regset_list; auto.
+ simpl. red; auto.
+ inv WTI. rewrite SIG. 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.
+ 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 plus_left. econstructor; eauto.
eapply star_right. eexact A1. econstructor; eauto.
- rewrite <- E. apply find_function_tailcall; auto.
+ 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.
@@ -2166,71 +2166,71 @@ Proof.
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 <- H6. apply wt_regset_list; auto.
+ inv WTI. rewrite <- H6. apply wt_regset_list; auto.
apply return_regs_agree_callee_save.
- rewrite SIG. inv WTI. rewrite <- H6. apply wt_regset_list; auto.
+ rewrite SIG. inv WTI. rewrite <- H6. apply wt_regset_list; auto.
(* builtin *)
- exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
- exploit add_equations_builtin_eval; eauto.
+ exploit add_equations_builtin_eval; eauto.
intros (C & vargs' & vres' & m'' & D & E & F & G).
assert (WTRS': wt_regset env (regmap_setres res vres rs)) by (eapply wt_exec_Ibuiltin; eauto).
set (ls2 := Locmap.setres res' vres' (undef_regs (destroyed_by_builtin ef) ls1)).
assert (satisf (regmap_setres res vres rs) ls2 e0).
- { eapply parallel_set_builtin_res_satisf; eauto.
+ { eapply parallel_set_builtin_res_satisf; eauto.
eapply can_undef_satisf; eauto. }
exploit (exec_moves mv2); eauto. intros [ls3 [A3 B3]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
- eapply star_trans. eexact A1.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
eapply star_left. econstructor.
eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
- eapply external_call_symbols_preserved. eauto.
+ eapply external_call_symbols_preserved. eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- instantiate (1 := ls2); auto.
+ instantiate (1 := ls2); auto.
eapply star_right. eexact A3.
- econstructor.
- reflexivity. reflexivity. reflexivity. traceEq.
+ econstructor.
+ reflexivity. reflexivity. reflexivity. traceEq.
exploit satisf_successors; eauto. simpl; eauto.
- intros [enext [U V]].
+ intros [enext [U V]].
econstructor; eauto.
(* cond *)
- exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
econstructor; split.
- eapply plus_left. econstructor; 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.
+ 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]].
+ intros [enext [U V]].
econstructor; eauto.
(* 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.
+ inv H2.
econstructor; split.
- eapply plus_left. econstructor; eauto.
+ eapply plus_left. econstructor; eauto.
eapply star_right. eexact A1.
- econstructor. eauto. eauto. eauto. eauto. traceEq.
+ econstructor. eauto. eauto. eauto. eauto. traceEq.
exploit satisf_successors; eauto.
- instantiate (1 := pc'). simpl. eapply list_nth_z_in; 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]].
+ intros [enext [U V]].
econstructor; eauto.
(* return *)
-- destruct (transf_function_inv _ _ FUN).
+- destruct (transf_function_inv _ _ FUN).
exploit Mem.free_parallel_extends; eauto. rewrite H10. intros [m'' [P Q]].
inv WTI; MonadInv.
+ (* without an argument *)
exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
+ eapply plus_left. econstructor; eauto.
eapply star_right. eexact A1.
econstructor. eauto. eauto. traceEq.
simpl. econstructor; eauto.
@@ -2240,7 +2240,7 @@ Proof.
+ (* with an argument *)
exploit (exec_moves mv); eauto. intros [ls1 [A1 B1]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
+ eapply plus_left. econstructor; eauto.
eapply star_right. eexact A1.
econstructor. eauto. eauto. traceEq.
simpl. econstructor; eauto. rewrite <- H11.
@@ -2250,25 +2250,25 @@ Proof.
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.
- unfold proj_sig_res. rewrite <- H11; rewrite H13. apply WTRS.
+ unfold proj_sig_res. rewrite <- H11; rewrite H13. apply WTRS.
(* internal function *)
- monadInv FUN. simpl in *.
- destruct (transf_function_inv _ _ EQ).
- exploit Mem.alloc_extends; eauto. apply Zle_refl. rewrite H8; apply Zle_refl.
+ destruct (transf_function_inv _ _ EQ).
+ exploit Mem.alloc_extends; eauto. apply Zle_refl. rewrite H8; apply Zle_refl.
intros [m'' [U V]].
assert (WTRS: wt_regset env (init_regs args (fn_params f))).
{ apply wt_init_regs. inv H0. rewrite wt_params. rewrite H9. auto. }
exploit (exec_moves mv). eauto. eauto.
- eapply can_undef_satisf; eauto. eapply compat_entry_satisf; eauto.
+ eapply can_undef_satisf; eauto. eapply compat_entry_satisf; eauto.
rewrite call_regs_param_values. rewrite H9. eexact ARGS.
exact WTRS.
intros [ls1 [A B]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
- eapply star_left. econstructor; eauto.
- eapply star_right. eexact A.
- econstructor; eauto.
+ eapply plus_left. econstructor; eauto.
+ eapply star_left. econstructor; eauto.
+ eapply star_right. eexact A.
+ econstructor; eauto.
eauto. eauto. traceEq.
econstructor; eauto.
@@ -2276,9 +2276,9 @@ Proof.
- exploit external_call_mem_extends; eauto. intros [v' [m'' [F [G [J K]]]]].
simpl in FUN; inv FUN.
econstructor; split.
- apply plus_one. econstructor; eauto.
- eapply external_call_symbols_preserved' with (ge1 := ge).
- econstructor; eauto.
+ apply plus_one. econstructor; eauto.
+ eapply external_call_symbols_preserved' with (ge1 := ge).
+ econstructor; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
econstructor; eauto. simpl.
replace (map
@@ -2290,11 +2290,11 @@ Proof.
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'.
+ 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.
+ simpl. eapply external_call_well_typed; eauto.
(* return *)
- inv STACKS.
@@ -2314,12 +2314,12 @@ Proof.
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.
+ eapply Genv.init_mem_transf_partial; eauto.
+ rewrite symbols_preserved.
rewrite (transform_partial_program_main _ _ TRANSF). auto.
congruence.
constructor; auto.
- constructor. rewrite SIG; rewrite H3; auto.
+ constructor. rewrite SIG; rewrite H3; auto.
rewrite SIG; rewrite H3; simpl; auto.
red; auto.
apply Mem.extends_refl.
@@ -2327,18 +2327,18 @@ Proof.
Qed.
Lemma final_states_simulation:
- forall st1 st2 r,
+ 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 STACKS.
- econstructor. simpl; reflexivity.
- unfold loc_result in RES; rewrite H in RES. simpl in RES. inv RES. inv H3; auto.
+ econstructor. simpl; reflexivity.
+ unfold loc_result in RES; rewrite H in RES. simpl in RES. inv RES. inv H3; auto.
Qed.
Lemma wt_prog: wt_program prog.
Proof.
- red; intros. exploit transform_partial_program_succeeds; eauto.
- intros [tfd TF]. destruct f; simpl in *.
+ red; intros. exploit transform_partial_program_succeeds; eauto.
+ intros [tfd TF]. destruct f; simpl in *.
- monadInv TF. unfold transf_function in EQ.
destruct (type_function f) as [env|] eqn:TF; try discriminate.
econstructor. eapply type_function_correct; eauto.
@@ -2349,16 +2349,16 @@ Theorem transf_program_correct:
forward_simulation (RTL.semantics prog) (LTL.semantics tprog).
Proof.
set (ms := fun s s' => wt_state s /\ match_states s s').
- eapply forward_simulation_plus with (match_states := ms).
+ eapply forward_simulation_plus with (match_states := ms).
- exact public_preserved.
-- intros. exploit initial_states_simulation; eauto. intros [st2 [A B]].
+- intros. exploit initial_states_simulation; eauto. intros [st2 [A B]].
exists st2; split; auto. split; auto.
- apply wt_initial_state with (p := prog); auto. exact wt_prog.
-- intros. destruct H. eapply final_states_simulation; eauto.
-- intros. destruct H0.
+ apply wt_initial_state with (p := prog); auto. exact wt_prog.
+- intros. destruct H. eapply final_states_simulation; eauto.
+- intros. destruct H0.
exploit step_simulation; eauto. intros [s2' [A B]].
exists s2'; split. exact A. split.
- eapply subject_reduction; eauto. eexact wt_prog. eexact H.
+ eapply subject_reduction; eauto. eexact wt_prog. eexact H.
auto.
Qed.
diff --git a/backend/Asmexpandaux.ml b/backend/Asmexpandaux.ml
index 25be9be3..7b8cc8c2 100644
--- a/backend/Asmexpandaux.ml
+++ b/backend/Asmexpandaux.ml
@@ -11,13 +11,13 @@
(* *)
(* *********************************************************************)
-(* Util functions used for the expansion of built-ins and some
+(* Util functions used for the expansion of built-ins and some
pseudo-instructions *)
open Asm
open AST
open Camlcoq
-
+
(* Buffering the expanded code *)
let current_code = ref ([]: instruction list)
@@ -25,7 +25,7 @@ let current_code = ref ([]: instruction list)
let emit i = current_code := i :: !current_code
(* Generation of fresh labels *)
-
+
let dummy_function = { fn_code = []; fn_sig = signature_main }
let current_function = ref dummy_function
let next_label = ref (None: label option)
@@ -46,7 +46,7 @@ let new_label () =
next_label := Some (P.succ lbl);
lbl
-
+
let set_current_function f =
current_function := f; next_label := None; current_code := []
@@ -64,7 +64,7 @@ let expand_scope id lbl oldscopes newscopes =
List.iter (fun i -> Debug.open_scope id i lbl) opening;
List.iter (fun i -> Debug.close_scope id i lbl) closing
-let translate_annot sp preg_to_dwarf annot =
+let translate_annot sp preg_to_dwarf annot =
let rec aux = function
| BA x -> Some (sp,BA (preg_to_dwarf x))
| BA_int _
@@ -75,7 +75,7 @@ let translate_annot sp preg_to_dwarf annot =
| BA_addrglobal _
| BA_loadstack _ -> None
| BA_addrstack ofs -> Some (sp,BA_addrstack ofs)
- | BA_splitlong (hi,lo) ->
+ | BA_splitlong (hi,lo) ->
begin
match (aux hi,aux lo) with
| Some (_,hi) ,Some (_,lo) -> Some (sp,BA_splitlong (hi,lo))
@@ -84,11 +84,11 @@ let translate_annot sp preg_to_dwarf annot =
(match annot with
| [] -> None
| a::_ -> aux a)
-
+
let expand_debug id sp preg simple l =
let get_lbl = function
- | None ->
+ | None ->
let lbl = new_label () in
emit (Plabel lbl);
lbl
@@ -100,12 +100,12 @@ let expand_debug id sp preg simple l =
let kind = (P.to_int kind) in
begin
match kind with
- | 1->
+ | 1->
emit i;aux lbl scopes rest
| 2 ->
aux lbl scopes rest
| 3 ->
- begin
+ begin
match translate_annot sp preg args with
| Some a ->
let lbl = get_lbl lbl in
diff --git a/backend/Asmgenproof0.v b/backend/Asmgenproof0.v
index 0533d561..cc27bd55 100644
--- a/backend/Asmgenproof0.v
+++ b/backend/Asmgenproof0.v
@@ -69,7 +69,7 @@ Hint Resolve data_diff: asmgen.
Lemma preg_of_not_SP:
forall r, preg_of r <> SP.
Proof.
- intros. unfold preg_of; destruct r; simpl; congruence.
+ intros. unfold preg_of; destruct r; simpl; congruence.
Qed.
Lemma preg_of_not_PC:
@@ -83,7 +83,7 @@ Hint Resolve preg_of_not_SP preg_of_not_PC: asmgen.
Lemma nextinstr_pc:
forall rs, (nextinstr rs)#PC = Val.add rs#PC Vone.
Proof.
- intros. apply Pregmap.gss.
+ intros. apply Pregmap.gss.
Qed.
Lemma nextinstr_inv:
@@ -102,16 +102,16 @@ Lemma nextinstr_set_preg:
forall rs m v,
(nextinstr (rs#(preg_of m) <- v))#PC = Val.add rs#PC Vone.
Proof.
- intros. unfold nextinstr. rewrite Pregmap.gss.
- rewrite Pregmap.gso. auto. apply sym_not_eq. apply preg_of_not_PC.
+ intros. unfold nextinstr. rewrite Pregmap.gss.
+ rewrite Pregmap.gso. auto. apply sym_not_eq. apply preg_of_not_PC.
Qed.
Lemma undef_regs_other:
- forall r rl rs,
+ forall r rl rs,
(forall r', In r' rl -> r <> r') ->
undef_regs rl rs r = rs r.
Proof.
- induction rl; simpl; intros. auto.
+ induction rl; simpl; intros. auto.
rewrite IHrl by auto. rewrite Pregmap.gso; auto.
Qed.
@@ -129,9 +129,9 @@ 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.
+ simpl. split. intros. intuition congruence. auto.
+ rewrite IHrl. split.
+ intros [A B]. intros. destruct H. congruence. auto.
auto.
Qed.
@@ -140,7 +140,7 @@ Lemma undef_regs_other_2:
preg_notin r rl ->
undef_regs (map preg_of rl) rs r = rs r.
Proof.
- intros. apply undef_regs_other. intros.
+ 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.
@@ -150,12 +150,12 @@ Lemma set_pregs_other_2:
preg_notin r rl ->
set_regs (map preg_of rl) vl rs r = rs r.
Proof.
- induction rl; simpl; intros.
+ 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.
+ rewrite IHrl by auto. apply Pregmap.gso; auto.
Qed.
(** * Agreement between Mach registers and processor registers *)
@@ -225,7 +225,7 @@ Lemma agree_set_mreg:
Proof.
intros. destruct H. split; auto.
rewrite H1; auto. apply sym_not_equal. apply preg_of_not_SP.
- intros. unfold Regmap.set. destruct (RegEq.eq r0 r). congruence.
+ intros. unfold Regmap.set. destruct (RegEq.eq r0 r). congruence.
rewrite H1. auto. apply preg_of_data.
red; intros; elim n. eapply preg_of_injective; eauto.
Qed.
@@ -253,12 +253,12 @@ Lemma agree_set_mregs:
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.
+ induction rl; simpl; intros.
auto.
- inv H0. auto. apply IHrl; auto.
- eapply agree_set_mreg. eexact H.
+ inv H0. auto. apply IHrl; auto.
+ eapply agree_set_mreg. eexact H.
rewrite Pregmap.gss. auto.
- intros. apply Pregmap.gso; auto.
+ intros. apply Pregmap.gso; auto.
Qed.
Lemma agree_undef_nondata_regs:
@@ -281,13 +281,13 @@ Lemma agree_undef_regs:
agree (Mach.undef_regs rl ms) sp rs'.
Proof.
intros. destruct H. split; auto.
- rewrite <- agree_sp0. apply H0; auto.
- rewrite preg_notin_charact. intros. apply not_eq_sym. apply preg_of_not_SP.
+ 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.
+ rewrite Mach.undef_regs_other; auto. rewrite H0; auto.
apply preg_of_data.
- rewrite preg_notin_charact. intros; red; intros. elim n.
+ rewrite preg_notin_charact. intros; red; intros. elim n.
exploit preg_of_injective; eauto. congruence.
Qed.
@@ -299,10 +299,10 @@ Lemma agree_set_undef_mreg:
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.
- 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.
+ 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.
Lemma agree_change_sp:
@@ -330,7 +330,7 @@ Proof.
exploit Mem.loadv_extends; eauto. intros [v' [A B]].
rewrite (sp_val _ _ _ H) in A.
exists v'; split; auto.
- econstructor. eauto. assumption.
+ econstructor. eauto. assumption.
Qed.
Lemma extcall_args_match:
@@ -339,7 +339,7 @@ Lemma extcall_args_match:
list_forall2 (Mach.extcall_arg ms m sp) ll vl ->
exists vl', list_forall2 (Asm.extcall_arg rs m') ll vl' /\ Val.lessdef_list vl vl'.
Proof.
- induction 3; intros.
+ induction 3; intros.
exists (@nil val); split. constructor. constructor.
exploit extcall_arg_match; eauto. intros [v1' [A B]].
destruct IHlist_forall2 as [vl' [C D]].
@@ -374,11 +374,11 @@ Lemma builtin_args_match:
Proof.
induction 3; intros; simpl.
exists (@nil val); split; constructor.
- exploit (@eval_builtin_arg_lessdef _ ge ms (fun r => rs (preg_of r))); eauto.
+ exploit (@eval_builtin_arg_lessdef _ ge ms (fun r => rs (preg_of r))); eauto.
intros; eapply preg_val; eauto.
intros (v1' & A & B).
destruct IHlist_forall2 as [vl' [C D]].
- exists (v1' :: vl'); split; constructor; auto. apply builtin_arg_match; auto.
+ exists (v1' :: vl'); split; constructor; auto. apply builtin_arg_match; auto.
Qed.
Lemma agree_set_res:
@@ -391,7 +391,7 @@ Proof.
- eapply agree_set_mreg; eauto. rewrite Pregmap.gss. auto.
intros. apply Pregmap.gso; auto.
- auto.
-- apply IHres2. apply IHres1. auto.
+- apply IHres2. apply IHres1. auto.
apply Val.hiword_lessdef; auto.
apply Val.loword_lessdef; auto.
Qed.
@@ -452,7 +452,7 @@ Remark code_tail_bounds_1:
code_tail ofs fn c -> 0 <= ofs <= list_length_z fn.
Proof.
induction 1; intros; simpl.
- generalize (list_length_z_pos c). omega.
+ generalize (list_length_z_pos c). omega.
rewrite list_length_z_cons. omega.
Qed.
@@ -462,7 +462,7 @@ Remark code_tail_bounds_2:
Proof.
assert (forall ofs fn c, code_tail ofs fn c ->
forall i c', c = i :: c' -> 0 <= ofs < list_length_z fn).
- induction 1; intros; simpl.
+ induction 1; intros; simpl.
rewrite H. rewrite list_length_z_cons. generalize (list_length_z_pos c'). omega.
rewrite list_length_z_cons. generalize (IHcode_tail _ _ H0). omega.
eauto.
@@ -490,7 +490,7 @@ Proof.
intros. rewrite Int.add_unsigned.
change (Int.unsigned Int.one) with 1.
rewrite Int.unsigned_repr. apply code_tail_next with i; auto.
- generalize (code_tail_bounds_2 _ _ _ _ H0). omega.
+ generalize (code_tail_bounds_2 _ _ _ _ H0). omega.
Qed.
(** [transl_code_at_pc pc fb f c ep tf tc] holds if the code pointer [pc] points
@@ -526,8 +526,8 @@ Lemma transl_code'_transl_code:
forall f il ep,
transl_code' f il ep = transl_code f il ep.
Proof.
- intros. unfold transl_code'. rewrite transl_code_rec_transl_code.
- destruct (transl_code f il ep); auto.
+ intros. unfold transl_code'. rewrite transl_code_rec_transl_code.
+ destruct (transl_code f il ep); auto.
Qed.
(** Predictor for return addresses in generated Asm code.
@@ -584,7 +584,7 @@ Hypothesis transf_function_inv:
Hypothesis transf_function_len:
forall f tf, transf_function f = OK tf -> list_length_z (fn_code tf) <= Int.max_unsigned.
-Lemma transl_code_tail:
+Lemma transl_code_tail:
forall f c1 c2, is_tail c1 c2 ->
forall tc2 ep2, transl_code f c2 ep2 = OK tc2 ->
exists tc1, exists ep1, transl_code f c1 ep1 = OK tc1 /\ is_tail tc1 tc2.
@@ -592,7 +592,7 @@ Proof.
induction 1; simpl; intros.
exists tc2; exists ep2; split; auto with coqlib.
monadInv H0. exploit IHis_tail; eauto. intros [tc1 [ep1 [A B]]].
- exists tc1; exists ep1; split. auto.
+ exists tc1; exists ep1; split. auto.
apply is_tail_trans with x. auto. eapply transl_instr_tail; eauto.
Qed.
@@ -604,17 +604,17 @@ Proof.
+ exploit transf_function_inv; eauto. intros (tc1 & ep1 & TR1 & TL1).
exploit transl_code_tail; eauto. intros (tc2 & ep2 & TR2 & TL2).
Opaque transl_instr.
- monadInv TR2.
+ monadInv TR2.
assert (TL3: is_tail x (fn_code tf)).
- { apply is_tail_trans with tc1; auto.
+ { apply is_tail_trans with tc1; auto.
apply is_tail_trans with tc2; auto.
eapply transl_instr_tail; eauto. }
exploit is_tail_code_tail. eexact TL3. intros [ofs CT].
- exists (Int.repr ofs). red; intros.
- rewrite Int.unsigned_repr. congruence.
+ exists (Int.repr ofs). red; intros.
+ rewrite Int.unsigned_repr. congruence.
exploit code_tail_bounds_1; eauto.
- apply transf_function_len in TF. omega.
-+ exists Int.zero; red; intros. congruence.
+ apply transf_function_len in TF. omega.
++ exists Int.zero; red; intros. congruence.
Qed.
End RETADDR_EXISTS.
@@ -641,9 +641,9 @@ Lemma return_address_offset_correct:
return_address_offset f c ofs' ->
ofs' = ofs.
Proof.
- intros. inv H. red in H0.
+ intros. inv H. red in H0.
exploit code_tail_unique. eexact H12. eapply H0; eauto. intro.
- rewrite <- (Int.repr_unsigned ofs).
+ rewrite <- (Int.repr_unsigned ofs).
rewrite <- (Int.repr_unsigned ofs').
congruence.
Qed.
@@ -662,26 +662,26 @@ Lemma label_pos_code_tail:
forall lbl c pos c',
find_label lbl c = Some c' ->
exists pos',
- label_pos lbl pos c = Some pos'
+ label_pos lbl pos c = Some pos'
/\ code_tail (pos' - pos) c c'
/\ pos < pos' <= pos + list_length_z c.
Proof.
- induction c.
+ induction c.
simpl; intros. discriminate.
simpl; intros until c'.
case (is_label lbl a).
intro EQ; injection EQ; intro; subst c'.
exists (pos + 1). split. auto. split.
- replace (pos + 1 - pos) with (0 + 1) by omega. constructor. constructor.
- rewrite list_length_z_cons. generalize (list_length_z_pos c). omega.
+ replace (pos + 1 - pos) with (0 + 1) by omega. constructor. constructor.
+ rewrite list_length_z_cons. generalize (list_length_z_pos c). omega.
intros. generalize (IHc (pos + 1) c' H). intros [pos' [A [B C]]].
exists pos'. split. auto. split.
replace (pos' - pos) with ((pos' - (pos + 1)) + 1) by omega.
- constructor. auto.
+ constructor. auto.
rewrite list_length_z_cons. omega.
Qed.
-(** Helper lemmas to reason about
+(** Helper lemmas to reason about
- the "code is tail of" property
- correct translation of labels. *)
@@ -697,7 +697,7 @@ Qed.
Lemma tail_nolabel_trans:
forall c1 c2 c3, tail_nolabel c2 c3 -> tail_nolabel c1 c2 -> tail_nolabel c1 c3.
Proof.
- intros. destruct H; destruct H0; split.
+ intros. destruct H; destruct H0; split.
eapply is_tail_trans; eauto.
intros. rewrite H1; auto.
Qed.
@@ -711,7 +711,7 @@ Lemma tail_nolabel_cons:
forall i c k,
nolabel i -> tail_nolabel k c -> tail_nolabel k (i :: c).
Proof.
- intros. destruct H0. split.
+ intros. destruct H0. split.
constructor; auto.
intros. simpl. rewrite <- H1. destruct i; reflexivity || contradiction.
Qed.
@@ -745,7 +745,7 @@ Variable fn: function.
Instructions are taken from the first list instead of being fetched
from memory. *)
-Inductive exec_straight: code -> regset -> mem ->
+Inductive exec_straight: code -> regset -> mem ->
code -> regset -> mem -> Prop :=
| exec_straight_one:
forall i1 c rs1 m1 rs2 m2,
@@ -811,18 +811,18 @@ Lemma exec_straight_steps_1:
Proof.
induction 1; intros.
apply plus_one.
- econstructor; eauto.
+ econstructor; eauto.
eapply find_instr_tail. eauto.
eapply plus_left'.
- econstructor; eauto.
+ econstructor; eauto.
eapply find_instr_tail. eauto.
- apply IHexec_straight with b (Int.add ofs Int.one).
+ apply IHexec_straight with b (Int.add ofs Int.one).
auto. rewrite H0. rewrite H3. reflexivity.
- auto.
+ auto.
apply code_tail_next_int with i; auto.
traceEq.
Qed.
-
+
Lemma exec_straight_steps_2:
forall c rs m c' rs' m',
exec_straight c rs m c' rs' m' ->
@@ -840,7 +840,7 @@ Proof.
rewrite H0. rewrite H2. auto.
apply code_tail_next_int with i1; auto.
apply IHexec_straight with (Int.add ofs Int.one).
- auto. rewrite H0. rewrite H3. reflexivity. auto.
+ auto. rewrite H0. rewrite H3. reflexivity. auto.
apply code_tail_next_int with i; auto.
Qed.
diff --git a/backend/Bounds.v b/backend/Bounds.v
index beb29965..2a63b1d5 100644
--- a/backend/Bounds.v
+++ b/backend/Bounds.v
@@ -161,7 +161,7 @@ Proof.
intros until valu. unfold max_over_list.
assert (forall l z, fold_left (fun x y => Zmax x (valu y)) l z >= z).
induction l; simpl; intros.
- omega. apply Zge_trans with (Zmax z (valu a)).
+ omega. apply Zge_trans with (Zmax z (valu a)).
auto. apply Zle_ge. apply Zmax1. auto.
Qed.
@@ -193,7 +193,7 @@ Program Definition function_bounds :=
_ _.
Next Obligation.
apply Zle_ge. eapply Zle_trans. 2: apply Zmax2.
- apply Zge_le. apply max_over_slots_of_funct_pos.
+ apply Zge_le. apply max_over_slots_of_funct_pos.
Qed.
Next Obligation.
apply Zle_ge. apply Zmax2.
@@ -211,10 +211,10 @@ Proof.
z <= f /\ (In x c -> valu x <= f)).
induction c; simpl; intros.
split. omega. tauto.
- elim (IHc (Zmax z (valu a))); intros.
- split. apply Zle_trans with (Zmax z (valu a)). apply Zmax1. auto.
- intro H1; elim H1; intro.
- subst a. apply Zle_trans with (Zmax z (valu x)).
+ elim (IHc (Zmax z (valu a))); intros.
+ split. apply Zle_trans with (Zmax z (valu a)). apply Zmax1. auto.
+ intro H1; elim H1; intro.
+ subst a. apply Zle_trans with (Zmax z (valu x)).
apply Zmax2. auto. auto.
intro. elim (H l 0); intros. auto.
Qed.
@@ -231,7 +231,7 @@ Lemma max_over_regs_of_funct_bound:
In i f.(fn_code) -> In r (regs_of_instr i) ->
valu r <= max_over_regs_of_funct valu.
Proof.
- intros. unfold max_over_regs_of_funct.
+ intros. unfold max_over_regs_of_funct.
apply Zle_trans with (max_over_regs_of_instr valu i).
unfold max_over_regs_of_instr. apply max_over_list_bound. auto.
apply max_over_instrs_bound. auto.
@@ -242,7 +242,7 @@ Lemma max_over_slots_of_funct_bound:
In i f.(fn_code) -> In s (slots_of_instr i) ->
valu s <= max_over_slots_of_funct valu.
Proof.
- intros. unfold max_over_slots_of_funct.
+ intros. unfold max_over_slots_of_funct.
apply Zle_trans with (max_over_slots_of_instr valu i).
unfold max_over_slots_of_instr. apply max_over_list_bound. auto.
apply max_over_instrs_bound. auto.
@@ -255,7 +255,7 @@ Lemma int_callee_save_bound:
Proof.
intros. apply Zlt_le_trans with (int_callee_save r).
unfold int_callee_save. omega.
- unfold function_bounds, bound_int_callee_save.
+ unfold function_bounds, bound_int_callee_save.
eapply max_over_regs_of_funct_bound; eauto.
Qed.
@@ -266,7 +266,7 @@ Lemma float_callee_save_bound:
Proof.
intros. apply Zlt_le_trans with (float_callee_save r).
unfold float_callee_save. omega.
- unfold function_bounds, bound_float_callee_save.
+ unfold function_bounds, bound_float_callee_save.
eapply max_over_regs_of_funct_bound; eauto.
Qed.
@@ -315,11 +315,11 @@ Proof.
Qed.
Lemma slot_is_within_bounds:
- forall i, In i f.(fn_code) ->
+ forall i, In i f.(fn_code) ->
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.
+ intros. unfold slot_within_bounds.
destruct sl.
eapply local_slot_bound; eauto.
auto.
@@ -329,12 +329,12 @@ Qed.
Lemma slots_of_locs_charact:
forall sl ofs ty l, In (sl, ofs, ty) (slots_of_locs l) <-> In (S sl ofs ty) l.
Proof.
- induction l; simpl; intros.
+ induction l; simpl; intros.
tauto.
destruct a; simpl; intuition congruence.
Qed.
-(** It follows that every instruction in the function is within bounds,
+(** It follows that every instruction in the function is within bounds,
in the sense of the [instr_within_bounds] predicate. *)
Lemma instr_is_within_bounds:
@@ -342,16 +342,16 @@ Lemma instr_is_within_bounds:
In i f.(fn_code) ->
instr_within_bounds function_bounds i.
Proof.
- intros;
+ intros;
destruct i;
- generalize (mreg_is_within_bounds _ H); generalize (slot_is_within_bounds _ H);
+ generalize (mreg_is_within_bounds _ H); generalize (slot_is_within_bounds _ H);
simpl; intros; auto.
(* call *)
eapply size_arguments_bound; eauto.
(* builtin *)
split; intros.
apply H1. apply in_or_app; auto.
- apply H0. rewrite slots_of_locs_charact; auto.
+ apply H0. rewrite slots_of_locs_charact; auto.
Qed.
Lemma function_is_within_bounds:
diff --git a/backend/CMlexer.mll b/backend/CMlexer.mll
index 7ed5b4ab..6695b6b7 100644
--- a/backend/CMlexer.mll
+++ b/backend/CMlexer.mll
@@ -21,8 +21,8 @@ exception Error of string
}
let blank = [' ' '\009' '\012' '\010' '\013']
-let floatlit =
- ("-"? (['0'-'9'] ['0'-'9' '_']*
+let floatlit =
+ ("-"? (['0'-'9'] ['0'-'9' '_']*
('.' ['0'-'9' '_']* )?
(['e' 'E'] ['+' '-']? ['0'-'9'] ['0'-'9' '_']*)? )) | "inf" | "nan"
let ident = ['A'-'Z' 'a'-'z' '_'] ['A'-'Z' 'a'-'z' '_' '$' '0'-'9']*
@@ -69,7 +69,7 @@ rule token = parse
| "floatofintu" { FLOATOFINTU }
| "floatoflong" { FLOATOFLONG }
| "floatoflongu" { FLOATOFLONGU }
- | "goto" { GOTO }
+ | "goto" { GOTO }
| ">" { GREATER }
| ">f" { GREATERF }
| ">l" { GREATERL }
diff --git a/backend/CMparser.mly b/backend/CMparser.mly
index 41bd35a1..5f189e7b 100644
--- a/backend/CMparser.mly
+++ b/backend/CMparser.mly
@@ -13,7 +13,7 @@
/* */
/* *********************************************************************/
-/* Note that this compiles a superset of the language defined by the AST,
+/* Note that this compiles a superset of the language defined by the AST,
including function calls in expressions, matches, while statements, etc. */
%{
@@ -123,7 +123,7 @@ let mkeval e =
let c1 = convert_rexpr e1 in
let cl = convert_rexpr_list el in
prepend_seq !convert_accu (Scall(None, sg, c1, cl))
- | Rbuiltin(sg, pef, el) ->
+ | Rbuiltin(sg, pef, el) ->
let ef = mkef sg pef in
let cl = convert_rexpr_list el in
prepend_seq !convert_accu (Sbuiltin(None, ef, cl))
@@ -322,9 +322,9 @@ let mkmatch expr cases =
%token LESSLESSL
%token LONG
%token <int64> LONGLIT
-%token LONGOFINT
+%token LONGOFINT
%token LONGOFINTU
-%token LONGOFFLOAT
+%token LONGOFFLOAT
%token LONGUOFFLOAT
%token LOOP
%token LPAREN
@@ -375,7 +375,7 @@ let mkmatch expr cases =
%left CARET CARETL
%left AMPERSAND AMPERSANDL
%left EQUALEQUAL BANGEQUAL LESS LESSEQUAL GREATER GREATEREQUAL EQUALEQUALU BANGEQUALU LESSU LESSEQUALU GREATERU GREATEREQUALU EQUALEQUALF BANGEQUALF LESSF LESSEQUALF GREATERF GREATEREQUALF EQUALEQUALL BANGEQUALL LESSL LESSEQUALL GREATERL GREATEREQUALL EQUALEQUALLU BANGEQUALLU LESSLU LESSEQUALLU GREATERLU GREATEREQUALLU
-%left LESSLESS GREATERGREATER GREATERGREATERU LESSLESSL GREATERGREATERL GREATERGREATERLU
+%left LESSLESS GREATERGREATER GREATERGREATERU LESSLESSL GREATERGREATERL GREATERGREATERLU
%left PLUS PLUSF PLUSL MINUS MINUSF MINUSL
%left STAR SLASH PERCENT STARF SLASHF SLASHU PERCENTU STARL SLASHL SLASHLU PERCENTL PERCENTLU
%nonassoc BANG TILDE TILDEL p_uminus ABSF INTOFFLOAT INTUOFFLOAT FLOATOFINT FLOATOFINTU INT8S INT8U INT16S INT16U FLOAT32 INTOFLONG LONGOFINT LONGOFINTU LONGOFFLOAT LONGUOFFLOAT FLOATOFLONG FLOATOFLONGU
@@ -432,12 +432,12 @@ init_data_list:
/* empty */ { [] }
| init_data_list_1 { $1 }
;
-
+
init_data_list_1:
init_data { [$1] }
| init_data_list_1 COMMA init_data { $3 :: $1 }
;
-
+
init_data:
INT8 INTLIT { Init_int8 (coqint_of_camlint $2) }
| INT16 INTLIT { Init_int16 (coqint_of_camlint $2) }
@@ -453,7 +453,7 @@ init_data:
| LBRACKET INTLIT RBRACKET { Init_space (Z.of_sint32 $2) }
| INTLIT LPAREN STRINGLIT RPAREN { Init_addrof (intern_string $3, coqint_of_camlint $1) }
;
-
+
/* Procedures */
proc:
@@ -472,14 +472,14 @@ proc:
fn_vars = List.rev (tmp @ $9);
fn_stackspace = $8;
fn_body = $10 })) }
- | EXTERN STRINGLIT COLON signature
+ | EXTERN STRINGLIT COLON signature
{ (intern_string $2, Gfun(External(EF_external(coqstring_of_camlstring $2,$4)))) }
| EXTERN STRINGLIT EQUAL eftoks COLON signature
{ (intern_string $2, Gfun(External(mkef $6 $4))) }
;
signature:
- type_
+ type_
{ {sig_args = []; sig_res = Some $1; sig_cc = cc_default} }
| VOID
{ {sig_args = []; sig_res = None; sig_cc = cc_default} }
diff --git a/backend/CMtypecheck.ml b/backend/CMtypecheck.ml
index aacbf86f..72bf9cb4 100644
--- a/backend/CMtypecheck.ml
+++ b/backend/CMtypecheck.ml
@@ -323,7 +323,7 @@ let rec type_stmt env blk ret s =
| Sreturn (Some e) ->
begin match ret with
| None -> raise (Error "return with argument")
- | Some tret ->
+ | Some tret ->
begin try
unify (type_expr env [] e) (ty_of_typ tret)
with Error s ->
diff --git a/backend/CSE.v b/backend/CSE.v
index ebeb921e..63dadbc7 100644
--- a/backend/CSE.v
+++ b/backend/CSE.v
@@ -181,7 +181,7 @@ Definition add_rhs (n: numbering) (rd: reg) (rh: rhs) : numbering :=
(** [add_op n rd op rs] specializes [add_rhs] for the case of an
arithmetic operation. The right-hand side corresponding to [op]
and the value numbers for the argument registers [rs] is built
- and added to [n] as described in [add_rhs].
+ and added to [n] as described in [add_rhs].
If [op] is a move instruction, we simply assign the value number of
the source register to the destination register, since we know that
@@ -190,7 +190,7 @@ Definition add_rhs (n: numbering) (rd: reg) (rh: rhs) : numbering :=
<<
z = add(x, y); u = x; v = add(u, y);
>>
- Since [u] and [x] have the same value number, the second [add]
+ Since [u] and [x] have the same value number, the second [add]
is recognized as computing the same result as the first [add],
and therefore [u] and [z] have the same value number. *)
@@ -212,13 +212,13 @@ Definition add_op (n: numbering) (rd: reg) (op: operation) (rs: list reg) :=
and the value numbers for the argument registers [rs] is built
and added to [n] as described in [add_rhs]. *)
-Definition add_load (n: numbering) (rd: reg)
+Definition add_load (n: numbering) (rd: reg)
(chunk: memory_chunk) (addr: addressing)
(rs: list reg) :=
let (n1, vs) := valnum_regs n rs in
add_rhs n1 rd (Load chunk addr vs).
-(** [set_unknown n rd] returns a numbering where [rd] is mapped to
+(** [set_unknown n rd] returns a numbering where [rd] is mapped to
no value number, and no equations are added. This is useful
to model instructions with unpredictable results such as [Ibuiltin]. *)
@@ -323,7 +323,7 @@ Definition kill_loads_after_storebytes
(app: VA.t) (n: numbering) (dst: aptr) (sz: Z) :=
kill_equations (filter_after_store app n dst sz) n.
-(** [add_memcpy app n1 n2 rsrc rdst sz] adds equations to [n2] that
+(** [add_memcpy app n1 n2 rsrc rdst sz] adds equations to [n2] that
represent the effect of a [memcpy] block copy operation of [sz] bytes
from the address denoted by [rsrc] to the address denoted by [rdst].
[n2] is the numbering returned by [kill_loads_after_storebytes]
@@ -415,7 +415,7 @@ End REDUCE.
Module Numbering.
Definition t := numbering.
Definition ge (n1 n2: numbering) : Prop :=
- forall valu ge sp rs m,
+ forall valu ge sp rs m,
numbering_holds valu ge sp rs m n2 ->
numbering_holds valu ge sp rs m n1.
Definition top := empty_numbering.
@@ -443,7 +443,7 @@ Module Solver := BBlock_solver(Numbering).
then add an equation for loads from the same location stored to.
For [Icall] instructions, we could simply associate a fresh, unconstrained by equations value number
to the result register. However, it is often undesirable to eliminate
- common subexpressions across a function call (there is a risk of
+ common subexpressions across a function call (there is a risk of
increasing too much the register pressure across the call), so we
just forget all equations and start afresh with an empty numbering.
Finally, for instructions that modify neither registers nor
diff --git a/backend/CSEdomain.v b/backend/CSEdomain.v
index 6a75d511..9b1243c8 100644
--- a/backend/CSEdomain.v
+++ b/backend/CSEdomain.v
@@ -78,7 +78,7 @@ Definition valnums_rhs (r: rhs): list valnum :=
end.
Definition wf_rhs (next: valnum) (r: rhs) : Prop :=
-forall v, In v (valnums_rhs r) -> Plt v next.
+forall v, In v (valnums_rhs r) -> Plt v next.
Definition wf_equation (next: valnum) (e: equation) : Prop :=
match e with Eq l str r => Plt l next /\ wf_rhs next r end.
diff --git a/backend/CSEproof.v b/backend/CSEproof.v
index 70f9bfc7..07c7008d 100644
--- a/backend/CSEproof.v
+++ b/backend/CSEproof.v
@@ -100,9 +100,9 @@ Lemma numbering_holds_exten:
Proof.
intros. destruct H. constructor; intros.
- auto.
-- apply equation_holds_exten. auto.
- eapply wf_equation_incr; eauto with cse.
-- rewrite AGREE. eauto. eapply Plt_le_trans; eauto. eapply wf_num_reg; eauto.
+- apply equation_holds_exten. auto.
+ eapply wf_equation_incr; eauto with cse.
+- rewrite AGREE. eauto. eapply Plt_le_trans; eauto. eapply wf_num_reg; eauto.
Qed.
End EXTEN.
@@ -136,16 +136,16 @@ Proof.
+ constructor; simpl; intros.
* constructor; simpl; intros.
apply wf_equation_incr with (num_next n). eauto with cse. xomega.
- rewrite PTree.gsspec in H0. destruct (peq r0 r).
+ rewrite PTree.gsspec in H0. destruct (peq r0 r).
inv H0; xomega.
apply Plt_trans_succ; eauto with cse.
rewrite PMap.gsspec in H0. destruct (peq v (num_next n)).
replace r0 with r by (simpl in H0; intuition). rewrite PTree.gss. subst; auto.
- exploit wf_num_val; eauto with cse. intro.
+ exploit wf_num_val; eauto with cse. intro.
rewrite PTree.gso by congruence. auto.
* eapply equation_holds_exten; eauto with cse.
* unfold valu2. rewrite PTree.gsspec in H0. destruct (peq r0 r).
- inv H0. rewrite peq_true; auto.
+ inv H0. rewrite peq_true; auto.
rewrite peq_false. eauto with cse. apply Plt_ne; eauto with cse.
+ unfold valu2. rewrite peq_true; auto.
+ auto.
@@ -169,9 +169,9 @@ Proof.
- destruct (valnum_reg n a) as [n1 v1] eqn:V1.
destruct (valnum_regs n1 rl) as [n2 vs] eqn:V2.
inv H0.
- exploit valnum_reg_holds; eauto.
+ exploit valnum_reg_holds; eauto.
intros (valu2 & A & B & C & D & E).
- exploit (IHrl valu2); eauto.
+ exploit (IHrl valu2); eauto.
intros (valu3 & P & Q & R & S & T).
exists valu3; splitall.
+ auto.
@@ -187,7 +187,7 @@ Lemma find_valnum_rhs_charact:
Proof.
induction eqs; simpl; intros.
- inv H.
-- destruct a. destruct (strict && eq_rhs rh r) eqn:T.
+- destruct a. destruct (strict && eq_rhs rh r) eqn:T.
+ InvBooleans. inv H. left; auto.
+ right; eauto.
Qed.
@@ -198,9 +198,9 @@ Lemma find_valnum_rhs'_charact:
Proof.
induction eqs; simpl; intros.
- inv H.
-- destruct a. destruct (eq_rhs rh r) eqn:T.
+- destruct a. destruct (eq_rhs rh r) eqn:T.
+ inv H. exists strict; auto.
- + exploit IHeqs; eauto. intros [s IN]. exists s; auto.
+ + exploit IHeqs; eauto. intros [s IN]. exists s; auto.
Qed.
Lemma find_valnum_num_charact:
@@ -208,8 +208,8 @@ Lemma find_valnum_num_charact:
Proof.
induction eqs; simpl; intros.
- inv H.
-- destruct a. destruct (strict && peq v v0) eqn:T.
- + InvBooleans. inv H. auto.
+- destruct a. destruct (strict && peq v v0) eqn:T.
+ + InvBooleans. inv H. auto.
+ eauto.
Qed.
@@ -220,7 +220,7 @@ Lemma reg_valnum_sound:
rs#r = valu v.
Proof.
unfold reg_valnum; intros. destruct (num_val n)#v as [ | r1 rl] eqn:E; inv H.
- eapply num_holds_reg; eauto. eapply wf_num_val; eauto with cse.
+ eapply num_holds_reg; eauto. eapply wf_num_val; eauto with cse.
rewrite E; auto with coqlib.
Qed.
@@ -235,7 +235,7 @@ Proof.
- inv H0; auto.
- destruct (reg_valnum n a) as [r1|] eqn:RV1; try discriminate.
destruct (regs_valnums n vl) as [rl1|] eqn:RVL; inv H0.
- simpl; f_equal. eapply reg_valnum_sound; eauto. eauto.
+ simpl; f_equal. eapply reg_valnum_sound; eauto. eauto.
Qed.
Lemma find_rhs_sound:
@@ -256,10 +256,10 @@ Remark in_remove:
forall (A: Type) (eq: forall (x y: A), {x=y}+{x<>y}) x y l,
In y (List.remove eq x l) <-> x <> y /\ In y l.
Proof.
- induction l; simpl.
+ induction l; simpl.
tauto.
- destruct (eq x a).
- subst a. rewrite IHl. tauto.
+ destruct (eq x a).
+ subst a. rewrite IHl. tauto.
simpl. rewrite IHl. intuition congruence.
Qed.
@@ -274,7 +274,7 @@ Proof.
+ subst v. rewrite in_remove in H0. intuition.
+ split; auto. exploit wf_num_val; eauto. congruence.
- split; auto. exploit wf_num_val; eauto. congruence.
-Qed.
+Qed.
Lemma update_reg_charact:
forall n rd vd r v,
@@ -285,7 +285,7 @@ Proof.
unfold update_reg; intros.
rewrite PMap.gsspec in H0.
destruct (peq v vd).
-- subst v. destruct H0.
+- subst v. destruct H0.
+ subst r. apply PTree.gss.
+ exploit forget_reg_charact; eauto. intros [A B].
rewrite PTree.gso by auto. eapply wf_num_val; eauto.
@@ -324,7 +324,7 @@ Proof.
eauto with cse.
* eapply update_reg_charact; eauto with cse.
+ eauto with cse.
-+ rewrite PTree.gsspec in H5. destruct (peq r rd).
++ rewrite PTree.gsspec in H5. destruct (peq r rd).
congruence.
rewrite H2 by auto. eauto with cse.
@@ -334,17 +334,17 @@ Proof.
{ red; intros. unfold valu2. apply peq_false. apply Plt_ne; auto. }
exists valu2; constructor; simpl; intros.
+ constructor; simpl; intros.
- * destruct H3. inv H3. simpl; split. xomega.
- red; intros. apply Plt_trans_succ; eauto.
- apply wf_equation_incr with (num_next n). eauto with cse. xomega.
+ * destruct H3. inv H3. simpl; split. xomega.
+ red; intros. apply Plt_trans_succ; eauto.
+ apply wf_equation_incr with (num_next n). eauto with cse. xomega.
* rewrite PTree.gsspec in H3. destruct (peq r rd).
inv H3. xomega.
apply Plt_trans_succ; eauto with cse.
* apply update_reg_charact; eauto with cse.
+ destruct H3. inv H3.
- constructor. unfold valu2 at 2; rewrite peq_true.
- eapply rhs_eval_to_exten; eauto.
- eapply equation_holds_exten; eauto with cse.
+ constructor. unfold valu2 at 2; rewrite peq_true.
+ eapply rhs_eval_to_exten; eauto.
+ eapply equation_holds_exten; eauto with cse.
+ rewrite PTree.gsspec in H3. unfold valu2. destruct (peq r rd).
inv H3. rewrite peq_true; auto.
rewrite peq_false. rewrite H2 by auto. eauto with cse.
@@ -363,7 +363,7 @@ Proof.
exploit is_move_operation_correct; eauto. intros [A B]; subst op args.
simpl in H0. inv H0.
destruct (valnum_reg n src) as [n1 vsrc] eqn:VN.
- exploit valnum_reg_holds; eauto.
+ exploit valnum_reg_holds; eauto.
intros (valu2 & A & B & C & D & E).
exists valu2; constructor; simpl; intros.
+ constructor; simpl; intros; eauto with cse.
@@ -372,15 +372,15 @@ Proof.
eauto with cse.
* eapply update_reg_charact; eauto with cse.
+ eauto with cse.
-+ rewrite PTree.gsspec in H0. rewrite Regmap.gsspec.
++ rewrite PTree.gsspec in H0. rewrite Regmap.gsspec.
destruct (peq r dst). congruence. eauto with cse.
- (* general case *)
destruct (valnum_regs n args) as [n1 vl] eqn:VN.
- exploit valnum_regs_holds; eauto.
+ exploit valnum_regs_holds; eauto.
intros (valu2 & A & B & C & D & E).
- eapply add_rhs_holds; eauto.
-+ constructor. rewrite Regmap.gss. congruence.
+ eapply add_rhs_holds; eauto.
++ constructor. rewrite Regmap.gss. congruence.
+ intros. apply Regmap.gso; auto.
Qed.
@@ -393,10 +393,10 @@ Lemma add_load_holds:
Proof.
unfold add_load; intros.
destruct (valnum_regs n args) as [n1 vl] eqn:VN.
- exploit valnum_regs_holds; eauto.
+ exploit valnum_regs_holds; eauto.
intros (valu2 & A & B & C & D & E).
- eapply add_rhs_holds; eauto.
-+ econstructor. rewrite <- B; eauto. rewrite Regmap.gss; auto.
+ eapply add_rhs_holds; eauto.
++ econstructor. rewrite <- B; eauto. rewrite Regmap.gss; auto.
+ intros. apply Regmap.gso; auto.
Qed.
@@ -408,13 +408,13 @@ Proof.
intros; constructor; simpl; intros.
- constructor; simpl; intros.
+ eauto with cse.
- + rewrite PTree.grspec in H0. destruct (PTree.elt_eq r0 r).
- discriminate.
+ + rewrite PTree.grspec in H0. destruct (PTree.elt_eq r0 r).
+ discriminate.
eauto with cse.
- + exploit forget_reg_charact; eauto with cse. intros [A B].
+ + exploit forget_reg_charact; eauto with cse. intros [A B].
rewrite PTree.gro; eauto with cse.
- eauto with cse.
-- rewrite PTree.grspec in H0. destruct (PTree.elt_eq r0 r).
+- rewrite PTree.grspec in H0. destruct (PTree.elt_eq r0 r).
discriminate.
rewrite Regmap.gso; eauto with cse.
Qed.
@@ -429,7 +429,7 @@ Qed.
Lemma kill_eqs_charact:
forall pred l strict r eqs,
- In (Eq l strict r) (kill_eqs pred eqs) ->
+ In (Eq l strict r) (kill_eqs pred eqs) ->
pred r = false /\ In (Eq l strict r) eqs.
Proof.
induction eqs; simpl; intros.
@@ -451,7 +451,7 @@ Proof.
intros; constructor; simpl; intros.
- constructor; simpl; intros; eauto with cse.
destruct e. exploit kill_eqs_charact; eauto. intros [A B]. eauto with cse.
-- destruct eq. exploit kill_eqs_charact; eauto. intros [A B].
+- destruct eq. exploit kill_eqs_charact; eauto. intros [A B].
exploit num_holds_eq; eauto. intro EH; inv EH; econstructor; eauto.
- eauto with cse.
Qed.
@@ -461,7 +461,7 @@ Lemma kill_all_loads_hold:
numbering_holds valu ge sp rs m n ->
numbering_holds valu ge sp rs m' (kill_all_loads n).
Proof.
- intros. eapply kill_equations_hold; eauto.
+ intros. eapply kill_equations_hold; eauto.
unfold filter_loads; intros. inv H1.
constructor. rewrite <- H2. apply op_depends_on_memory_correct; auto.
discriminate.
@@ -486,11 +486,11 @@ Proof.
econstructor; eauto. rewrite <- H9.
destruct a; simpl in H1; try discriminate.
destruct a0; simpl in H9; try discriminate.
- simpl.
+ simpl.
rewrite negb_false_iff in H6. unfold aaddressing in H6.
eapply Mem.load_store_other. eauto.
- eapply pdisjoint_sound. eauto.
- apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
+ eapply pdisjoint_sound. eauto.
+ apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
erewrite <- regs_valnums_sound by eauto. eauto with va.
apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto with va.
Qed.
@@ -516,22 +516,22 @@ Lemma add_store_result_hold:
approx = VA.State ae am ->
exists valu2, numbering_holds valu2 ge sp rs m' (add_store_result approx n chunk addr args src).
Proof.
- unfold add_store_result; intros.
- unfold avalue; rewrite H3.
+ unfold add_store_result; intros.
+ unfold avalue; rewrite H3.
destruct (vincl (AE.get src ae) (store_normalized_range chunk)) eqn:INCL.
- destruct (valnum_reg n src) as [n1 vsrc] eqn:VR1.
destruct (valnum_regs n1 args) as [n2 vargs] eqn:VR2.
- exploit valnum_reg_holds; eauto. intros (valu2 & A & B & C & D & E).
+ exploit valnum_reg_holds; eauto. intros (valu2 & A & B & C & D & E).
exploit valnum_regs_holds; eauto. intros (valu3 & P & Q & R & S & T).
exists valu3. constructor; simpl; intros.
+ constructor; simpl; intros; eauto with cse.
- destruct H4; eauto with cse. subst e. split.
- eapply Plt_le_trans; eauto.
+ destruct H4; eauto with cse. subst e. split.
+ eapply Plt_le_trans; eauto.
red; simpl; intros. auto.
+ destruct H4; eauto with cse. subst eq. apply eq_holds_lessdef with (Val.load_result chunk rs#src).
apply load_eval_to with a. rewrite <- Q; auto.
destruct a; try discriminate. simpl. eapply Mem.load_store_same; eauto.
- rewrite B. rewrite R by auto. apply store_normalized_range_sound with bc.
+ rewrite B. rewrite R by auto. apply store_normalized_range_sound with bc.
rewrite <- B. eapply vmatch_ge. apply vincl_ge; eauto. apply H2.
+ eauto with cse.
@@ -557,12 +557,12 @@ Proof.
- destruct (regs_valnums n vl) as [rl|] eqn:RV; try discriminate.
econstructor; eauto. rewrite <- H11.
destruct a; simpl in H10; try discriminate.
- simpl.
+ simpl.
rewrite negb_false_iff in H8.
eapply Mem.load_storebytes_other. eauto.
- rewrite H6. rewrite nat_of_Z_eq by auto.
- eapply pdisjoint_sound. eauto.
- unfold aaddressing. apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
+ rewrite H6. rewrite nat_of_Z_eq by auto.
+ eapply pdisjoint_sound. eauto.
+ unfold aaddressing. apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
erewrite <- regs_valnums_sound by eauto. eauto with va.
auto.
Qed.
@@ -576,14 +576,14 @@ Lemma load_memcpy:
(align_chunk chunk | ofs2 - ofs1) ->
Mem.load chunk m' b2 (i + (ofs2 - ofs1)) = Some v.
Proof.
- intros.
+ intros.
generalize (size_chunk_pos chunk); intros SPOS.
set (n1 := i - ofs1).
set (n2 := size_chunk chunk).
set (n3 := sz - (n1 + n2)).
replace sz with (n1 + (n2 + n3)) in H by (unfold n3, n2, n1; omega).
- exploit Mem.loadbytes_split; eauto.
- unfold n1; omega.
+ exploit Mem.loadbytes_split; eauto.
+ unfold n1; omega.
unfold n3, n2, n1; omega.
intros (bytes1 & bytes23 & LB1 & LB23 & EQ).
clear H.
@@ -591,7 +591,7 @@ Proof.
unfold n2; omega.
unfold n3, n2, n1; omega.
intros (bytes2 & bytes3 & LB2 & LB3 & EQ').
- subst bytes23; subst bytes.
+ subst bytes23; subst bytes.
exploit Mem.load_loadbytes; eauto. intros (bytes2' & A & B).
assert (bytes2' = bytes2).
{ replace (ofs1 + n1) with i in LB2 by (unfold n1; omega). unfold n2 in LB2. congruence. }
@@ -604,17 +604,17 @@ Proof.
{ erewrite Mem.loadbytes_length by eauto. apply nat_of_Z_eq. unfold n1; omega. }
assert (L2: Z.of_nat (length bytes2) = n2).
{ erewrite Mem.loadbytes_length by eauto. apply nat_of_Z_eq. unfold n2; omega. }
- rewrite L1 in *. rewrite L2 in *.
+ rewrite L1 in *. rewrite L2 in *.
assert (LB': Mem.loadbytes m2 b2 (ofs2 + n1) n2 = Some bytes2).
{ rewrite <- L2. eapply Mem.loadbytes_storebytes_same; eauto. }
assert (LB'': Mem.loadbytes m' b2 (ofs2 + n1) n2 = Some bytes2).
- { rewrite <- LB'. eapply Mem.loadbytes_storebytes_other; eauto.
- unfold n2; omega.
+ { rewrite <- LB'. eapply Mem.loadbytes_storebytes_other; eauto.
+ unfold n2; omega.
right; left; omega. }
- exploit Mem.load_valid_access; eauto. intros [P Q].
+ exploit Mem.load_valid_access; eauto. intros [P Q].
rewrite B. apply Mem.loadbytes_load.
- replace (i + (ofs2 - ofs1)) with (ofs2 + n1) by (unfold n1; omega).
- exact LB''.
+ replace (i + (ofs2 - ofs1)) with (ofs2 + n1) by (unfold n1; omega).
+ exact LB''.
apply Z.divide_add_r; auto.
Qed.
@@ -625,7 +625,7 @@ Lemma shift_memcpy_eq_wf:
wf_equation next e'.
Proof with (try discriminate).
unfold shift_memcpy_eq; intros.
- destruct e. destruct r... destruct a...
+ destruct e. destruct r... destruct a...
destruct (zle src (Int.unsigned i) &&
zle (Int.unsigned i + size_chunk m) (src + sz) &&
zeq (delta mod align_chunk m) 0 && zle 0 (Int.unsigned i + delta) &&
@@ -642,7 +642,7 @@ Lemma shift_memcpy_eq_holds:
equation_holds valu ge (Vptr sp Int.zero) m' e'.
Proof with (try discriminate).
intros. set (delta := dst - src) in *. unfold shift_memcpy_eq in H.
- destruct e as [l strict rhs] eqn:E.
+ destruct e as [l strict rhs] eqn:E.
destruct rhs as [op vl | chunk addr vl]...
destruct addr...
set (i1 := Int.unsigned i) in *. set (j := i1 + delta) in *.
@@ -656,16 +656,16 @@ Proof with (try discriminate).
Mem.loadv chunk m (Vptr sp i) = Some v ->
Mem.loadv chunk m' (Vptr sp (Int.repr j)) = Some v).
{
- simpl; intros. rewrite Int.unsigned_repr by omega.
- unfold j, delta. eapply load_memcpy; eauto.
+ simpl; intros. rewrite Int.unsigned_repr by omega.
+ unfold j, delta. eapply load_memcpy; eauto.
apply Zmod_divide; auto. generalize (align_chunk_pos chunk); omega.
}
inv H2.
+ inv H3. destruct vl... simpl in H6. rewrite Int.add_zero_l in H6. inv H6.
- apply eq_holds_strict. econstructor. simpl. rewrite Int.add_zero_l. eauto.
+ apply eq_holds_strict. econstructor. simpl. rewrite Int.add_zero_l. eauto.
apply LD; auto.
+ inv H4. destruct vl... simpl in H7. rewrite Int.add_zero_l in H7. inv H7.
- apply eq_holds_lessdef with v; auto.
+ apply eq_holds_lessdef with v; auto.
econstructor. simpl. rewrite Int.add_zero_l. eauto. apply LD; auto.
Qed.
@@ -677,7 +677,7 @@ Proof.
induction eqs1; simpl; intros.
- auto.
- destruct (shift_memcpy_eq src sz delta a) as [e''|] eqn:SHIFT.
- + destruct H. subst e''. right; exists a; auto.
+ + destruct H. subst e''. right; exists a; auto.
destruct IHeqs1 as [A | [e [A B]]]; auto. right; exists e; auto.
+ destruct IHeqs1 as [A | [e [A B]]]; auto. right; exists e; auto.
Qed.
@@ -695,7 +695,7 @@ Lemma add_memcpy_holds:
numbering_holds valu ge (Vptr sp Int.zero) rs m' (add_memcpy n1 n2 asrc adst sz).
Proof.
intros. unfold add_memcpy.
- destruct asrc; auto; destruct adst; auto.
+ destruct asrc; auto; destruct adst; auto.
assert (A: forall b o i, pmatch bc b o (Stk i) -> b = sp /\ i = o).
{
intros. inv H7. split; auto. eapply bc_stack; eauto.
@@ -703,11 +703,11 @@ Proof.
apply A in H3; destruct H3. subst bsrc ofs.
apply A in H4; destruct H4. subst bdst ofs0.
constructor; simpl; intros; eauto with cse.
-- constructor; simpl; eauto with cse.
+- constructor; simpl; eauto with cse.
intros. exploit add_memcpy_eqs_charact; eauto. intros [X | (e0 & X & Y)].
eauto with cse.
- apply wf_equation_incr with (num_next n1); auto.
- eapply shift_memcpy_eq_wf; eauto with cse.
+ apply wf_equation_incr with (num_next n1); auto.
+ eapply shift_memcpy_eq_wf; eauto with cse.
- exploit add_memcpy_eqs_charact; eauto. intros [X | (e0 & X & Y)].
eauto with cse.
eapply shift_memcpy_eq_holds; eauto with cse.
@@ -747,7 +747,7 @@ Proof.
assert (sem op1 (map valu args1) = Some res).
rewrite <- H0. eapply f_sound; eauto.
simpl; intros.
- exploit num_holds_eq; eauto.
+ exploit num_holds_eq; eauto.
eapply find_valnum_num_charact; eauto with cse.
intros EH; inv EH; auto.
destruct (reduce_rec A f n niter op1 args1) as [[op2 rl2] | ] eqn:?.
@@ -765,7 +765,7 @@ Lemma reduce_sound:
sem op rs##rl = Some res ->
sem op' rs##rl' = Some res.
Proof.
- unfold reduce; intros.
+ unfold reduce; intros.
destruct (reduce_rec A f n 4%nat op vl) as [[op1 rl1] | ] eqn:?; inv H.
eapply reduce_rec_sound; eauto. congruence.
auto.
@@ -775,8 +775,8 @@ End REDUCE.
(** The numberings associated to each instruction by the static analysis
are inductively satisfiable, in the following sense: the numbering
- at the function entry point is satisfiable, and for any RTL execution
- from [pc] to [pc'], satisfiability at [pc] implies
+ at the function entry point is satisfiable, and for any RTL execution
+ from [pc] to [pc'], satisfiability at [pc] implies
satisfiability at [pc']. *)
Theorem analysis_correct_1:
@@ -797,7 +797,7 @@ Theorem analysis_correct_entry:
analyze f vapprox = Some approx ->
exists valu, numbering_holds valu ge sp rs m approx!!(f.(fn_entrypoint)).
Proof.
- intros.
+ intros.
replace (approx!!(f.(fn_entrypoint))) with Solver.L.top.
exists (fun v => Vundef). apply empty_numbering_holds.
symmetry. eapply Solver.fixpoint_entry; eauto.
@@ -843,7 +843,7 @@ Lemma sig_preserved:
Proof.
unfold transf_fundef; intros. destruct f; monadInv H; auto.
unfold transf_function in EQ.
- destruct (analyze f (vanalyze rm f)); try discriminate. inv EQ; auto.
+ destruct (analyze f (vanalyze rm f)); try discriminate. inv EQ; auto.
Qed.
Definition transf_function' (f: function) (approxs: PMap.t numbering) : function :=
@@ -868,7 +868,7 @@ Lemma set_reg_lessdef:
forall r v1 v2 rs1 rs2,
Val.lessdef v1 v2 -> regs_lessdef rs1 rs2 -> regs_lessdef (rs1#r <- v1) (rs2#r <- v2).
Proof.
- intros; red; intros. repeat rewrite Regmap.gsspec.
+ intros; red; intros. repeat rewrite Regmap.gsspec.
destruct (peq r0 r); auto.
Qed.
@@ -958,7 +958,7 @@ Ltac TransfInstr :=
| H1: (PTree.get ?pc ?c = Some ?instr), f: function, approx: PMap.t numbering |- _ =>
cut ((transf_function' f approx).(fn_code)!pc = Some(transf_instr approx!!pc instr));
[ simpl transf_instr
- | unfold transf_function', transf_code; simpl; rewrite PTree.gmap;
+ | unfold transf_function', transf_code; simpl; rewrite PTree.gmap;
unfold option_map; rewrite H1; reflexivity ]
end.
@@ -975,8 +975,8 @@ Proof.
(* Inop *)
- econstructor; split.
eapply exec_Inop; eauto.
- econstructor; eauto.
- eapply analysis_correct_1; eauto. simpl; auto.
+ econstructor; eauto.
+ eapply analysis_correct_1; eauto. simpl; auto.
unfold transfer; rewrite H; auto.
(* Iop *)
@@ -987,9 +987,9 @@ Proof.
econstructor; split.
eapply exec_Iop with (v := v'); eauto.
rewrite <- A. apply eval_operation_preserved. exact symbols_preserved.
- econstructor; eauto.
+ econstructor; eauto.
eapply analysis_correct_1; eauto. simpl; auto.
- unfold transfer; rewrite H.
+ unfold transfer; rewrite H.
destruct SAT as [valu NH]. eapply add_op_holds; eauto.
apply set_reg_lessdef; auto.
+ (* possibly optimized *)
@@ -998,31 +998,31 @@ Proof.
exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
destruct (find_rhs n1 (Op op vl)) as [r|] eqn:?.
* (* replaced by move *)
- exploit find_rhs_sound; eauto. intros (v' & EV & LD).
+ exploit find_rhs_sound; eauto. intros (v' & EV & LD).
assert (v' = v) by (inv EV; congruence). subst v'.
econstructor; split.
eapply exec_Iop; eauto. simpl; eauto.
- econstructor; eauto.
+ econstructor; eauto.
eapply analysis_correct_1; eauto. simpl; auto.
- unfold transfer; rewrite H.
- eapply add_op_holds; eauto.
+ unfold transfer; rewrite H.
+ eapply add_op_holds; eauto.
apply set_reg_lessdef; auto.
eapply Val.lessdef_trans; eauto.
* (* possibly simplified *)
destruct (reduce operation combine_op n1 op args vl) as [op' args'] eqn:?.
assert (RES: eval_operation ge sp op' rs##args' m = Some v).
- eapply reduce_sound with (sem := fun op vl => eval_operation ge sp op vl m); eauto.
+ eapply reduce_sound with (sem := fun op vl => eval_operation ge sp op vl m); eauto.
intros; eapply combine_op_sound; eauto.
exploit eval_operation_lessdef. eapply regs_lessdef_regs; eauto. eauto. eauto.
intros [v' [A B]].
econstructor; split.
- eapply exec_Iop with (v := v'); eauto.
+ eapply exec_Iop with (v := v'); eauto.
rewrite <- A. apply eval_operation_preserved. exact symbols_preserved.
econstructor; eauto.
eapply analysis_correct_1; eauto. simpl; auto.
- unfold transfer; rewrite H.
- eapply add_op_holds; eauto.
- apply set_reg_lessdef; auto.
+ unfold transfer; rewrite H.
+ eapply add_op_holds; eauto.
+ apply set_reg_lessdef; auto.
- (* Iload *)
destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?.
@@ -1030,31 +1030,31 @@ Proof.
exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
destruct (find_rhs n1 (Load chunk addr vl)) as [r|] eqn:?.
+ (* replaced by move *)
- exploit find_rhs_sound; eauto. intros (v' & EV & LD).
+ exploit find_rhs_sound; eauto. intros (v' & EV & LD).
assert (v' = v) by (inv EV; congruence). subst v'.
econstructor; split.
eapply exec_Iop; eauto. simpl; eauto.
- econstructor; eauto.
+ econstructor; eauto.
eapply analysis_correct_1; eauto. simpl; auto.
- unfold transfer; rewrite H.
- eapply add_load_holds; eauto.
+ unfold transfer; rewrite H.
+ eapply add_load_holds; eauto.
apply set_reg_lessdef; auto. eapply Val.lessdef_trans; eauto.
+ (* load is preserved, but addressing is possibly simplified *)
destruct (reduce addressing combine_addr n1 addr args vl) as [addr' args'] eqn:?.
assert (ADDR: eval_addressing ge sp addr' rs##args' = Some a).
- { eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto.
+ { eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto.
intros; eapply combine_addr_sound; eauto. }
exploit eval_addressing_lessdef. apply regs_lessdef_regs; eauto. eexact ADDR.
intros [a' [A B]].
assert (ADDR': eval_addressing tge sp addr' rs'##args' = Some a').
{ rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved. }
- exploit Mem.loadv_extends; eauto.
+ exploit Mem.loadv_extends; eauto.
intros [v' [X Y]].
econstructor; split.
eapply exec_Iload; eauto.
econstructor; eauto.
- eapply analysis_correct_1; eauto. simpl; auto.
- unfold transfer; rewrite H.
+ eapply analysis_correct_1; eauto. simpl; auto.
+ unfold transfer; rewrite H.
eapply add_load_holds; eauto.
apply set_reg_lessdef; auto.
@@ -1064,7 +1064,7 @@ Proof.
exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
destruct (reduce addressing combine_addr n1 addr args vl) as [addr' args'] eqn:?.
assert (ADDR: eval_addressing ge sp addr' rs##args' = Some a).
- { eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto.
+ { eapply reduce_sound with (sem := fun addr vl => eval_addressing ge sp addr vl); eauto.
intros; eapply combine_addr_sound; eauto. }
exploit eval_addressing_lessdef. apply regs_lessdef_regs; eauto. eexact ADDR.
intros [a' [A B]].
@@ -1074,35 +1074,35 @@ Proof.
econstructor; split.
eapply exec_Istore; eauto.
econstructor; eauto.
- eapply analysis_correct_1; eauto. simpl; auto.
+ eapply analysis_correct_1; eauto. simpl; auto.
unfold transfer; rewrite H.
inv SOUND.
- eapply add_store_result_hold; eauto.
+ eapply add_store_result_hold; eauto.
eapply kill_loads_after_store_holds; eauto.
- (* Icall *)
- exploit find_function_translated; eauto. intros [tf [FIND' TRANSF']].
+ exploit find_function_translated; eauto. intros [tf [FIND' TRANSF']].
econstructor; split.
eapply exec_Icall; eauto.
apply sig_preserved; auto.
- econstructor; eauto.
- econstructor; eauto.
- intros. eapply analysis_correct_1; eauto. simpl; auto.
- unfold transfer; rewrite H.
+ econstructor; eauto.
+ econstructor; eauto.
+ intros. eapply analysis_correct_1; eauto. simpl; auto.
+ unfold transfer; rewrite H.
exists (fun _ => Vundef); apply empty_numbering_holds.
apply regs_lessdef_regs; auto.
- (* Itailcall *)
- exploit find_function_translated; eauto. intros [tf [FIND' TRANSF']].
+ exploit find_function_translated; eauto. intros [tf [FIND' TRANSF']].
exploit Mem.free_parallel_extends; eauto. intros [m'' [A B]].
econstructor; split.
eapply exec_Itailcall; eauto.
apply sig_preserved; auto.
- econstructor; eauto.
+ econstructor; eauto.
apply regs_lessdef_regs; auto.
- (* Ibuiltin *)
- exploit (@eval_builtin_args_lessdef _ ge (fun r => rs#r) (fun r => rs'#r)); eauto.
+ exploit (@eval_builtin_args_lessdef _ ge (fun r => rs#r) (fun r => rs'#r)); eauto.
intros (vargs' & A & B).
exploit external_call_mem_extends; eauto.
intros (v' & m1' & P & Q & R & S).
@@ -1124,7 +1124,7 @@ Proof.
{ exists valu. apply set_res_unknown_holds. eapply kill_all_loads_hold; eauto. }
destruct ef.
+ apply CASE1.
- + apply CASE3.
+ + apply CASE3.
+ apply CASE2; inv H1; auto.
+ apply CASE3.
+ apply CASE1.
@@ -1133,15 +1133,15 @@ Proof.
simpl in H1. inv H1.
exists valu.
apply set_res_unknown_holds.
- inv SOUND. unfold vanalyze, rm; rewrite AN.
+ inv SOUND. unfold vanalyze, rm; rewrite AN.
assert (pmatch bc bsrc osrc (aaddr_arg (VA.State ae am) a0))
- by (eapply aaddr_arg_sound_1; eauto).
+ by (eapply aaddr_arg_sound_1; eauto).
assert (pmatch bc bdst odst (aaddr_arg (VA.State ae am) a1))
- by (eapply aaddr_arg_sound_1; eauto).
- eapply add_memcpy_holds; eauto.
- eapply kill_loads_after_storebytes_holds; eauto.
- eapply Mem.loadbytes_length; eauto.
- simpl. apply Ple_refl.
+ by (eapply aaddr_arg_sound_1; eauto).
+ eapply add_memcpy_holds; eauto.
+ eapply kill_loads_after_storebytes_holds; eauto.
+ eapply Mem.loadbytes_length; eauto.
+ simpl. apply Ple_refl.
+ apply CASE2; inv H1; auto.
+ apply CASE2; inv H1; auto.
+ apply CASE1.
@@ -1154,10 +1154,10 @@ Proof.
exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
destruct (reduce condition combine_cond n1 cond args vl) as [cond' args'] eqn:?.
assert (RES: eval_condition cond' rs##args' m = Some b).
- { eapply reduce_sound with (sem := fun cond vl => eval_condition cond vl m); eauto.
+ { eapply reduce_sound with (sem := fun cond vl => eval_condition cond vl m); eauto.
intros; eapply combine_cond_sound; eauto. }
econstructor; split.
- eapply exec_Icond; eauto.
+ eapply exec_Icond; eauto.
eapply eval_condition_lessdef; eauto. apply regs_lessdef_regs; auto.
econstructor; eauto.
destruct b; eapply analysis_correct_1; eauto; simpl; auto;
@@ -1166,7 +1166,7 @@ Proof.
- (* Ijumptable *)
generalize (RLD arg); rewrite H0; intro LD; inv LD.
econstructor; split.
- eapply exec_Ijumptable; eauto.
+ eapply exec_Ijumptable; eauto.
econstructor; eauto.
eapply analysis_correct_1; eauto. simpl. eapply list_nth_z_in; eauto.
unfold transfer; rewrite H; auto.
@@ -1176,21 +1176,21 @@ Proof.
econstructor; split.
eapply exec_Ireturn; eauto.
econstructor; eauto.
- destruct or; simpl; auto.
+ destruct or; simpl; auto.
- (* internal function *)
- monadInv H6. unfold transf_function in EQ.
- destruct (analyze f (vanalyze rm f)) as [approx|] eqn:?; inv EQ.
- exploit Mem.alloc_extends; eauto. apply Zle_refl. apply Zle_refl.
+ monadInv H6. unfold transf_function in EQ.
+ destruct (analyze f (vanalyze rm f)) as [approx|] eqn:?; inv EQ.
+ exploit Mem.alloc_extends; eauto. apply Zle_refl. apply Zle_refl.
intros (m'' & A & B).
econstructor; split.
- eapply exec_function_internal; simpl; eauto.
+ eapply exec_function_internal; simpl; eauto.
simpl. econstructor; eauto.
eapply analysis_correct_entry; eauto.
apply init_regs_lessdef; auto.
- (* external function *)
- monadInv H6.
+ monadInv H6.
exploit external_call_mem_extends; eauto.
intros (v' & m1' & P & Q & R & S).
econstructor; split.
@@ -1211,7 +1211,7 @@ Lemma transf_initial_states:
forall st1, initial_state prog st1 ->
exists st2, initial_state tprog st2 /\ match_states st1 st2.
Proof.
- intros. inversion H.
+ intros. inversion H.
exploit funct_ptr_translated; eauto. intros [tf [A B]].
exists (Callstate nil tf nil m0); split.
econstructor; eauto.
@@ -1224,10 +1224,10 @@ Proof.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
- intros. inv H0. inv H. inv H5. inv H3. constructor.
+ intros. inv H0. inv H. inv H5. inv H3. constructor.
Qed.
Theorem transf_program_correct:
@@ -1236,10 +1236,10 @@ Proof.
eapply forward_simulation_step with
(match_states := fun s1 s2 => sound_state prog s1 /\ match_states s1 s2).
- eexact public_preserved.
-- intros. exploit transf_initial_states; eauto. intros [s2 [A B]].
+- intros. exploit transf_initial_states; eauto. intros [s2 [A B]].
exists s2. split. auto. split. apply sound_initial; auto. auto.
- intros. destruct H. eapply transf_final_states; eauto.
-- intros. destruct H0. exploit transf_step_correct; eauto.
+- intros. destruct H0. exploit transf_step_correct; eauto.
intros [s2' [A B]]. exists s2'; split. auto. split. eapply sound_step; eauto. auto.
Qed.
diff --git a/backend/CleanupLabels.v b/backend/CleanupLabels.v
index 5eaa81e6..759201b2 100644
--- a/backend/CleanupLabels.v
+++ b/backend/CleanupLabels.v
@@ -11,12 +11,12 @@
(* *********************************************************************)
(** Removal of useless labels introduced by the linearization pass.
-
- The linearization pass introduces one label for each node of the
+
+ The linearization pass introduces one label for each node of the
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 Linear code
+ further optimizations, and to make the generated Linear code
better-looking, the present pass removes labels that cannot be
branched to. *)
diff --git a/backend/CleanupLabelsproof.v b/backend/CleanupLabelsproof.v
index 1e93dd7a..6f33c9c2 100644
--- a/backend/CleanupLabelsproof.v
+++ b/backend/CleanupLabelsproof.v
@@ -39,21 +39,21 @@ Let tge := Genv.globalenv tprog.
Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
- intros; unfold ge, tge, tprog, transf_program.
+ intros; unfold ge, tge, tprog, transf_program.
apply Genv.find_symbol_transf.
Qed.
Lemma public_preserved:
forall (s: ident), Genv.public_symbol tge s = Genv.public_symbol ge s.
Proof.
- intros; unfold ge, tge, tprog, transf_program.
+ intros; unfold ge, tge, tprog, transf_program.
apply Genv.public_symbol_transf.
Qed.
Lemma varinfo_preserved:
forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
Proof.
- intros; unfold ge, tge, tprog, transf_program.
+ intros; unfold ge, tge, tprog, transf_program.
apply Genv.find_var_info_transf.
Qed.
@@ -61,7 +61,7 @@ Lemma functions_translated:
forall (v: val) (f: fundef),
Genv.find_funct ge v = Some f ->
Genv.find_funct tge v = Some (transf_fundef f).
-Proof.
+Proof.
intros.
exact (Genv.find_funct_transf transf_fundef _ _ H).
Qed.
@@ -70,8 +70,8 @@ Lemma function_ptr_translated:
forall (b: block) (f: fundef),
Genv.find_funct_ptr ge b = Some f ->
Genv.find_funct_ptr tge b = Some (transf_fundef f).
-Proof.
- intros.
+Proof.
+ intros.
exact (Genv.find_funct_ptr_transf transf_fundef _ _ H).
Qed.
@@ -121,7 +121,7 @@ Proof.
destruct i; simpl; intros; try contradiction.
apply Labelset.add_1; auto.
apply Labelset.add_1; auto.
- revert H. induction l; simpl; intros.
+ revert H. induction l; simpl; intros.
contradiction.
destruct H. apply Labelset.add_1; auto. apply Labelset.add_2; auto.
Qed.
@@ -141,8 +141,8 @@ Proof.
In i c' -> Labelset.In lbl (fold_left add_label_branched_to c' bto)).
induction c'; simpl; intros.
contradiction.
- destruct H2.
- subst a. apply H1. apply add_label_branched_to_contains; auto.
+ destruct H2.
+ subst a. apply H1. apply add_label_branched_to_contains; auto.
apply IHc'; auto.
unfold labels_branched_to. auto.
@@ -152,7 +152,7 @@ Qed.
Lemma remove_unused_labels_cons:
forall bto i c,
- remove_unused_labels bto (i :: c) =
+ remove_unused_labels bto (i :: c) =
match i with
| Llabel lbl =>
if Labelset.mem lbl bto then i :: remove_unused_labels bto c else remove_unused_labels bto c
@@ -160,7 +160,7 @@ Lemma remove_unused_labels_cons:
i :: remove_unused_labels bto c
end.
Proof.
- unfold remove_unused_labels; intros. rewrite list_fold_right_eq. auto.
+ unfold remove_unused_labels; intros. rewrite list_fold_right_eq. auto.
Qed.
@@ -176,9 +176,9 @@ Proof.
rewrite remove_unused_labels_cons.
unfold is_label in H0. destruct a; simpl; auto.
destruct (peq lbl l). subst l. inv H0.
- rewrite Labelset.mem_1; auto.
+ rewrite Labelset.mem_1; auto.
simpl. rewrite peq_true. auto.
- destruct (Labelset.mem l bto); auto. simpl. rewrite peq_false; auto.
+ destruct (Labelset.mem l bto); auto. simpl. rewrite peq_false; auto.
Qed.
Corollary find_label_translated:
@@ -189,8 +189,8 @@ Corollary find_label_translated:
find_label lbl (fn_code (transf_function f)) =
Some (remove_unused_labels (labels_branched_to (fn_code f)) c).
Proof.
- intros. unfold transf_function; unfold cleanup_labels; simpl.
- apply find_label_commut. eapply labels_branched_to_correct; eauto.
+ intros. unfold transf_function; unfold cleanup_labels; simpl.
+ apply find_label_commut. eapply labels_branched_to_correct; eauto.
apply H; auto with coqlib.
auto.
Qed.
@@ -211,7 +211,7 @@ Inductive match_stackframes: stackframe -> stackframe -> Prop :=
incl c f.(fn_code) ->
match_stackframes
(Stackframe f sp ls c)
- (Stackframe (transf_function f) sp ls
+ (Stackframe (transf_function f) sp ls
(remove_unused_labels (labels_branched_to f.(fn_code)) c)).
Inductive match_states: state -> state -> Prop :=
@@ -252,14 +252,14 @@ Theorem transf_step_correct:
(exists s2', step tge s1' t s2' /\ match_states s2 s2')
\/ (measure s2 < measure s1 /\ t = E0 /\ match_states s2 s1')%nat.
Proof.
- induction 1; intros; inv MS; try rewrite remove_unused_labels_cons.
+ induction 1; intros; inv MS; try rewrite remove_unused_labels_cons.
(* Lgetstack *)
left; econstructor; split.
- econstructor; eauto.
+ econstructor; eauto.
econstructor; eauto with coqlib.
(* Lsetstack *)
left; econstructor; split.
- econstructor; eauto.
+ econstructor; eauto.
econstructor; eauto with coqlib.
(* Lop *)
left; econstructor; split.
@@ -270,7 +270,7 @@ Proof.
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.
econstructor; eauto with coqlib.
(* Lstore *)
assert (eval_addressing tge sp addr (LTL.reglist rs args) = Some a).
@@ -280,14 +280,14 @@ Proof.
econstructor; eauto with coqlib.
(* Lcall *)
left; econstructor; split.
- econstructor. eapply find_function_translated; eauto.
+ econstructor. eapply find_function_translated; eauto.
symmetry; apply sig_function_translated.
econstructor; eauto. constructor; auto. constructor; eauto with coqlib.
(* Ltailcall *)
left; econstructor; split.
- econstructor. erewrite match_parent_locset; eauto. eapply find_function_translated; eauto.
+ econstructor. erewrite match_parent_locset; eauto. eapply find_function_translated; eauto.
symmetry; apply sig_function_translated.
- simpl. eauto.
+ simpl. eauto.
econstructor; eauto.
(* Lbuiltin *)
left; econstructor; split.
@@ -307,11 +307,11 @@ Proof.
right. split. simpl. omega. split. auto. econstructor; eauto with coqlib.
(* Lgoto *)
left; econstructor; split.
- econstructor. eapply find_label_translated; eauto. red; auto.
+ econstructor. eapply find_label_translated; eauto. red; auto.
econstructor; eauto. eapply find_label_incl; eauto.
(* Lcond taken *)
left; econstructor; split.
- econstructor. auto. eauto. 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.
@@ -319,8 +319,8 @@ Proof.
econstructor; eauto with coqlib.
(* Ljumptable *)
left; econstructor; split.
- econstructor. eauto. eauto. eapply find_label_translated; eauto.
- red. eapply list_nth_z_in; eauto. eauto.
+ econstructor. eauto. eauto. eapply find_label_translated; eauto.
+ red. eapply list_nth_z_in; eauto. eauto.
econstructor; eauto. eapply find_label_incl; eauto.
(* Lreturn *)
left; econstructor; split.
@@ -329,7 +329,7 @@ Proof.
econstructor; eauto with coqlib.
(* internal function *)
left; econstructor; split.
- econstructor; simpl; eauto.
+ econstructor; simpl; eauto.
econstructor; eauto with coqlib.
(* external function *)
left; econstructor; split.
@@ -338,7 +338,7 @@ Proof.
econstructor; eauto with coqlib.
(* return *)
inv H3. inv H1. left; econstructor; split.
- econstructor; eauto.
+ econstructor; eauto.
econstructor; eauto.
Qed.
@@ -349,7 +349,7 @@ Proof.
intros. inv H.
econstructor; split.
eapply initial_state_intro with (f := transf_fundef f).
- eapply Genv.init_mem_transf; eauto.
+ eapply Genv.init_mem_transf; eauto.
rewrite symbols_preserved; eauto.
apply function_ptr_translated; auto.
rewrite sig_function_translated. auto.
@@ -357,7 +357,7 @@ Proof.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
intros. inv H0. inv H. inv H6. econstructor; eauto.
diff --git a/backend/Cminor.v b/backend/Cminor.v
index 7ac23bfa..0d959531 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -241,7 +241,7 @@ Inductive state: Type :=
(k: cont) (**r what to do next *)
(m: mem), (**r memory state *)
state.
-
+
Section RELSEM.
Variable ge: genv.
@@ -378,7 +378,7 @@ Inductive eval_expr: expr -> val -> Prop :=
eval_expr (Ebinop op a1 a2) v
| eval_Eload: forall chunk addr vaddr v,
eval_expr addr vaddr ->
- Mem.loadv chunk m vaddr = Some v ->
+ Mem.loadv chunk m vaddr = Some v ->
eval_expr (Eload chunk addr) v.
Inductive eval_exprlist: list expr -> list val -> Prop :=
@@ -406,10 +406,10 @@ Definition is_call_cont (k: cont) : Prop :=
| _ => False
end.
-(** Find the statement and manufacture the continuation
+(** Find the statement and manufacture the continuation
corresponding to a label *)
-Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
+Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
{struct s}: option (stmt * cont) :=
match s with
| Sseq s1 s2 =>
@@ -543,7 +543,7 @@ Inductive step: state -> trace -> state -> Prop :=
| step_external_function: forall ef vargs k m t vres m',
external_call ef ge vargs m t vres m' ->
step (Callstate (External ef) vargs k m)
- t (Returnstate vres k m')
+ t (Returnstate vres k m')
| step_return: forall v optid f sp e k m,
step (Returnstate v (Kcall optid f sp e k) m)
@@ -586,9 +586,9 @@ Proof.
assert (t1 = E0 -> exists s2, step (Genv.globalenv p) s t2 s2).
intros. subst. inv H0. exists s1; auto.
inversion H; subst; auto.
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
- exists (State f Sskip k sp (set_optvar optid vres2 e) m2). econstructor; eauto.
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exists (State f Sskip k sp (set_optvar optid vres2 e) m2). econstructor; eauto.
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate vres2 k m2). econstructor; eauto.
(* trace length *)
red; intros; inv H; simpl; try omega; eapply external_call_trace_length; eauto.
@@ -598,7 +598,7 @@ Qed.
(** We now define another semantics for Cminor without [goto] that follows
the ``big-step'' style of semantics, also known as natural semantics.
- In this style, just like expressions evaluate to values,
+ In this style, just like expressions evaluate to values,
statements evaluate to``outcomes'' indicating how execution should
proceed afterwards. *)
@@ -758,7 +758,7 @@ with exec_stmt:
Scheme eval_funcall_ind2 := Minimality for eval_funcall Sort Prop
with exec_stmt_ind2 := Minimality for exec_stmt Sort Prop.
-Combined Scheme eval_funcall_exec_stmt_ind2
+Combined Scheme eval_funcall_exec_stmt_ind2
from eval_funcall_ind2, exec_stmt_ind2.
(** Coinductive semantics for divergence.
@@ -871,22 +871,22 @@ Inductive outcome_state_match
(sp: val) (e: env) (m: mem) (f: function) (k: cont):
outcome -> state -> Prop :=
| osm_normal:
- outcome_state_match sp e m f k
+ outcome_state_match sp e m f k
Out_normal
(State f Sskip k sp e m)
| osm_exit: forall n,
- outcome_state_match sp e m f k
+ outcome_state_match sp e m f k
(Out_exit n)
(State f (Sexit n) k sp e m)
| osm_return_none: forall k',
call_cont k' = call_cont k ->
- outcome_state_match sp e m f k
+ outcome_state_match sp e m f k
(Out_return None)
(State f (Sreturn None) k' sp e m)
| osm_return_some: forall k' a v,
call_cont k' = call_cont k ->
eval_expr ge sp e m a v ->
- outcome_state_match sp e m f k
+ outcome_state_match sp e m f k
(Out_return (Some v))
(State f (Sreturn (Some a)) k' sp e m)
| osm_tail: forall v,
@@ -925,11 +925,11 @@ Proof.
(* funcall internal *)
destruct (H2 k) as [S [A B]].
assert (call_cont k = k) by (apply call_cont_is_call_cont; auto).
- eapply star_left. econstructor; eauto.
+ eapply star_left. econstructor; eauto.
eapply star_trans. eexact A.
inversion B; clear B; subst out; simpl in H3; simpl; try contradiction.
(* Out normal *)
- subst vres. apply star_one. apply step_skip_call; auto.
+ subst vres. apply star_one. apply step_skip_call; auto.
(* Out_return None *)
subst vres. replace k with (call_cont k') by congruence.
apply star_one. apply step_return_0; auto.
@@ -943,11 +943,11 @@ Proof.
reflexivity. traceEq.
(* funcall external *)
- apply star_one. constructor; auto.
+ apply star_one. constructor; auto.
(* skip *)
econstructor; split.
- apply star_refl.
+ apply star_refl.
constructor.
(* assign *)
@@ -962,14 +962,14 @@ Proof.
(* call *)
econstructor; split.
- eapply star_left. econstructor; eauto.
- eapply star_right. apply H4. red; auto.
+ eapply star_left. econstructor; eauto.
+ eapply star_right. apply H4. red; auto.
constructor. reflexivity. traceEq.
subst e'. constructor.
(* builtin *)
econstructor; split.
- apply star_one. econstructor; eauto.
+ apply star_one. econstructor; eauto.
subst e'. constructor.
(* ifthenelse *)
@@ -985,8 +985,8 @@ Proof.
destruct (H2 k) as [S2 [A2 B2]].
inv B1.
exists S2; split.
- eapply star_left. constructor.
- eapply star_trans. eexact A1.
+ eapply star_left. constructor.
+ eapply star_trans. eexact A1.
eapply star_left. constructor. eexact A2.
reflexivity. reflexivity. traceEq.
auto.
@@ -1010,8 +1010,8 @@ Proof.
destruct (H2 k) as [S2 [A2 B2]].
inv B1.
exists S2; split.
- eapply star_left. constructor.
- eapply star_trans. eexact A1.
+ eapply star_left. constructor.
+ eapply star_trans. eexact A1.
eapply star_left. constructor. eexact A2.
reflexivity. reflexivity. traceEq.
auto.
@@ -1063,9 +1063,9 @@ Proof.
(* tailcall *)
econstructor; split.
- eapply star_left. econstructor; eauto.
+ eapply star_left. econstructor; eauto.
apply H5. apply is_call_cont_call_cont. traceEq.
- econstructor.
+ econstructor.
Qed.
Lemma eval_funcall_steps:
@@ -1100,12 +1100,12 @@ Proof.
(* call *)
eapply forever_plus_intro.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
apply CIH_FUN. eauto. traceEq.
(* ifthenelse *)
eapply forever_plus_intro with (s2 := State f (if b then s1 else s2) k sp e m).
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
apply CIH_STMT. eauto. traceEq.
(* seq 1 *)
@@ -1118,9 +1118,9 @@ Proof.
as [S [A B]]. inv B.
eapply forever_plus_intro.
eapply plus_left. constructor.
- eapply star_right. eexact A. constructor.
+ eapply star_right. eexact A. constructor.
reflexivity. reflexivity.
- apply CIH_STMT. eauto. traceEq.
+ apply CIH_STMT. eauto. traceEq.
(* loop body *)
eapply forever_plus_intro.
@@ -1132,7 +1132,7 @@ Proof.
as [S [A B]]. inv B.
eapply forever_plus_intro.
eapply plus_left. constructor.
- eapply star_right. eexact A. constructor.
+ eapply star_right. eexact A. constructor.
reflexivity. reflexivity.
apply CIH_STMT. eauto. traceEq.
@@ -1143,14 +1143,14 @@ Proof.
(* tailcall *)
eapply forever_plus_intro.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
apply CIH_FUN. eauto. traceEq.
(* function call *)
intros. inv H0.
eapply forever_plus_intro.
apply plus_one. econstructor; eauto.
- apply H. eauto.
+ apply H. eauto.
traceEq.
Qed.
@@ -1160,12 +1160,12 @@ Proof.
constructor; intros.
(* termination *)
inv H. econstructor; econstructor.
- split. econstructor; eauto.
- split. apply eval_funcall_steps. eauto. red; auto.
+ split. econstructor; eauto.
+ split. apply eval_funcall_steps. eauto. red; auto.
econstructor.
(* divergence *)
inv H. econstructor.
- split. econstructor; eauto.
+ split. econstructor; eauto.
eapply forever_plus_forever.
eapply evalinf_funcall_forever; eauto.
Qed.
diff --git a/backend/CminorSel.v b/backend/CminorSel.v
index 6a43eccd..d654502b 100644
--- a/backend/CminorSel.v
+++ b/backend/CminorSel.v
@@ -174,7 +174,7 @@ Inductive eval_expr: letenv -> expr -> val -> Prop :=
| eval_Eload: forall le chunk addr al vl vaddr v,
eval_exprlist le al vl ->
eval_addressing ge sp addr vl = Some vaddr ->
- Mem.loadv chunk m vaddr = Some v ->
+ Mem.loadv chunk m vaddr = Some v ->
eval_expr le (Eload chunk addr al) v
| eval_Econdition: forall le a b c va v,
eval_condexpr le a va ->
@@ -300,10 +300,10 @@ Definition is_call_cont (k: cont) : Prop :=
| _ => False
end.
-(** Find the statement and manufacture the continuation
+(** Find the statement and manufacture the continuation
corresponding to a label *)
-Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
+Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
{struct s}: option (stmt * cont) :=
match s with
| Sseq s1 s2 =>
@@ -436,7 +436,7 @@ Inductive step: state -> trace -> state -> Prop :=
| step_external_function: forall ef vargs k m t vres m',
external_call ef ge vargs m t vres m' ->
step (Callstate (External ef) vargs k m)
- t (Returnstate vres k m')
+ t (Returnstate vres k m')
| step_return: forall v optid f sp e k m,
step (Returnstate v (Kcall optid f sp e k) m)
@@ -519,7 +519,7 @@ Lemma insert_lenv_lookup1:
Proof.
induction 1; intros.
omegaContradiction.
- destruct n; simpl; simpl in H0. auto.
+ destruct n; simpl; simpl in H0. auto.
apply IHinsert_lenv. auto. omega.
Qed.
@@ -532,7 +532,7 @@ Lemma insert_lenv_lookup2:
Proof.
induction 1; intros.
simpl. assumption.
- simpl. destruct n. omegaContradiction.
+ simpl. destruct n. omegaContradiction.
apply IHinsert_lenv. exact H0. omega.
Qed.
@@ -559,7 +559,7 @@ Proof.
eapply eval_Elet. eauto. apply H2. apply insert_lenv_S; auto.
- case (le_gt_dec p n); intro.
+ case (le_gt_dec p n); intro.
apply eval_Eletvar. eapply insert_lenv_lookup2; eauto.
apply eval_Eletvar. eapply insert_lenv_lookup1; eauto.
@@ -573,7 +573,7 @@ Lemma eval_lift:
eval_expr ge sp e m (w::le) (lift a) v.
Proof.
intros. unfold lift. eapply eval_lift_expr.
- eexact H. apply insert_lenv_0.
+ eexact H. apply insert_lenv_0.
Qed.
Hint Resolve eval_lift: evalexpr.
diff --git a/backend/Constprop.v b/backend/Constprop.v
index 8f4cb76d..5ca69183 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -183,12 +183,12 @@ Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem)
Iop cop nil dst s
| None =>
let (addr', args') := addr_strength_reduction addr args aargs in
- Iload chunk addr' args' dst s
+ Iload chunk addr' args' dst s
end
| Istore chunk addr args src s =>
let aargs := aregs ae args in
let (addr', args') := addr_strength_reduction addr args aargs in
- Istore chunk addr' args' src s
+ Istore chunk addr' args' src s
| Icall sig ros args res s =>
Icall sig (transf_ros ae ros) args res s
| Itailcall sig ros args =>
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index eafefed5..ad9068ab 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -50,21 +50,21 @@ Let rm := romem_for_program prog.
Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
- intros; unfold ge, tge, tprog, transf_program.
+ intros; unfold ge, tge, tprog, transf_program.
apply Genv.find_symbol_transf.
Qed.
Lemma public_preserved:
forall (s: ident), Genv.public_symbol tge s = Genv.public_symbol ge s.
Proof.
- intros; unfold ge, tge, tprog, transf_program.
+ intros; unfold ge, tge, tprog, transf_program.
apply Genv.public_symbol_transf.
Qed.
Lemma varinfo_preserved:
forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
Proof.
- intros; unfold ge, tge, tprog, transf_program.
+ intros; unfold ge, tge, tprog, transf_program.
apply Genv.find_var_info_transf.
Qed.
@@ -72,7 +72,7 @@ Lemma functions_translated:
forall (v: val) (f: fundef),
Genv.find_funct ge v = Some f ->
Genv.find_funct tge v = Some (transf_fundef rm f).
-Proof.
+Proof.
intros.
exact (Genv.find_funct_transf (transf_fundef rm) _ _ H).
Qed.
@@ -81,8 +81,8 @@ Lemma function_ptr_translated:
forall (b: block) (f: fundef),
Genv.find_funct_ptr ge b = Some f ->
Genv.find_funct_ptr tge b = Some (transf_fundef rm f).
-Proof.
- intros.
+Proof.
+ intros.
exact (Genv.find_funct_ptr_transf (transf_fundef rm) _ _ H).
Qed.
@@ -117,19 +117,19 @@ Proof.
generalize (EM r); fold (areg ae r); intro VM. generalize (RLD r); intro LD.
assert (DEFAULT: find_function tge (inl _ r) rs' = Some (transf_fundef rm f)).
{
- simpl. inv LD. apply functions_translated; auto. rewrite <- H0 in FF; discriminate.
+ simpl. inv LD. apply functions_translated; auto. rewrite <- H0 in FF; discriminate.
}
- destruct (areg ae r); auto. destruct p; auto.
- predSpec Int.eq Int.eq_spec ofs Int.zero; intros; auto.
+ destruct (areg ae r); auto. destruct p; auto.
+ predSpec Int.eq Int.eq_spec ofs Int.zero; intros; auto.
subst ofs. exploit vmatch_ptr_gl; eauto. intros LD'. inv LD'; try discriminate.
- rewrite H1 in FF. unfold Genv.symbol_address in FF.
+ rewrite H1 in FF. unfold Genv.symbol_address in FF.
simpl. rewrite symbols_preserved.
destruct (Genv.find_symbol ge id) as [b|]; try discriminate.
simpl in FF. rewrite dec_eq_true in FF.
apply function_ptr_translated; auto.
rewrite <- H0 in FF; discriminate.
- (* function symbol *)
- rewrite symbols_preserved.
+ rewrite symbols_preserved.
destruct (Genv.find_symbol ge i) as [b|]; try discriminate.
apply function_ptr_translated; auto.
Qed.
@@ -155,11 +155,11 @@ Proof.
+ (* global *)
inv H. exists (Genv.symbol_address ge id ofs); split.
unfold Genv.symbol_address. rewrite <- symbols_preserved. reflexivity.
- eapply vmatch_ptr_gl; eauto.
+ eapply vmatch_ptr_gl; eauto.
+ (* stack *)
- inv H. exists (Vptr sp ofs); split.
- simpl; rewrite Int.add_zero_l; auto.
- eapply vmatch_ptr_stk; eauto.
+ inv H. exists (Vptr sp ofs); split.
+ simpl; rewrite Int.add_zero_l; auto.
+ eapply vmatch_ptr_stk; eauto.
Qed.
Inductive match_pc (f: function) (ae: AE.t): nat -> node -> node -> Prop :=
@@ -200,14 +200,14 @@ Lemma builtin_arg_reduction_correct:
eval_builtin_arg ge (fun r => rs#r) sp m (builtin_arg_reduction ae a) v.
Proof.
induction 2; simpl; eauto with barg.
-- specialize (H x). unfold areg. destruct (AE.get x ae); try constructor.
+- specialize (H x). unfold areg. destruct (AE.get x ae); try constructor.
+ inv H. constructor.
+ inv H. constructor.
+ destruct (Compopts.generate_float_constants tt); [inv H|idtac]; constructor.
+ destruct (Compopts.generate_float_constants tt); [inv H|idtac]; constructor.
- destruct (builtin_arg_reduction ae hi); auto with barg.
destruct (builtin_arg_reduction ae lo); auto with barg.
- inv IHeval_builtin_arg1; inv IHeval_builtin_arg2. constructor.
+ inv IHeval_builtin_arg1; inv IHeval_builtin_arg2. constructor.
Qed.
Lemma builtin_arg_strength_reduction_correct:
@@ -216,7 +216,7 @@ Lemma builtin_arg_strength_reduction_correct:
eval_builtin_arg ge (fun r => rs#r) sp m a v ->
eval_builtin_arg ge (fun r => rs#r) sp m (builtin_arg_strength_reduction ae a c) v.
Proof.
- intros. unfold builtin_arg_strength_reduction.
+ intros. unfold builtin_arg_strength_reduction.
destruct (builtin_arg_ok (builtin_arg_reduction ae a) c).
eapply builtin_arg_reduction_correct; eauto.
auto.
@@ -231,7 +231,7 @@ Lemma builtin_args_strength_reduction_correct:
Proof.
induction 2; simpl; constructor.
eapply builtin_arg_strength_reduction_correct; eauto.
- apply IHlist_forall2.
+ apply IHlist_forall2.
Qed.
Lemma debug_strength_reduction_correct:
@@ -247,7 +247,7 @@ Proof.
(a1 :: debug_strength_reduction ae al) (b1 :: vl'))
by (constructor; eauto).
destruct a1; try (econstructor; eassumption).
- destruct (builtin_arg_reduction ae (BA x)); repeat (eauto; econstructor).
+ destruct (builtin_arg_reduction ae (BA x)); repeat (eauto; econstructor).
Qed.
Lemma builtin_strength_reduction_correct:
@@ -259,7 +259,7 @@ Lemma builtin_strength_reduction_correct:
eval_builtin_args ge (fun r => rs#r) sp m (builtin_strength_reduction ae ef args) vargs'
/\ external_call ef ge vargs' m t vres m'.
Proof.
- intros.
+ intros.
assert (DEFAULT: forall cl,
exists vargs',
eval_builtin_args ge (fun r => rs#r) sp m (builtin_args_strength_reduction ae args cl) vargs'
@@ -267,8 +267,8 @@ Proof.
{ exists vargs; split; auto. eapply builtin_args_strength_reduction_correct; eauto. }
unfold builtin_strength_reduction.
destruct ef; auto.
- exploit debug_strength_reduction_correct; eauto. intros (vargs' & P).
- exists vargs'; split; auto.
+ exploit debug_strength_reduction_correct; eauto. intros (vargs' & P).
+ exists vargs'; split; auto.
inv H1; constructor.
Qed.
@@ -341,8 +341,8 @@ Lemma match_states_succ:
match_states O (State s f sp pc rs m)
(State s' (transf_function rm f) sp pc rs' m').
Proof.
- intros. inv H.
- apply match_states_intro with (bc := bc) (ae := ae); auto.
+ intros. inv H.
+ apply match_states_intro with (bc := bc) (ae := ae); auto.
constructor.
Qed.
@@ -351,7 +351,7 @@ Lemma transf_instr_at:
f.(fn_code)!pc = Some i ->
(transf_function rm f).(fn_code)!pc = Some(transf_instr f (analyze rm f) rm pc i).
Proof.
- intros. simpl. rewrite PTree.gmap. rewrite H. auto.
+ intros. simpl. rewrite PTree.gmap. rewrite H. auto.
Qed.
Ltac TransfInstr :=
@@ -374,7 +374,7 @@ Proof.
induction 1; intros; inv SS1; inv MS; try (inv PC; try congruence).
(* Inop, preserved *)
- rename pc'0 into pc. TransfInstr; intros.
+ rename pc'0 into pc. TransfInstr; intros.
left; econstructor; econstructor; split.
eapply exec_Inop; eauto.
eapply match_states_succ; eauto.
@@ -389,14 +389,14 @@ Proof.
set (a := eval_static_operation op (aregs ae args)).
set (ae' := AE.set res a ae).
assert (VMATCH: vmatch bc v a) by (eapply eval_static_operation_sound; eauto with va).
- assert (MATCH': ematch bc (rs#res <- v) ae') by (eapply ematch_update; eauto).
+ assert (MATCH': ematch bc (rs#res <- v) ae') by (eapply ematch_update; eauto).
destruct (const_for_result a) as [cop|] eqn:?; intros.
(* constant is propagated *)
exploit const_for_result_correct; eauto. intros (v' & A & B).
left; econstructor; econstructor; split.
- eapply exec_Iop; eauto.
+ eapply exec_Iop; eauto.
apply match_states_intro with bc ae'; auto.
- apply match_successor.
+ apply match_successor.
apply set_reg_lessdef; auto.
(* operator is strength-reduced *)
assert(OP:
@@ -421,8 +421,8 @@ Proof.
rename pc'0 into pc. TransfInstr.
set (aa := eval_static_addressing addr (aregs ae args)).
assert (VM1: vmatch bc a aa) by (eapply eval_static_addressing_sound; eauto with va).
- set (av := loadv chunk rm am aa).
- assert (VM2: vmatch bc v av) by (eapply loadv_sound; eauto).
+ set (av := loadv chunk rm am aa).
+ assert (VM2: vmatch bc v av) by (eapply loadv_sound; eauto).
destruct (const_for_result av) as [cop|] eqn:?; intros.
(* constant-propagated *)
exploit const_for_result_correct; eauto. intros (v' & A & B).
@@ -439,7 +439,7 @@ Proof.
{ eapply addr_strength_reduction_correct with (ae := ae); eauto with va. }
destruct (addr_strength_reduction addr args (aregs ae args)) as [addr' args'].
destruct ADDR as (a' & P & Q).
- exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P.
+ exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P.
intros (a'' & U & V).
assert (W: eval_addressing tge (Vptr sp0 Int.zero) addr' rs'##args' = Some a'').
{ rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. }
@@ -459,11 +459,11 @@ Proof.
{ eapply addr_strength_reduction_correct with (ae := ae); eauto with va. }
destruct (addr_strength_reduction addr args (aregs ae args)) as [addr' args'].
destruct ADDR as (a' & P & Q).
- exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P.
+ exploit eval_addressing_lessdef. eapply regs_lessdef_regs; eauto. eexact P.
intros (a'' & U & V).
assert (W: eval_addressing tge (Vptr sp0 Int.zero) addr' rs'##args' = Some a'').
{ rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved. }
- exploit Mem.storev_extends. eauto. eauto. apply Val.lessdef_trans with a'; eauto. apply REGS.
+ exploit Mem.storev_extends. eauto. eauto. apply Val.lessdef_trans with a'; eauto. apply REGS.
intros (m2' & X & Y).
left; econstructor; econstructor; split.
eapply exec_Istore; eauto.
@@ -477,7 +477,7 @@ Proof.
eapply exec_Icall; eauto. apply sig_function_translated; auto.
constructor; auto. constructor; auto.
econstructor; eauto.
- apply regs_lessdef_regs; auto.
+ apply regs_lessdef_regs; auto.
(* Itailcall *)
exploit Mem.free_parallel_extends; eauto. intros [m2' [A B]].
@@ -485,20 +485,20 @@ Proof.
TransfInstr; intro.
left; econstructor; econstructor; split.
eapply exec_Itailcall; eauto. apply sig_function_translated; auto.
- constructor; auto.
- apply regs_lessdef_regs; auto.
+ constructor; auto.
+ apply regs_lessdef_regs; auto.
(* Ibuiltin *)
rename pc'0 into pc. clear MATCH. TransfInstr; intros.
Opaque builtin_strength_reduction.
exploit builtin_strength_reduction_correct; eauto. intros (vargs' & P & Q).
- exploit (@eval_builtin_args_lessdef _ ge (fun r => rs#r) (fun r => rs'#r)).
+ exploit (@eval_builtin_args_lessdef _ ge (fun r => rs#r) (fun r => rs'#r)).
apply REGS. eauto. eexact P.
intros (vargs'' & U & V).
exploit external_call_mem_extends; eauto.
intros [v' [m2' [A [B [C D]]]]].
left; econstructor; econstructor; split.
- eapply exec_Ibuiltin; eauto.
+ eapply exec_Ibuiltin; eauto.
eapply eval_builtin_args_preserved. eexact symbols_preserved. eauto.
eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
@@ -511,24 +511,24 @@ Opaque builtin_strength_reduction.
assert (C: cmatch (eval_condition cond rs ## args m) ac)
by (eapply eval_static_condition_sound; eauto with va).
rewrite H0 in C.
- generalize (cond_strength_reduction_correct bc ae rs m EM cond args (aregs ae args) (refl_equal _)).
+ generalize (cond_strength_reduction_correct bc ae rs m EM cond args (aregs ae args) (refl_equal _)).
destruct (cond_strength_reduction cond args (aregs ae args)) as [cond' args'].
intros EV1 TCODE.
- left; exists O; exists (State s' (transf_function rm f) (Vptr sp0 Int.zero) (if b then ifso else ifnot) rs' m'); split.
- destruct (resolve_branch ac) eqn: RB.
- assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0.
- destruct b; eapply exec_Inop; eauto.
+ left; exists O; exists (State s' (transf_function rm f) (Vptr sp0 Int.zero) (if b then ifso else ifnot) rs' m'); split.
+ destruct (resolve_branch ac) eqn: RB.
+ assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0.
+ destruct b; eapply exec_Inop; eauto.
eapply exec_Icond; eauto.
eapply eval_condition_lessdef with (vl1 := rs##args'); eauto. eapply regs_lessdef_regs; eauto. congruence.
- eapply match_states_succ; eauto.
+ eapply match_states_succ; eauto.
(* Icond, skipped over *)
- rewrite H1 in H; inv H.
+ rewrite H1 in H; inv H.
set (ac := eval_static_condition cond (aregs ae0 args)) in *.
assert (C: cmatch (eval_condition cond rs ## args m) ac)
by (eapply eval_static_condition_sound; eauto with va).
rewrite H0 in C.
- assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0.
+ assert (b0 = b) by (eapply resolve_branch_sound; eauto). subst b0.
right; exists n; split. omega. split. auto.
econstructor; eauto.
@@ -537,11 +537,11 @@ Opaque builtin_strength_reduction.
assert (A: (fn_code (transf_function rm f))!pc = Some(Ijumptable arg tbl)
\/ (fn_code (transf_function rm f))!pc = Some(Inop pc')).
{ TransfInstr.
- destruct (areg ae arg) eqn:A; auto.
- generalize (EM arg). fold (areg ae arg); rewrite A.
- intros V; inv V. replace n0 with n by congruence.
+ destruct (areg ae arg) eqn:A; auto.
+ generalize (EM arg). fold (areg ae arg); rewrite A.
+ intros V; inv V. replace n0 with n by congruence.
rewrite H1. auto. }
- assert (rs'#arg = Vint n).
+ assert (rs'#arg = Vint n).
{ generalize (REGS arg). rewrite H0. intros LD; inv LD; auto. }
left; exists O; exists (State s' (transf_function rm f) (Vptr sp0 Int.zero) pc' rs' m'); split.
destruct A. eapply exec_Ijumptable; eauto. eapply exec_Inop; eauto.
@@ -552,7 +552,7 @@ Opaque builtin_strength_reduction.
left; exists O; exists (Returnstate s' (regmap_optget or Vundef rs') m2'); split.
eapply exec_Ireturn; eauto. TransfInstr; auto.
constructor; auto.
- destruct or; simpl; auto.
+ destruct or; simpl; auto.
(* internal function *)
exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl.
@@ -564,11 +564,11 @@ Opaque builtin_strength_reduction.
left; exists O; econstructor; split.
eapply exec_function_internal; simpl; eauto.
simpl. econstructor; eauto.
- constructor.
+ constructor.
apply init_regs_lessdef; auto.
(* external function *)
- exploit external_call_mem_extends; eauto.
+ exploit external_call_mem_extends; eauto.
intros [v' [m2' [A [B [C D]]]]].
simpl. left; econstructor; econstructor; split.
eapply exec_function_external; eauto.
@@ -580,10 +580,10 @@ Opaque builtin_strength_reduction.
assert (X: exists bc ae, ematch bc (rs#res <- vres) ae).
{ inv SS2. exists bc0; exists ae; auto. }
destruct X as (bc1 & ae1 & MATCH).
- inv H4. inv H1.
+ inv H4. inv H1.
left; exists O; econstructor; split.
- eapply exec_return; eauto.
- econstructor; eauto. constructor. apply set_reg_lessdef; auto.
+ eapply exec_return; eauto.
+ econstructor; eauto. constructor. apply set_reg_lessdef; auto.
Qed.
Lemma transf_initial_states:
@@ -603,10 +603,10 @@ Proof.
Qed.
Lemma transf_final_states:
- forall n st1 st2 r,
+ forall n st1 st2 r,
match_states n st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
- intros. inv H0. inv H. inv STACKS. inv RES. constructor.
+ intros. inv H0. inv H. inv STACKS. inv RES. constructor.
Qed.
(** The preservation of the observable behavior of the program then
@@ -620,15 +620,15 @@ Proof.
(fsim_match_states := fun n s1 s2 => sound_state prog s1 /\ match_states n s1 s2).
- apply lt_wf.
- simpl; intros. exploit transf_initial_states; eauto. intros (n & st2 & A & B).
- exists n, st2; intuition. eapply sound_initial; eauto.
-- simpl; intros. destruct H. eapply transf_final_states; eauto.
+ exists n, st2; intuition. eapply sound_initial; eauto.
+- simpl; intros. destruct H. eapply transf_final_states; eauto.
- simpl; intros. destruct H0.
assert (sound_state prog s1') by (eapply sound_step; eauto).
fold ge; fold tge.
- exploit transf_step_correct; eauto.
+ exploit transf_step_correct; eauto.
intros [ [n2 [s2' [A B]]] | [n2 [A [B C]]]].
exists n2; exists s2'; split; auto. left; apply plus_one; auto.
- exists n2; exists s2; split; auto. right; split; auto. subst t; apply star_refl.
+ exists n2; exists s2; split; auto. right; split; auto. subst t; apply star_refl.
- eexact public_preserved.
Qed.
diff --git a/backend/Conventions.v b/backend/Conventions.v
index abfe4eee..69cdd07d 100644
--- a/backend/Conventions.v
+++ b/backend/Conventions.v
@@ -10,7 +10,7 @@
(* *)
(* *********************************************************************)
-(** Function calling conventions and other conventions regarding the use of
+(** Function calling conventions and other conventions regarding the use of
machine registers and stack slots. *)
Require Import Coqlib.
@@ -44,14 +44,14 @@ Lemma incoming_slot_in_parameters:
In (S Outgoing ofs ty) (loc_arguments sg).
Proof.
intros.
- unfold loc_parameters in H.
+ unfold loc_parameters 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 sl; try contradiction.
+ destruct sl; try contradiction.
inv A. auto.
-Qed.
+Qed.
(** * Tail calls *)
@@ -82,16 +82,16 @@ Proof.
intro s. unfold tailcall_is_possible, tailcall_possible.
generalize (loc_arguments s). induction l; simpl; intros.
elim H0.
- destruct a.
+ destruct a.
destruct H0. subst l0. auto. apply IHl. auto. auto. discriminate.
Qed.
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.
+ 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.
+ generalize (typesize_pos ty). omega.
Qed.
diff --git a/backend/Deadcode.v b/backend/Deadcode.v
index 9bf17d1d..fa99915d 100644
--- a/backend/Deadcode.v
+++ b/backend/Deadcode.v
@@ -168,7 +168,7 @@ Definition analyze (approx: PMap.t VA.t) (f: function): option (PMap.t NA.t) :=
(** * Part 2: the code transformation *)
-Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t)
+Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t)
(pc: node) (instr: instruction) :=
match instr with
| Iop op args res s =>
@@ -177,7 +177,7 @@ Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t)
Inop s
else if is_int_zero nres then
Iop (Ointconst Int.zero) nil res s
- else if operation_is_redundant op nres then
+ else if operation_is_redundant op nres then
match args with
| arg :: _ => Iop Omove (arg :: nil) res s
| nil => instr
diff --git a/backend/Deadcodeproof.v b/backend/Deadcodeproof.v
index a45869d7..6bbf0ae7 100644
--- a/backend/Deadcodeproof.v
+++ b/backend/Deadcodeproof.v
@@ -73,8 +73,8 @@ Lemma mextends_agree:
Proof.
intros. destruct H. destruct mext_inj. constructor; intros.
- replace ofs with (ofs + 0) by omega. eapply mi_perm; eauto. auto.
-- exploit mi_memval; eauto. unfold inject_id; eauto.
- rewrite Zplus_0_r. auto.
+- exploit mi_memval; eauto. unfold inject_id; eauto.
+ rewrite Zplus_0_r. auto.
- auto.
Qed.
@@ -84,8 +84,8 @@ Lemma magree_extends:
magree m1 m2 P -> Mem.extends m1 m2.
Proof.
intros. destruct H0. constructor; auto. constructor; unfold inject_id; intros.
-- inv H0. rewrite Zplus_0_r. eauto.
-- inv H0. apply Zdivide_0.
+- inv H0. rewrite Zplus_0_r. eauto.
+- inv H0. apply Zdivide_0.
- inv H0. rewrite Zplus_0_r. eapply ma_memval0; eauto.
Qed.
@@ -100,20 +100,20 @@ Proof.
(forall i, ofs <= i < ofs + Z.of_nat n -> memval_lessdef (ZMap.get i c1) (ZMap.get i c2)) ->
list_forall2 memval_lessdef (Mem.getN n ofs c1) (Mem.getN n ofs c2)).
{
- induction n; intros; simpl.
+ induction n; intros; simpl.
constructor.
rewrite inj_S in H. constructor.
- apply H. omega.
+ apply H. omega.
apply IHn. intros; apply H; omega.
}
Local Transparent Mem.loadbytes.
- unfold Mem.loadbytes; intros. destruct H.
+ unfold Mem.loadbytes; intros. destruct H.
destruct (Mem.range_perm_dec m1 b ofs (ofs + n) Cur Readable); inv H0.
rewrite pred_dec_true. econstructor; split; eauto.
apply GETN. intros. rewrite nat_of_Z_max in H.
- assert (ofs <= i < ofs + n) by xomega.
+ assert (ofs <= i < ofs + n) by xomega.
apply ma_memval0; auto.
- red; intros; eauto.
+ red; intros; eauto.
Qed.
Lemma magree_load:
@@ -123,12 +123,12 @@ Lemma magree_load:
(forall i, ofs <= i < ofs + size_chunk chunk -> P b i) ->
exists v', Mem.load chunk m2 b ofs = Some v' /\ Val.lessdef v v'.
Proof.
- intros. exploit Mem.load_valid_access; eauto. intros [A B].
+ intros. exploit Mem.load_valid_access; eauto. intros [A B].
exploit Mem.load_loadbytes; eauto. intros [bytes [C D]].
exploit magree_loadbytes; eauto. intros [bytes' [E F]].
- exists (decode_val chunk bytes'); split.
- apply Mem.loadbytes_load; auto.
- apply val_inject_id. subst v. apply decode_val_inject; auto.
+ exists (decode_val chunk bytes'); split.
+ apply Mem.loadbytes_load; auto.
+ apply val_inject_id. subst v. apply decode_val_inject; auto.
Qed.
Lemma magree_storebytes_parallel:
@@ -151,20 +151,20 @@ Proof.
{
induction 1; intros; simpl.
- apply H; auto. simpl. omega.
- - simpl length in H1; rewrite inj_S in H1.
- apply IHlist_forall2; auto.
- intros. rewrite ! ZMap.gsspec. destruct (ZIndexed.eq i p). auto.
- apply H1; auto. unfold ZIndexed.t in *; omega.
+ - simpl length in H1; rewrite inj_S in H1.
+ apply IHlist_forall2; auto.
+ intros. rewrite ! ZMap.gsspec. destruct (ZIndexed.eq i p). auto.
+ apply H1; auto. unfold ZIndexed.t in *; omega.
}
- intros.
+ intros.
destruct (Mem.range_perm_storebytes m2 b ofs bytes2) as [m2' ST2].
{ erewrite <- list_forall2_length by eauto. red; intros.
- eapply ma_perm; eauto.
+ eapply ma_perm; eauto.
eapply Mem.storebytes_range_perm; eauto. }
- exists m2'; split; auto.
+ exists m2'; split; auto.
constructor; intros.
- eapply Mem.perm_storebytes_1; eauto. eapply ma_perm; eauto.
- eapply Mem.perm_storebytes_2; eauto.
+ eapply Mem.perm_storebytes_2; eauto.
- rewrite (Mem.storebytes_mem_contents _ _ _ _ _ H0).
rewrite (Mem.storebytes_mem_contents _ _ _ _ _ ST2).
rewrite ! PMap.gsspec. destruct (peq b0 b).
@@ -175,7 +175,7 @@ Proof.
+ eapply ma_memval; eauto. eapply Mem.perm_storebytes_2; eauto. apply H1; auto.
- rewrite (Mem.nextblock_storebytes _ _ _ _ _ H0).
rewrite (Mem.nextblock_storebytes _ _ _ _ _ ST2).
- eapply ma_nextblock; eauto.
+ eapply ma_nextblock; eauto.
Qed.
Lemma magree_store_parallel:
@@ -188,16 +188,16 @@ Lemma magree_store_parallel:
P b' i) ->
exists m2', Mem.store chunk m2 b ofs v2 = Some m2' /\ magree m1' m2' Q.
Proof.
- intros.
- exploit Mem.store_valid_access_3; eauto. intros [A B].
+ intros.
+ exploit Mem.store_valid_access_3; eauto. intros [A B].
exploit Mem.store_storebytes; eauto. intros SB1.
- exploit magree_storebytes_parallel. eauto. eauto.
+ exploit magree_storebytes_parallel. eauto. eauto.
instantiate (1 := Q). intros. rewrite encode_val_length in H4.
- rewrite <- size_chunk_conv in H4. apply H2; auto.
- eapply store_argument_sound; eauto.
- intros [m2' [SB2 AG]].
+ rewrite <- size_chunk_conv in H4. apply H2; auto.
+ eapply store_argument_sound; eauto.
+ intros [m2' [SB2 AG]].
exists m2'; split; auto.
- apply Mem.storebytes_store; auto.
+ apply Mem.storebytes_store; auto.
Qed.
Lemma magree_storebytes_left:
@@ -208,15 +208,15 @@ Lemma magree_storebytes_left:
magree m1' m2 P.
Proof.
intros. constructor; intros.
-- eapply ma_perm; eauto. eapply Mem.perm_storebytes_2; eauto.
+- eapply ma_perm; eauto. eapply Mem.perm_storebytes_2; eauto.
- rewrite (Mem.storebytes_mem_contents _ _ _ _ _ H0).
rewrite PMap.gsspec. destruct (peq b0 b).
+ subst b0. rewrite Mem.setN_outside. eapply ma_memval; eauto. eapply Mem.perm_storebytes_2; eauto.
destruct (zlt ofs0 ofs); auto. destruct (zle (ofs + Z.of_nat (length bytes1)) ofs0); try omega.
- elim (H1 ofs0). omega. auto.
+ elim (H1 ofs0). omega. auto.
+ eapply ma_memval; eauto. eapply Mem.perm_storebytes_2; eauto.
- rewrite (Mem.nextblock_storebytes _ _ _ _ _ H0).
- eapply ma_nextblock; eauto.
+ eapply ma_nextblock; eauto.
Qed.
Lemma magree_store_left:
@@ -227,9 +227,9 @@ Lemma magree_store_left:
magree m1' m2 P.
Proof.
intros. eapply magree_storebytes_left; eauto.
- eapply Mem.store_storebytes; eauto.
+ eapply Mem.store_storebytes; eauto.
intros. rewrite encode_val_length in H2.
- rewrite <- size_chunk_conv in H2. apply H1; auto.
+ rewrite <- size_chunk_conv in H2. apply H1; auto.
Qed.
Lemma magree_free:
@@ -241,21 +241,21 @@ Lemma magree_free:
P b' i) ->
exists m2', Mem.free m2 b lo hi = Some m2' /\ magree m1' m2' Q.
Proof.
- intros.
+ intros.
destruct (Mem.range_perm_free m2 b lo hi) as [m2' FREE].
- red; intros. eapply ma_perm; eauto. eapply Mem.free_range_perm; eauto.
+ red; intros. eapply ma_perm; eauto. eapply Mem.free_range_perm; eauto.
exists m2'; split; auto.
constructor; intros.
- (* permissions *)
assert (Mem.perm m2 b0 ofs k p). { eapply ma_perm; eauto. eapply Mem.perm_free_3; eauto. }
exploit Mem.perm_free_inv; eauto. intros [[A B] | A]; auto.
- subst b0. eelim Mem.perm_free_2. eexact H0. eauto. eauto.
+ subst b0. eelim Mem.perm_free_2. eexact H0. eauto. eauto.
- (* contents *)
rewrite (Mem.free_result _ _ _ _ _ H0).
- rewrite (Mem.free_result _ _ _ _ _ FREE).
+ rewrite (Mem.free_result _ _ _ _ _ FREE).
simpl. eapply ma_memval; eauto. eapply Mem.perm_free_3; eauto.
apply H1; auto. destruct (eq_block b0 b); auto.
- subst b0. right. red; intros. eelim Mem.perm_free_2. eexact H0. eauto. eauto.
+ subst b0. right. red; intros. eelim Mem.perm_free_2. eexact H0. eauto. eauto.
- (* nextblock *)
rewrite (Mem.free_result _ _ _ _ _ H0).
rewrite (Mem.free_result _ _ _ _ _ FREE).
@@ -268,9 +268,9 @@ Lemma magree_valid_access:
Mem.valid_access m1 chunk b ofs p ->
Mem.valid_access m2 chunk b ofs p.
Proof.
- intros. destruct H0; split; auto.
+ intros. destruct H0; split; auto.
red; intros. eapply ma_perm; eauto.
-Qed.
+Qed.
(** * Properties of the need environment *)
@@ -278,15 +278,15 @@ Lemma add_need_all_eagree:
forall e e' r ne,
eagree e e' (add_need_all r ne) -> eagree e e' ne.
Proof.
- intros; red; intros. generalize (H r0). unfold add_need_all.
- rewrite NE.gsspec. destruct (peq r0 r); auto with na.
+ intros; red; intros. generalize (H r0). unfold add_need_all.
+ rewrite NE.gsspec. destruct (peq r0 r); auto with na.
Qed.
Lemma add_need_all_lessdef:
forall e e' r ne,
eagree e e' (add_need_all r ne) -> Val.lessdef e#r e'#r.
Proof.
- intros. generalize (H r); unfold add_need_all.
+ intros. generalize (H r); unfold add_need_all.
rewrite NE.gsspec, peq_true. auto with na.
Qed.
@@ -313,17 +313,17 @@ Lemma add_needs_all_eagree:
Proof.
induction rl; simpl; intros.
auto.
- apply IHrl. eapply add_need_all_eagree; eauto.
+ apply IHrl. eapply add_need_all_eagree; eauto.
Qed.
Lemma add_needs_all_lessdef:
forall rl e e' ne,
eagree e e' (add_needs_all rl ne) -> Val.lessdef_list e##rl e'##rl.
Proof.
- induction rl; simpl; intros.
+ induction rl; simpl; intros.
constructor.
- constructor. eapply add_need_all_lessdef; eauto.
- eapply IHrl. eapply add_need_all_eagree; eauto.
+ constructor. eapply add_need_all_lessdef; eauto.
+ eapply IHrl. eapply add_need_all_eagree; eauto.
Qed.
Lemma add_needs_eagree:
@@ -333,7 +333,7 @@ Proof.
induction rl; simpl; intros.
auto.
destruct nvl. apply add_needs_all_eagree with (a :: rl); auto.
- eapply IHrl. eapply add_need_eagree; eauto.
+ eapply IHrl. eapply add_need_eagree; eauto.
Qed.
Lemma add_needs_vagree:
@@ -344,14 +344,14 @@ Proof.
constructor.
destruct nvl.
apply vagree_lessdef_list. eapply add_needs_all_lessdef with (rl := a :: rl); eauto.
- constructor. eapply add_need_vagree; eauto.
- eapply IHrl. eapply add_need_eagree; eauto.
+ constructor. eapply add_need_vagree; eauto.
+ eapply IHrl. eapply add_need_eagree; eauto.
Qed.
Lemma add_ros_need_eagree:
forall e e' ros ne, eagree e e' (add_ros_need_all ros ne) -> eagree e e' ne.
Proof.
- intros. destruct ros; simpl in *. eapply add_need_all_eagree; eauto. auto.
+ intros. destruct ros; simpl in *. eapply add_need_all_eagree; eauto. auto.
Qed.
Hint Resolve add_need_all_eagree add_need_all_lessdef
@@ -362,13 +362,13 @@ Hint Resolve add_need_all_eagree add_need_all_lessdef
Lemma eagree_init_regs:
forall rl vl1 vl2 ne,
- Val.lessdef_list vl1 vl2 ->
+ Val.lessdef_list vl1 vl2 ->
eagree (init_regs vl1 rl) (init_regs vl2 rl) ne.
Proof.
induction rl; intros until ne; intros LD; simpl.
- red; auto with na.
-- inv LD.
- + red; auto with na.
+- inv LD.
+ + red; auto with na.
+ apply eagree_update; auto with na.
Qed.
@@ -427,8 +427,8 @@ Lemma sig_function_translated:
funsig tf = funsig f.
Proof.
intros; destruct f; monadInv H.
- unfold transf_function in EQ.
- destruct (analyze (vanalyze rm f) f); inv EQ; auto.
+ unfold transf_function in EQ.
+ destruct (analyze (vanalyze rm f) f); inv EQ; auto.
auto.
Qed.
@@ -446,14 +446,14 @@ Lemma transf_function_at:
f.(fn_code)!pc = Some instr ->
tf.(fn_code)!pc = Some(transf_instr (vanalyze rm f) an pc instr).
Proof.
- intros. unfold transf_function in H. rewrite H0 in H. inv H; simpl.
- rewrite PTree.gmap. rewrite H1; auto.
+ intros. unfold transf_function in H. rewrite H0 in H. inv H; simpl.
+ rewrite PTree.gmap. rewrite H1; auto.
Qed.
Lemma is_dead_sound_1:
forall nv, is_dead nv = true -> nv = Nothing.
Proof.
- destruct nv; simpl; congruence.
+ destruct nv; simpl; congruence.
Qed.
Lemma is_dead_sound_2:
@@ -469,7 +469,7 @@ Lemma is_int_zero_sound:
Proof.
unfold is_int_zero; destruct nv; try discriminate.
predSpec Int.eq Int.eq_spec m Int.zero; congruence.
-Qed.
+Qed.
Lemma find_function_translated:
forall ros rs fd trs ne,
@@ -551,10 +551,10 @@ Lemma match_succ_states:
match_states (State s f (Vptr sp Int.zero) pc' e m)
(State ts tf (Vptr sp Int.zero) pc' te tm).
Proof.
- intros. exploit analyze_successors; eauto. rewrite ANPC; simpl. intros [A B].
- econstructor; eauto.
- eapply eagree_ge; eauto.
- eapply magree_monotone; eauto. intros; apply B; auto.
+ intros. exploit analyze_successors; eauto. rewrite ANPC; simpl. intros [A B].
+ econstructor; eauto.
+ eapply eagree_ge; eauto.
+ eapply magree_monotone; eauto. intros; apply B; auto.
Qed.
(** Builtin arguments and results *)
@@ -565,7 +565,7 @@ Lemma eagree_set_res:
eagree e1 e2 (kill_builtin_res res ne) ->
eagree (regmap_setres res v1 e1) (regmap_setres res v2 e2) ne.
Proof.
- intros. destruct res; simpl in *; auto.
+ intros. destruct res; simpl in *; auto.
apply eagree_update; eauto. apply vagree_lessdef; auto.
Qed.
@@ -590,19 +590,19 @@ Proof.
- exists (Vlong n); intuition auto. constructor. apply vagree_same.
- exists (Vfloat n); intuition auto. constructor. apply vagree_same.
- exists (Vsingle n); intuition auto. constructor. apply vagree_same.
-- simpl in H. exploit magree_load; eauto.
+- simpl in H. exploit magree_load; eauto.
intros. eapply nlive_add; eauto with va. rewrite Int.add_zero_l in H0; auto.
intros (v' & A & B).
exists v'; intuition auto. constructor; auto. apply vagree_lessdef; auto.
- eapply magree_monotone; eauto. intros; eapply incl_nmem_add; eauto.
-- exists (Vptr sp (Int.add Int.zero ofs)); intuition auto with na. constructor.
-- unfold Senv.symbol_address in H; simpl in H.
+ eapply magree_monotone; eauto. intros; eapply incl_nmem_add; eauto.
+- exists (Vptr sp (Int.add Int.zero ofs)); intuition auto with na. constructor.
+- unfold Senv.symbol_address in H; simpl in H.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; simpl in H; try discriminate.
exploit magree_load; eauto.
- intros. eapply nlive_add; eauto. constructor. apply GM; auto.
+ intros. eapply nlive_add; eauto. constructor. apply GM; auto.
intros (v' & A & B).
exists v'; intuition auto.
- constructor. simpl. unfold Senv.symbol_address; simpl; rewrite FS; auto.
+ constructor. simpl. unfold Senv.symbol_address; simpl; rewrite FS; auto.
apply vagree_lessdef; auto.
eapply magree_monotone; eauto. intros; eapply incl_nmem_add; eauto.
- exists (Senv.symbol_address ge id ofs); intuition auto with na. constructor.
@@ -635,8 +635,8 @@ Local Opaque transfer_builtin_arg.
- inv H. exists (@nil val); intuition auto. constructor.
- destruct (transfer_builtin_arg All (ne1, nm1) a1) as [ne' nm'] eqn:TR.
exploit IHlist_forall2; eauto. intros (vs' & A1 & B1 & C1 & D1).
- exploit transfer_builtin_arg_sound; eauto. intros (v1' & A2 & B2 & C2 & D2).
- exists (v1' :: vs'); intuition auto. constructor; auto.
+ exploit transfer_builtin_arg_sound; eauto. intros (v1' & A2 & B2 & C2 & D2).
+ exists (v1' :: vs'); intuition auto. constructor; auto.
Qed.
Lemma can_eval_builtin_arg:
@@ -651,13 +651,13 @@ Proof.
Mem.loadv chunk m addr = Some v ->
exists v', Mem.loadv chunk m' addr = Some v').
{
- intros. destruct addr; simpl in H; try discriminate.
- eapply Mem.valid_access_load. eapply magree_valid_access; eauto.
+ intros. destruct addr; simpl in H; try discriminate.
+ eapply Mem.valid_access_load. eapply magree_valid_access; eauto.
eapply Mem.load_valid_access; eauto. }
induction 1; try (econstructor; now constructor).
- exploit LD; eauto. intros (v' & A). exists v'; constructor; auto.
- exploit LD; eauto. intros (v' & A). exists v'; constructor.
- unfold Senv.symbol_address, Senv.find_symbol. rewrite symbols_preserved. assumption.
+ unfold Senv.symbol_address, Senv.find_symbol. rewrite symbols_preserved. assumption.
- destruct IHeval_builtin_arg1 as (v1' & A1).
destruct IHeval_builtin_arg2 as (v2' & A2).
exists (Val.longofwords v1' v2'); constructor; auto.
@@ -692,11 +692,11 @@ Proof.
intros. inv H. split; auto.
inv H0. inv H9.
- (* volatile *)
- exists tm; split; auto. econstructor. econstructor; eauto.
+ exists tm; split; auto. econstructor. econstructor; eauto.
eapply eventval_match_lessdef; eauto. apply store_argument_load_result; auto.
- (* not volatile *)
exploit magree_store_parallel. eauto. eauto. eauto.
- instantiate (1 := nlive ge sp nm). auto.
+ instantiate (1 := nlive ge sp nm). auto.
intros (tm' & P & Q).
exists tm'; split. econstructor. econstructor; eauto. auto.
Qed.
@@ -740,7 +740,7 @@ Ltac UseTransfer :=
- (* nop *)
TransfInstr; UseTransfer.
econstructor; split.
- eapply exec_Inop; eauto.
+ eapply exec_Inop; eauto.
eapply match_succ_states; eauto. simpl; auto.
- (* op *)
@@ -750,26 +750,26 @@ Ltac UseTransfer :=
[idtac|destruct (operation_is_redundant op (nreg ne res)) eqn:REDUNDANT]].
+ (* dead instruction, turned into a nop *)
econstructor; split.
- eapply exec_Inop; eauto.
+ eapply exec_Inop; eauto.
eapply match_succ_states; eauto. simpl; auto.
- apply eagree_update_dead; auto with na.
+ apply eagree_update_dead; auto with na.
+ (* instruction with needs = [I Int.zero], turned into a load immediate of zero. *)
econstructor; split.
- eapply exec_Iop with (v := Vint Int.zero); eauto.
+ eapply exec_Iop with (v := Vint Int.zero); eauto.
eapply match_succ_states; eauto. simpl; auto.
- apply eagree_update; auto.
+ apply eagree_update; auto.
rewrite is_int_zero_sound by auto.
destruct v; simpl; auto. apply iagree_zero.
+ (* redundant operation *)
destruct args.
* (* kept as is because no arguments -- should never happen *)
- simpl in *.
- exploit needs_of_operation_sound. eapply ma_perm; eauto.
+ simpl in *.
+ exploit needs_of_operation_sound. eapply ma_perm; eauto.
eauto. instantiate (1 := nreg ne res). eauto with na. eauto with na. intros [tv [A B]].
- econstructor; split.
+ econstructor; split.
eapply exec_Iop with (v := tv); eauto.
rewrite <- A. apply eval_operation_preserved. exact symbols_preserved.
- eapply match_succ_states; eauto. simpl; auto.
+ eapply match_succ_states; eauto. simpl; auto.
apply eagree_update; auto.
* (* turned into a move *)
unfold fst in ENV. unfold snd in MEM. simpl in H0.
@@ -777,17 +777,17 @@ Ltac UseTransfer :=
{ eapply operation_is_redundant_sound with (arg1' := te#r) (args' := te##args).
eauto. eauto. exploit add_needs_vagree; eauto. }
econstructor; split.
- eapply exec_Iop; eauto. simpl; reflexivity.
- eapply match_succ_states; eauto. simpl; auto.
- eapply eagree_update; eauto 2 with na.
+ eapply exec_Iop; eauto. simpl; reflexivity.
+ eapply match_succ_states; eauto. simpl; auto.
+ eapply eagree_update; eauto 2 with na.
+ (* preserved operation *)
simpl in *.
exploit needs_of_operation_sound. eapply ma_perm; eauto. eauto. eauto 2 with na. eauto with na.
intros [tv [A B]].
- econstructor; split.
+ econstructor; split.
eapply exec_Iop with (v := tv); eauto.
rewrite <- A. apply eval_operation_preserved. exact symbols_preserved.
- eapply match_succ_states; eauto. simpl; auto.
+ eapply match_succ_states; eauto. simpl; auto.
apply eagree_update; eauto 2 with na.
- (* load *)
@@ -797,87 +797,87 @@ Ltac UseTransfer :=
simpl in *.
+ (* dead instruction, turned into a nop *)
econstructor; split.
- eapply exec_Inop; eauto.
+ eapply exec_Inop; eauto.
eapply match_succ_states; eauto. simpl; auto.
- apply eagree_update_dead; auto with na.
+ apply eagree_update_dead; auto with na.
+ (* instruction with needs = [I Int.zero], turned into a load immediate of zero. *)
econstructor; split.
- eapply exec_Iop with (v := Vint Int.zero); eauto.
+ eapply exec_Iop with (v := Vint Int.zero); eauto.
eapply match_succ_states; eauto. simpl; auto.
- apply eagree_update; auto.
+ apply eagree_update; auto.
rewrite is_int_zero_sound by auto.
destruct v; simpl; auto. apply iagree_zero.
+ (* preserved *)
- exploit eval_addressing_lessdef. eapply add_needs_all_lessdef; eauto. eauto.
+ exploit eval_addressing_lessdef. eapply add_needs_all_lessdef; eauto. eauto.
intros (ta & U & V). inv V; try discriminate.
destruct ta; simpl in H1; try discriminate.
- exploit magree_load; eauto.
+ exploit magree_load; eauto.
exploit aaddressing_sound; eauto. intros (bc & A & B & C).
- intros. apply nlive_add with bc i; assumption.
+ intros. apply nlive_add with bc i; assumption.
intros (tv & P & Q).
econstructor; split.
- eapply exec_Iload with (a := Vptr b i). eauto.
+ eapply exec_Iload with (a := Vptr b i). eauto.
rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved.
- eauto.
+ eauto.
eapply match_succ_states; eauto. simpl; auto.
- apply eagree_update; eauto 2 with na.
- eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+ apply eagree_update; eauto 2 with na.
+ eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
- (* store *)
TransfInstr; UseTransfer.
destruct (nmem_contains nm (aaddressing (vanalyze rm f) # pc addr args)
(size_chunk chunk)) eqn:CONTAINS.
+ (* preserved *)
- simpl in *.
- exploit eval_addressing_lessdef. eapply add_needs_all_lessdef; eauto. eauto.
+ simpl in *.
+ exploit eval_addressing_lessdef. eapply add_needs_all_lessdef; eauto. eauto.
intros (ta & U & V). inv V; try discriminate.
destruct ta; simpl in H1; try discriminate.
exploit magree_store_parallel. eauto. eauto. instantiate (1 := te#src). eauto with na.
- instantiate (1 := nlive ge sp0 nm).
+ instantiate (1 := nlive ge sp0 nm).
exploit aaddressing_sound; eauto. intros (bc & A & B & C).
- intros. apply nlive_remove with bc b i; assumption.
+ intros. apply nlive_remove with bc b i; assumption.
intros (tm' & P & Q).
econstructor; split.
- eapply exec_Istore with (a := Vptr b i). eauto.
+ eapply exec_Istore with (a := Vptr b i). eauto.
rewrite <- U. apply eval_addressing_preserved. exact symbols_preserved.
eauto.
eapply match_succ_states; eauto. simpl; auto.
- eauto 3 with na.
+ eauto 3 with na.
+ (* dead instruction, turned into a nop *)
destruct a; simpl in H1; try discriminate.
econstructor; split.
- eapply exec_Inop; eauto.
+ eapply exec_Inop; eauto.
eapply match_succ_states; eauto. simpl; auto.
eapply magree_store_left; eauto.
exploit aaddressing_sound; eauto. intros (bc & A & B & C).
- intros. eapply nlive_contains; eauto.
+ intros. eapply nlive_contains; eauto.
- (* call *)
TransfInstr; UseTransfer.
exploit find_function_translated; eauto 2 with na. intros (tfd & A & B).
econstructor; split.
- eapply exec_Icall; eauto. apply sig_function_translated; auto.
- constructor.
- constructor; auto. econstructor; eauto.
+ eapply exec_Icall; eauto. apply sig_function_translated; auto.
+ constructor.
+ constructor; auto. econstructor; eauto.
intros.
- edestruct analyze_successors; eauto. simpl; eauto.
- eapply eagree_ge; eauto. rewrite ANPC. simpl.
+ edestruct analyze_successors; eauto. simpl; eauto.
+ eapply eagree_ge; eauto. rewrite ANPC. simpl.
apply eagree_update; eauto with na.
- auto. eauto 2 with na. eapply magree_extends; eauto. apply nlive_all.
+ auto. eauto 2 with na. eapply magree_extends; eauto. apply nlive_all.
- (* tailcall *)
TransfInstr; UseTransfer.
exploit find_function_translated; eauto 2 with na. intros (tfd & A & B).
- exploit magree_free. eauto. eauto. instantiate (1 := nlive ge stk nmem_all).
- intros; eapply nlive_dead_stack; eauto.
- intros (tm' & C & D).
+ exploit magree_free. eauto. eauto. instantiate (1 := nlive ge stk nmem_all).
+ intros; eapply nlive_dead_stack; eauto.
+ intros (tm' & C & D).
econstructor; split.
- eapply exec_Itailcall; eauto. apply sig_function_translated; auto.
+ eapply exec_Itailcall; eauto. apply sig_function_translated; auto.
erewrite stacksize_translated by eauto. eexact C.
constructor; eauto 2 with na. eapply magree_extends; eauto. apply nlive_all.
- (* builtin *)
- TransfInstr; UseTransfer. revert ENV MEM TI.
+ TransfInstr; UseTransfer. revert ENV MEM TI.
functional induction (transfer_builtin (vanalyze rm f)#pc ef args res ne nm);
simpl in *; intros.
+ (* volatile load *)
@@ -886,18 +886,18 @@ Ltac UseTransfer :=
(kill_builtin_res res ne,
nmem_add nm (aaddr_arg (vanalyze rm f) # pc a1)
(size_chunk chunk)) a1) as (ne1, nm1) eqn: TR.
- inversion SS; subst. exploit transfer_builtin_arg_sound; eauto.
- intros (tv1 & A & B & C & D).
- inv H1. simpl in B. inv B.
+ inversion SS; subst. exploit transfer_builtin_arg_sound; eauto.
+ intros (tv1 & A & B & C & D).
+ inv H1. simpl in B. inv B.
assert (X: exists tvres, volatile_load ge chunk tm b ofs t tvres /\ Val.lessdef vres tvres).
{
- inv H2.
- * exists (Val.load_result chunk v); split; auto. constructor; auto.
- * exploit magree_load; eauto.
- exploit aaddr_arg_sound_1; eauto. rewrite <- AN. intros.
- intros. eapply nlive_add; eassumption.
- intros (tv & P & Q).
- exists tv; split; auto. constructor; auto.
+ inv H2.
+ * exists (Val.load_result chunk v); split; auto. constructor; auto.
+ * exploit magree_load; eauto.
+ exploit aaddr_arg_sound_1; eauto. rewrite <- AN. intros.
+ intros. eapply nlive_add; eassumption.
+ intros (tv & P & Q).
+ exists tv; split; auto. constructor; auto.
}
destruct X as (tvres & P & Q).
econstructor; split.
@@ -905,31 +905,31 @@ Ltac UseTransfer :=
apply eval_builtin_args_preserved with (ge1 := ge). exact symbols_preserved.
constructor. eauto. constructor.
eapply external_call_symbols_preserved.
- constructor. simpl. eauto.
+ constructor. simpl. eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
eapply match_succ_states; eauto. simpl; auto.
apply eagree_set_res; auto.
- eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+ eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+ (* volatile store *)
inv H0. inv H6. inv H7. rename b1 into v1. rename b0 into v2.
destruct (transfer_builtin_arg (store_argument chunk)
(kill_builtin_res res ne, nm) a2) as (ne2, nm2) eqn: TR2.
destruct (transfer_builtin_arg All (ne2, nm2) a1) as (ne1, nm1) eqn: TR1.
- inversion SS; subst.
+ inversion SS; subst.
exploit transfer_builtin_arg_sound. eexact H4. eauto. eauto. eauto. eauto. eauto.
intros (tv1 & A1 & B1 & C1 & D1).
exploit transfer_builtin_arg_sound. eexact H3. eauto. eauto. eauto. eauto. eauto.
intros (tv2 & A2 & B2 & C2 & D2).
exploit transf_volatile_store; eauto.
- intros (EQ & tm' & P & Q). subst vres.
+ intros (EQ & tm' & P & Q). subst vres.
econstructor; split.
eapply exec_Ibuiltin; eauto.
apply eval_builtin_args_preserved with (ge1 := ge). exact symbols_preserved.
constructor. eauto. constructor. eauto. constructor.
- eapply external_call_symbols_preserved. simpl; eauto.
+ eapply external_call_symbols_preserved. simpl; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
eapply match_succ_states; eauto. simpl; auto.
- apply eagree_set_res; auto.
+ apply eagree_set_res; auto.
+ (* memcpy *)
rewrite e1 in TI.
inv H0. inv H6. inv H7. rename b1 into v1. rename b0 into v2.
@@ -947,27 +947,27 @@ Ltac UseTransfer :=
intros (tv2 & A2 & B2 & C2 & D2).
inv H1.
exploit magree_loadbytes. eauto. eauto.
- intros. eapply nlive_add; eauto.
+ intros. eapply nlive_add; eauto.
unfold asrc, vanalyze, rm; rewrite AN; eapply aaddr_arg_sound_1; eauto.
intros (tbytes & P & Q).
exploit magree_storebytes_parallel.
- eapply magree_monotone. eexact D2.
+ eapply magree_monotone. eexact D2.
instantiate (1 := nlive ge sp0 (nmem_remove nm adst sz)).
intros. apply incl_nmem_add; auto.
- eauto.
+ eauto.
instantiate (1 := nlive ge sp0 nm).
- intros. eapply nlive_remove; eauto.
+ intros. eapply nlive_remove; eauto.
unfold adst, vanalyze, rm; rewrite AN; eapply aaddr_arg_sound_1; eauto.
- erewrite Mem.loadbytes_length in H1 by eauto.
+ erewrite Mem.loadbytes_length in H1 by eauto.
rewrite nat_of_Z_eq in H1 by omega. auto.
- eauto.
+ eauto.
intros (tm' & A & B).
econstructor; split.
eapply exec_Ibuiltin; eauto.
apply eval_builtin_args_preserved with (ge1 := ge). exact symbols_preserved.
constructor. eauto. constructor. eauto. constructor.
eapply external_call_symbols_preserved. simpl.
- simpl in B1; inv B1. simpl in B2; inv B2. econstructor; eauto.
+ simpl in B1; inv B1. simpl in B2; inv B2. econstructor; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
eapply match_succ_states; eauto. simpl; auto.
apply eagree_set_res; auto.
@@ -978,24 +978,24 @@ Ltac UseTransfer :=
set (asrc := aaddr_arg (vanalyze rm f) # pc src) in *.
inv H1.
econstructor; split.
- eapply exec_Inop; eauto.
+ eapply exec_Inop; eauto.
eapply match_succ_states; eauto. simpl; auto.
destruct res; auto. apply eagree_set_undef; auto.
eapply magree_storebytes_left; eauto.
- exploit aaddr_arg_sound. eauto. eauto.
+ exploit aaddr_arg_sound. eauto. eauto.
intros (bc & A & B & C).
intros. eapply nlive_contains; eauto.
- erewrite Mem.loadbytes_length in H0 by eauto.
+ erewrite Mem.loadbytes_length in H0 by eauto.
rewrite nat_of_Z_eq in H0 by omega. auto.
+ (* annot *)
destruct (transfer_builtin_args (kill_builtin_res res ne, nm) _x1) as (ne1, nm1) eqn:TR.
inversion SS; subst.
exploit transfer_builtin_args_sound; eauto. intros (tvl & A & B & C & D).
inv H1.
- econstructor; split.
+ econstructor; split.
eapply exec_Ibuiltin; eauto.
apply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
- eapply external_call_symbols_preserved. simpl; constructor.
+ eapply external_call_symbols_preserved. simpl; constructor.
eapply eventval_list_match_lessdef; eauto 2 with na.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
eapply match_succ_states; eauto. simpl; auto.
@@ -1005,10 +1005,10 @@ Ltac UseTransfer :=
inversion SS; subst.
exploit transfer_builtin_args_sound; eauto. intros (tvl & A & B & C & D).
inv H1. inv B. inv H6.
- econstructor; split.
- eapply exec_Ibuiltin; eauto.
+ econstructor; split.
+ eapply exec_Ibuiltin; eauto.
apply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
- eapply external_call_symbols_preserved. simpl; constructor.
+ eapply external_call_symbols_preserved. simpl; constructor.
eapply eventval_match_lessdef; eauto 2 with na.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
eapply match_succ_states; eauto. simpl; auto.
@@ -1023,15 +1023,15 @@ Ltac UseTransfer :=
+ (* all other builtins *)
assert ((fn_code tf)!pc = Some(Ibuiltin _x _x0 res pc')).
{
- destruct _x; auto. destruct _x0; auto. destruct _x0; auto. destruct _x0; auto. contradiction.
+ destruct _x; auto. destruct _x0; auto. destruct _x0; auto. destruct _x0; auto. contradiction.
}
clear y TI.
destruct (transfer_builtin_args (kill_builtin_res res ne, nmem_all) _x0) as (ne1, nm1) eqn:TR.
inversion SS; subst.
exploit transfer_builtin_args_sound; eauto. intros (tvl & A & B & C & D).
- exploit external_call_mem_extends; eauto 2 with na.
- eapply magree_extends; eauto. intros. apply nlive_all.
- intros (v' & tm' & P & Q & R & S & T).
+ exploit external_call_mem_extends; eauto 2 with na.
+ eapply magree_extends; eauto. intros. apply nlive_all.
+ intros (v' & tm' & P & Q & R & S & T).
econstructor; split.
eapply exec_Ibuiltin; eauto.
apply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
@@ -1039,33 +1039,33 @@ Ltac UseTransfer :=
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
eapply match_succ_states; eauto. simpl; auto.
apply eagree_set_res; auto.
- eapply mextends_agree; eauto.
+ eapply mextends_agree; eauto.
- (* conditional *)
TransfInstr; UseTransfer.
econstructor; split.
- eapply exec_Icond; eauto.
- eapply needs_of_condition_sound. eapply ma_perm; eauto. eauto. eauto with na.
- eapply match_succ_states; eauto 2 with na.
- simpl; destruct b; auto.
+ eapply exec_Icond; eauto.
+ eapply needs_of_condition_sound. eapply ma_perm; eauto. eauto. eauto with na.
+ eapply match_succ_states; eauto 2 with na.
+ simpl; destruct b; auto.
- (* jumptable *)
TransfInstr; UseTransfer.
- assert (LD: Val.lessdef rs#arg te#arg) by eauto 2 with na.
- rewrite H0 in LD. inv LD.
+ assert (LD: Val.lessdef rs#arg te#arg) by eauto 2 with na.
+ rewrite H0 in LD. inv LD.
econstructor; split.
- eapply exec_Ijumptable; eauto.
- eapply match_succ_states; eauto 2 with na.
- simpl. eapply list_nth_z_in; eauto.
+ eapply exec_Ijumptable; eauto.
+ eapply match_succ_states; eauto 2 with na.
+ simpl. eapply list_nth_z_in; eauto.
- (* return *)
TransfInstr; UseTransfer.
- exploit magree_free. eauto. eauto. instantiate (1 := nlive ge stk nmem_all).
- intros; eapply nlive_dead_stack; eauto.
- intros (tm' & A & B).
+ exploit magree_free. eauto. eauto. instantiate (1 := nlive ge stk nmem_all).
+ intros; eapply nlive_dead_stack; eauto.
+ intros (tm' & A & B).
econstructor; split.
- eapply exec_Ireturn; eauto.
- erewrite stacksize_translated by eauto. eexact A.
+ eapply exec_Ireturn; eauto.
+ erewrite stacksize_translated by eauto. eexact A.
constructor; auto.
destruct or; simpl; eauto 2 with na.
eapply magree_extends; eauto. apply nlive_all.
@@ -1074,28 +1074,28 @@ Ltac UseTransfer :=
monadInv FUN. generalize EQ. unfold transf_function. intros EQ'.
destruct (analyze (vanalyze rm f) f) as [an|] eqn:AN; inv EQ'.
exploit Mem.alloc_extends; eauto. apply Zle_refl. apply Zle_refl.
- intros (tm' & A & B).
+ intros (tm' & A & B).
econstructor; split.
- econstructor; simpl; eauto.
- simpl. econstructor; eauto.
- apply eagree_init_regs; auto.
- apply mextends_agree; auto.
+ econstructor; simpl; eauto.
+ simpl. econstructor; eauto.
+ apply eagree_init_regs; auto.
+ apply mextends_agree; auto.
- (* external function *)
exploit external_call_mem_extends; eauto.
- intros (res' & tm' & A & B & C & D & E).
+ intros (res' & tm' & A & B & C & D & E).
simpl in FUN. inv FUN.
econstructor; split.
econstructor; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- econstructor; eauto.
+ econstructor; eauto.
- (* return *)
- inv STACKS. inv H1.
+ inv STACKS. inv H1.
econstructor; split.
- constructor.
- econstructor; eauto. apply mextends_agree; auto.
+ constructor.
+ econstructor; eauto. apply mextends_agree; auto.
Qed.
Lemma transf_initial_states:
@@ -1114,10 +1114,10 @@ Proof.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
- intros. inv H0. inv H. inv STACKS. inv RES. constructor.
+ intros. inv H0. inv H. inv STACKS. inv RES. constructor.
Qed.
(** * Semantic preservation *)
@@ -1130,12 +1130,12 @@ Proof.
(match_states := fun s1 s2 => sound_state prog s1 /\ match_states s1 s2).
- exact public_preserved.
- simpl; intros. exploit transf_initial_states; eauto. intros [st2 [A B]].
- exists st2; intuition. eapply sound_initial; eauto.
-- simpl; intros. destruct H. eapply transf_final_states; eauto.
+ exists st2; intuition. eapply sound_initial; eauto.
+- simpl; intros. destruct H. eapply transf_final_states; eauto.
- simpl; intros. destruct H0.
assert (sound_state prog s1') by (eapply sound_step; eauto).
fold ge; fold tge. exploit step_simulation; eauto. intros [st2' [A B]].
- exists st2'; auto.
+ exists st2'; auto.
Qed.
End PRESERVATION.
diff --git a/backend/Debugvar.v b/backend/Debugvar.v
index 314f43fd..dcc4327a 100644
--- a/backend/Debugvar.v
+++ b/backend/Debugvar.v
@@ -10,7 +10,7 @@
(* *)
(* *********************************************************************)
-(** Computation of live ranges for local variables that carry
+(** Computation of live ranges for local variables that carry
debugging information. *)
Require Import Coqlib.
@@ -154,7 +154,7 @@ Definition eq_debuginfo (i1 i2: debuginfo) : {i1=i2} + {i1 <> i2}.
Proof.
destruct (eq_arg (proj1_sig i1) (proj1_sig i2)).
left. destruct i1, i2; simpl in *. subst x0. f_equal. apply proof_irr.
- right. congruence.
+ right. congruence.
Defined.
Global Opaque eq_debuginfo.
@@ -177,7 +177,7 @@ Fixpoint join (s1: avail) (s2: avail) {struct s1} : avail :=
Definition eq_state (s1 s2: avail) : {s1=s2} + {s1<>s2}.
Proof.
- apply list_eq_dec. decide equality. apply eq_debuginfo. apply ident_eq.
+ apply list_eq_dec. decide equality. apply eq_debuginfo. apply ident_eq.
Defined.
Global Opaque eq_state.
@@ -273,7 +273,7 @@ Definition transfer (lm: labelmap) (before: option avail) (i: instruction):
end
end.
-(** One pass of forward analysis over the code [c].
+(** One pass of forward analysis over the code [c].
Return an updated label map. *)
Fixpoint ana_code (lm: labelmap) (before: option avail) (c: code) : labelmap :=
diff --git a/backend/Debugvarproof.v b/backend/Debugvarproof.v
index 6f0b8cda..73e32103 100644
--- a/backend/Debugvarproof.v
+++ b/backend/Debugvarproof.v
@@ -58,18 +58,18 @@ Qed.
Lemma transf_code_match:
forall lm c before, match_code c (transf_code lm before c).
Proof.
- intros lm. fix REC 1. destruct c; intros before; simpl.
+ intros lm. fix REC 1. destruct c; intros before; simpl.
- constructor.
- assert (DEFAULT: forall before after,
match_code (i :: c)
(i :: add_delta_ranges before after (transf_code lm after c))).
{ intros. constructor. apply REC. }
- destruct i; auto. destruct c; auto. destruct i; auto.
+ destruct i; auto. destruct c; auto. destruct i; auto.
set (after := get_label l0 lm).
set (c1 := Llabel l0 :: add_delta_ranges before after (transf_code lm after c)).
replace c1 with (add_delta_ranges before before c1).
constructor. constructor. apply REC.
- unfold add_delta_ranges. rewrite delta_state_same. auto.
+ unfold add_delta_ranges. rewrite delta_state_same. auto.
Qed.
Inductive match_function: function -> function -> Prop :=
@@ -80,15 +80,15 @@ Inductive match_function: function -> function -> Prop :=
Lemma transf_function_match:
forall f tf, transf_function f = OK tf -> match_function f tf.
Proof.
- unfold transf_function; intros.
- destruct (ana_function f) as [lm|]; inv H.
- constructor. apply transf_code_match.
+ unfold transf_function; intros.
+ destruct (ana_function f) as [lm|]; inv H.
+ constructor. apply transf_code_match.
Qed.
Remark find_label_add_delta_ranges:
forall lbl c before after, find_label lbl (add_delta_ranges before after c) = find_label lbl c.
Proof.
- intros. unfold add_delta_ranges.
+ intros. unfold add_delta_ranges.
destruct (delta_state before after) as [killed born].
induction killed as [ | [v i] l]; simpl; auto.
induction born as [ | [v i] l]; simpl; auto.
@@ -104,7 +104,7 @@ Proof.
- discriminate.
- destruct (is_label lbl i).
inv H0. econstructor; econstructor; econstructor; eauto.
- rewrite find_label_add_delta_ranges. auto.
+ rewrite find_label_add_delta_ranges. auto.
Qed.
Lemma find_label_match:
@@ -113,7 +113,7 @@ Lemma find_label_match:
find_label lbl f.(fn_code) = Some c ->
exists before after tc, find_label lbl tf.(fn_code) = Some (add_delta_ranges before after tc) /\ match_code c tc.
Proof.
- intros. inv H. eapply find_label_match_rec; eauto.
+ intros. inv H. eapply find_label_match_rec; eauto.
Qed.
(** * Properties of availability sets *)
@@ -135,9 +135,9 @@ Inductive wf_avail: avail -> Prop :=
Lemma set_state_1:
forall v i s, In (v, i) (set_state v i s).
Proof.
- induction s as [ | [v' i'] s]; simpl.
+ induction s as [ | [v' i'] s]; simpl.
- auto.
-- destruct (Pos.compare v v'); simpl; auto.
+- destruct (Pos.compare v v'); simpl; auto.
Qed.
Lemma set_state_2:
@@ -153,7 +153,7 @@ Proof.
Qed.
Lemma set_state_3:
- forall v i v' i' s,
+ forall v i v' i' s,
wf_avail s ->
In (v', i') (set_state v i s) ->
(v' = v /\ i' = i) \/ (v' <> v /\ In (v', i') s).
@@ -162,7 +162,7 @@ Proof.
- intuition congruence.
- destruct (Pos.compare_spec v v0); simpl in H1.
+ subst v0. destruct H1. inv H1; auto. right; split.
- apply sym_not_equal. apply Plt_ne. eapply H; eauto.
+ apply sym_not_equal. apply Plt_ne. eapply H; eauto.
auto.
+ destruct H1. inv H1; auto.
destruct H1. inv H1. right; split; auto. apply sym_not_equal. apply Plt_ne. auto.
@@ -177,12 +177,12 @@ Proof.
induction 1; simpl.
- constructor. red; simpl; tauto. constructor.
- destruct (Pos.compare_spec v v0).
-+ subst v0. constructor; auto.
-+ constructor.
- red; simpl; intros. destruct H2.
- inv H2. auto. apply Plt_trans with v0; eauto.
++ subst v0. constructor; auto.
++ constructor.
+ red; simpl; intros. destruct H2.
+ inv H2. auto. apply Plt_trans with v0; eauto.
constructor; auto.
-+ constructor.
++ constructor.
red; intros. exploit set_state_3. eexact H0. eauto. intros [[A B] | [A B]]; subst; eauto.
auto.
Qed.
@@ -194,8 +194,8 @@ Proof.
- auto.
- destruct (Pos.compare_spec v v0); simpl in *.
+ subst v0. elim (Plt_strict v); eauto.
-+ destruct H1. inv H1. elim (Plt_strict v); eauto.
- elim (Plt_strict v). apply Plt_trans with v0; eauto.
++ destruct H1. inv H1. elim (Plt_strict v); eauto.
+ elim (Plt_strict v). apply Plt_trans with v0; eauto.
+ destruct H1. inv H1. elim (Plt_strict v); eauto. tauto.
Qed.
@@ -219,7 +219,7 @@ Proof.
+ subst v0. split; auto. apply sym_not_equal; apply Plt_ne; eauto.
+ destruct H1. inv H1. split; auto. apply sym_not_equal; apply Plt_ne; eauto.
split; auto. apply sym_not_equal; apply Plt_ne. apply Plt_trans with v0; eauto.
-+ destruct H1. inv H1. split; auto. apply Plt_ne; auto.
++ destruct H1. inv H1. split; auto. apply Plt_ne; auto.
destruct IHwf_avail as [A B] ; auto.
Qed.
@@ -240,9 +240,9 @@ Lemma wf_filter:
Proof.
induction 1; simpl.
- constructor.
-- destruct (pred (v, i)) eqn:P; auto.
- constructor; auto.
- red; intros. apply filter_In in H1. destruct H1. eauto.
+- destruct (pred (v, i)) eqn:P; auto.
+ constructor; auto.
+ red; intros. apply filter_In in H1. destruct H1. eauto.
Qed.
Lemma join_1:
@@ -252,12 +252,12 @@ Proof.
induction 1; simpl; try tauto; induction 1; simpl; intros I1 I2; auto.
destruct I1, I2.
- inv H3; inv H4. rewrite Pos.compare_refl. rewrite dec_eq_true; auto with coqlib.
-- inv H3.
- assert (L: Plt v1 v) by eauto. apply Pos.compare_gt_iff in L. rewrite L. auto.
+- inv H3.
+ assert (L: Plt v1 v) by eauto. apply Pos.compare_gt_iff in L. rewrite L. auto.
- inv H4.
assert (L: Plt v0 v) by eauto. apply Pos.compare_lt_iff in L. rewrite L. apply IHwf_avail. constructor; auto. auto. auto with coqlib.
- destruct (Pos.compare v0 v1).
-+ destruct (eq_debuginfo i0 i1); auto with coqlib.
++ destruct (eq_debuginfo i0 i1); auto with coqlib.
+ apply IHwf_avail; auto with coqlib. constructor; auto.
+ eauto.
Qed.
@@ -266,12 +266,12 @@ Lemma join_2:
forall v i s1, wf_avail s1 -> forall s2, wf_avail s2 ->
In (v, i) (join s1 s2) -> In (v, i) s1 /\ In (v, i) s2.
Proof.
- induction 1; simpl; try tauto; induction 1; simpl; intros I; try tauto.
+ induction 1; simpl; try tauto; induction 1; simpl; intros I; try tauto.
destruct (Pos.compare_spec v0 v1).
- subst v1. destruct (eq_debuginfo i0 i1).
+ subst i1. destruct I. auto. exploit IHwf_avail; eauto. tauto.
+ exploit IHwf_avail; eauto. tauto.
-- exploit (IHwf_avail ((v1, i1) :: s0)); eauto. constructor; auto.
+- exploit (IHwf_avail ((v1, i1) :: s0)); eauto. constructor; auto.
simpl. tauto.
- exploit IHwf_avail0; eauto. tauto.
Qed.
@@ -281,10 +281,10 @@ Lemma wf_join:
Proof.
induction 1; simpl; induction 1; simpl; try constructor.
destruct (Pos.compare_spec v v0).
-- subst v0. destruct (eq_debuginfo i i0); auto. constructor; auto.
+- subst v0. destruct (eq_debuginfo i i0); auto. constructor; auto.
red; intros. apply join_2 in H3; auto. destruct H3. eauto.
-- apply IHwf_avail. constructor; auto.
-- apply IHwf_avail0.
+- apply IHwf_avail. constructor; auto.
+- apply IHwf_avail0.
Qed.
(** * Semantic preservation *)
@@ -334,7 +334,7 @@ Lemma sig_preserved:
Proof.
unfold transf_fundef, transf_partial_fundef; intros.
destruct f. monadInv H.
- exploit transf_function_match; eauto. intros M; inv M; auto.
+ exploit transf_function_match; eauto. intros M; inv M; auto.
inv H. reflexivity.
Qed.
@@ -360,7 +360,7 @@ Proof.
induction a; simpl; intros; try contradiction;
try (econstructor; now eauto with barg).
destruct H as [S1 S2].
- destruct (IHa1 S1) as [v1 E1]. destruct (IHa2 S2) as [v2 E2].
+ destruct (IHa1 S1) as [v1 E1]. destruct (IHa2 S2) as [v2 E2].
exists (Val.longofwords v1 v2); auto with barg.
Qed.
@@ -369,24 +369,24 @@ Lemma eval_add_delta_ranges:
star step tge (State s f sp (add_delta_ranges before after c) rs m)
E0 (State s f sp c rs m).
Proof.
- intros. unfold add_delta_ranges.
+ intros. unfold add_delta_ranges.
destruct (delta_state before after) as [killed born].
induction killed as [ | [v i] l]; simpl.
- induction born as [ | [v i] l]; simpl.
+ apply star_refl.
-+ destruct i as [a SAFE]; simpl.
- exploit can_eval_safe_arg; eauto. intros [v1 E1].
- eapply star_step; eauto.
- econstructor.
++ destruct i as [a SAFE]; simpl.
+ exploit can_eval_safe_arg; eauto. intros [v1 E1].
+ eapply star_step; eauto.
+ econstructor.
constructor. eexact E1. constructor.
simpl; constructor.
- simpl; auto.
+ simpl; auto.
traceEq.
-- eapply star_step; eauto.
- econstructor.
+- eapply star_step; eauto.
+ econstructor.
constructor.
simpl; constructor.
- simpl; auto.
+ simpl; auto.
traceEq.
Qed.
@@ -426,7 +426,7 @@ Lemma parent_locset_match:
list_forall2 match_stackframes s ts ->
parent_locset ts = parent_locset s.
Proof.
- induction 1; simpl. auto. inv H; auto.
+ induction 1; simpl. auto. inv H; auto.
Qed.
(** The simulation diagram. *)
@@ -455,9 +455,9 @@ Proof.
- (* load *)
econstructor; split.
eapply plus_left.
- eapply exec_Lload with (a := a).
+ eapply exec_Lload with (a := a).
rewrite <- H; apply eval_addressing_preserved; exact symbols_preserved.
- eauto. eauto.
+ eauto. eauto.
apply eval_add_delta_ranges. traceEq.
constructor; auto.
- (* store *)
@@ -465,7 +465,7 @@ Proof.
eapply plus_left.
eapply exec_Lstore with (a := a).
rewrite <- H; apply eval_addressing_preserved; exact symbols_preserved.
- eauto. eauto.
+ eauto. eauto.
apply eval_add_delta_ranges. traceEq.
constructor; auto.
- (* call *)
@@ -473,16 +473,16 @@ Proof.
econstructor; split.
apply plus_one.
econstructor. eexact A. symmetry; apply sig_preserved; auto. traceEq.
- constructor; auto. constructor; auto. constructor; auto.
+ constructor; auto. constructor; auto. constructor; auto.
- (* tailcall *)
exploit find_function_translated; eauto. intros (tf' & A & B).
exploit parent_locset_match; eauto. intros PLS.
econstructor; split.
- apply plus_one.
- econstructor. eauto. rewrite PLS. eexact A.
+ apply plus_one.
+ econstructor. eauto. rewrite PLS. eexact A.
symmetry; apply sig_preserved; auto.
inv TRF; eauto. traceEq.
- rewrite PLS. constructor; auto.
+ rewrite PLS. constructor; auto.
- (* builtin *)
econstructor; split.
eapply plus_left.
@@ -513,28 +513,28 @@ Proof.
- (* jumptable *)
exploit find_label_match; eauto. intros (before' & after' & tc' & A & B).
econstructor; split.
- eapply plus_left. econstructor; eauto.
+ eapply plus_left. econstructor; eauto.
apply eval_add_delta_ranges. reflexivity. traceEq.
constructor; auto.
- (* return *)
econstructor; split.
apply plus_one. constructor. inv TRF; eauto. traceEq.
- rewrite (parent_locset_match _ _ STACKS). constructor; auto.
+ rewrite (parent_locset_match _ _ STACKS). constructor; auto.
- (* internal function *)
- monadInv H7. rename x into tf.
+ monadInv H7. rename x into tf.
assert (MF: match_function f tf) by (apply transf_function_match; auto).
inversion MF; subst.
econstructor; split.
- apply plus_one. constructor. simpl; eauto. reflexivity.
- constructor; auto.
+ apply plus_one. constructor. simpl; eauto. reflexivity.
+ constructor; auto.
- (* external function *)
monadInv H8. econstructor; split.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
eapply external_call_symbols_preserved'. eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
constructor; auto.
- (* return *)
- inv H3. inv H1.
+ inv H3. inv H1.
econstructor; split.
eapply plus_left. econstructor. apply eval_add_delta_ranges. traceEq.
constructor; auto.
@@ -545,18 +545,18 @@ Lemma transf_initial_states:
exists st2, initial_state tprog st2 /\ match_states st1 st2.
Proof.
intros. inversion H.
- exploit function_ptr_translated; eauto. intros [tf [A B]].
+ exploit function_ptr_translated; eauto. intros [tf [A B]].
exists (Callstate nil tf (Locmap.init Vundef) m0); split.
- econstructor; eauto. eapply Genv.init_mem_transf_partial; eauto.
+ 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).
+ symmetry. apply (transform_partial_program_main transf_fundef _ TRANSF).
rewrite <- H3. apply sig_preserved. auto.
constructor. constructor. auto.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
intros. inv H0. inv H. inv H6. econstructor; eauto.
diff --git a/backend/IRC.ml b/backend/IRC.ml
index dcd8624a..eb677069 100644
--- a/backend/IRC.ml
+++ b/backend/IRC.ml
@@ -163,7 +163,7 @@ module DLinkMove = struct
type t = move
let make state =
let rec empty =
- { src = DLinkNode.dummy; dst = DLinkNode.dummy;
+ { src = DLinkNode.dummy; dst = DLinkNode.dummy;
mstate = state; mprev = empty; mnext = empty }
in empty
let dummy = make CoalescedMoves
@@ -301,7 +301,7 @@ let init costs =
worklistMoves = DLinkMove.make WorklistMoves;
activeMoves = DLinkMove.make ActiveMoves
}
-
+
(* Create nodes corresponding to XTL variables *)
let weightedSpillCost st =
@@ -312,7 +312,7 @@ let weightedSpillCost st =
let newNodeOfReg g r ty =
let st = g.stats_of_reg r in
g.nextIdent <- g.nextIdent + 1;
- { ident = g.nextIdent; typ = ty;
+ { ident = g.nextIdent; typ = ty;
var = V(r, ty);
regclass = if st.cost >= 0 then class_of_type ty else no_spill_class;
accesses = st.usedefs;
@@ -328,7 +328,7 @@ let newNodeOfReg g r ty =
let newNodeOfLoc g l =
let ty = Loc.coq_type l in
g.nextIdent <- g.nextIdent + 1;
- { ident = g.nextIdent; typ = ty;
+ { 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 = [];
@@ -608,9 +608,9 @@ let canCoalesceGeorge g u v =
- 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
+ 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
+ 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,
@@ -690,7 +690,7 @@ let coalesce g =
combine g u v;
addWorkList g u
end else begin
- DLinkMove.insert m g.activeMoves
+ DLinkMove.insert m g.activeMoves
end
(* Freeze moves associated with node [u] *)
diff --git a/backend/Inlining.v b/backend/Inlining.v
index 08f2bfc4..566ab27c 100644
--- a/backend/Inlining.v
+++ b/backend/Inlining.v
@@ -107,7 +107,7 @@ Definition initstate :=
mkstate 1%positive 1%positive (PTree.empty instruction) 0.
Program Definition set_instr (pc: node) (i: instruction): mon unit :=
- fun s =>
+ fun s =>
R tt
(mkstate s.(st_nextreg) s.(st_nextnode) (PTree.set pc i s.(st_code)) s.(st_stksize))
_.
@@ -144,7 +144,7 @@ Next Obligation.
Qed.
Program Definition request_stack (sz: Z): mon unit :=
- fun s =>
+ fun s =>
R tt
(mkstate s.(st_nextreg) s.(st_nextnode) s.(st_code) (Zmax s.(st_stksize) sz))
_.
@@ -154,14 +154,14 @@ Qed.
Program Definition ptree_mfold {A: Type} (f: positive -> A -> mon unit) (t: PTree.t A): mon unit :=
fun s =>
- R tt
+ 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.
+ intros. destruct (f k v a). eapply sincr_trans; eauto.
Qed.
(** ** Inlining contexts *)
@@ -280,7 +280,7 @@ Section EXPAND_CFG.
Variable fenv: funenv.
-(** The [rec] parameter is the recursor: [rec fenv' P ctx f] copies
+(** The [rec] parameter is the recursor: [rec fenv' P ctx f] copies
the body of function [f], with inline expansion within, as governed
by context [ctx]. It can only be called for function environments
[fenv'] strictly smaller than the current environment [fenv]. *)
@@ -311,7 +311,7 @@ Definition inline_function (ctx: context) (id: ident) (f: function)
let nreg := max_reg_function f in
do dpc <- reserve_nodes npc;
do dreg <- reserve_regs nreg;
- let ctx' := callcontext ctx dpc dreg nreg f.(fn_stacksize) retpc retreg in
+ let ctx' := callcontext ctx dpc dreg nreg f.(fn_stacksize) retpc retreg in
do x <- rec (PTree.remove id fenv) (PTree_Properties.cardinal_remove P) ctx' f;
add_moves (sregs ctx args) (sregs ctx' f.(fn_params)) (spc ctx' f.(fn_entrypoint)).
@@ -325,7 +325,7 @@ Definition inline_tail_function (ctx: context) (id: ident) (f: function)
let nreg := max_reg_function f in
do dpc <- reserve_nodes npc;
do dreg <- reserve_regs nreg;
- let ctx' := tailcontext ctx dpc dreg nreg f.(fn_stacksize) in
+ let ctx' := tailcontext ctx dpc dreg nreg f.(fn_stacksize) in
do x <- rec (PTree.remove id fenv) (PTree_Properties.cardinal_remove P) ctx' f;
add_moves (sregs ctx args) (sregs ctx' f.(fn_params)) (spc ctx' f.(fn_entrypoint)).
@@ -341,7 +341,7 @@ Definition inline_return (ctx: context) (or: option reg) (retinfo: node * reg) :
(** Expansion and copying of an instruction. For most instructions,
its registers and successor PC are shifted as per the context [ctx],
then the instruction is inserted in the final CFG at its final position
- [spc ctx pc].
+ [spc ctx pc].
[Icall] instructions are either replaced by a "goto" to the expansion
of the called function, or shifted as described above.
@@ -393,7 +393,7 @@ Definition expand_instr (ctx: context) (pc: node) (i: instruction): mon unit :=
| Can_inline id f P Q =>
do n <- inline_tail_function ctx id f Q args;
set_instr (spc ctx pc) (Inop n)
- end
+ end
| Ibuiltin ef args res s =>
set_instr (spc ctx pc)
(Ibuiltin ef (map (sbuiltinarg ctx) args) (sbuiltinres ctx res) (spc ctx s))
@@ -434,7 +434,7 @@ Definition expand_function (fenv: funenv) (f: function): mon context :=
let nreg := max_reg_function f in
do dpc <- reserve_nodes npc;
do dreg <- reserve_regs nreg;
- let ctx := initcontext dpc dreg nreg f.(fn_stacksize) in
+ let ctx := initcontext dpc dreg nreg f.(fn_stacksize) in
do x <- expand_cfg fenv ctx f;
ret ctx.
@@ -450,7 +450,7 @@ Local Open Scope string_scope.
Definition transf_function (fenv: funenv) (f: function) : Errors.res function :=
let '(R ctx s _) := expand_function fenv f initstate in
if zlt s.(st_stksize) Int.max_unsigned then
- OK (mkfunction f.(fn_sig)
+ OK (mkfunction f.(fn_sig)
(sregs ctx f.(fn_params))
s.(st_stksize)
s.(st_code)
diff --git a/backend/Inliningproof.v b/backend/Inliningproof.v
index c7cc8d8a..ad861543 100644
--- a/backend/Inliningproof.v
+++ b/backend/Inliningproof.v
@@ -71,7 +71,7 @@ Lemma sig_function_translated:
forall f f', transf_fundef fenv f = OK f' -> funsig f' = funsig f.
Proof.
intros. destruct f; Errors.monadInv H.
- exploit transf_function_spec; eauto. intros SP; inv SP. auto.
+ exploit transf_function_spec; eauto. intros SP; inv SP. auto.
auto.
Qed.
@@ -80,7 +80,7 @@ Qed.
Remark sreg_below_diff:
forall ctx r r', Plt r' ctx.(dreg) -> sreg ctx r <> r'.
Proof.
- intros. zify. unfold sreg; rewrite shiftpos_eq. xomega.
+ intros. zify. unfold sreg; rewrite shiftpos_eq. xomega.
Qed.
Remark context_below_diff:
@@ -93,7 +93,7 @@ Qed.
Remark context_below_lt:
forall ctx1 ctx2 r, context_below ctx1 ctx2 -> Ple r ctx1.(mreg) -> Plt (sreg ctx1 r) ctx2.(dreg).
Proof.
- intros. red in H. unfold Plt; zify. unfold sreg; rewrite shiftpos_eq.
+ intros. red in H. unfold Plt; zify. unfold sreg; rewrite shiftpos_eq.
xomega.
Qed.
@@ -101,7 +101,7 @@ Qed.
Remark context_below_le:
forall ctx1 ctx2 r, context_below ctx1 ctx2 -> Ple r ctx1.(mreg) -> Ple (sreg ctx1 r) ctx2.(dreg).
Proof.
- intros. red in H. unfold Ple; zify. unfold sreg; rewrite shiftpos_eq.
+ intros. red in H. unfold Ple; zify. unfold sreg; rewrite shiftpos_eq.
xomega.
Qed.
*)
@@ -125,8 +125,8 @@ Lemma agree_val_reg_gen:
forall F ctx rs rs' r, agree_regs F ctx rs rs' -> val_reg_charact F ctx rs' rs#r r.
Proof.
intros. destruct H as [A B].
- destruct (Plt_Ple_dec (mreg ctx) r).
- left. rewrite B; auto.
+ destruct (Plt_Ple_dec (mreg ctx) r).
+ left. rewrite B; auto.
right. auto.
Qed.
@@ -159,10 +159,10 @@ Lemma agree_set_reg:
agree_regs F ctx (rs#r <- v) (rs'#(sreg ctx r) <- v').
Proof.
unfold agree_regs; intros. destruct H. split; intros.
- repeat rewrite Regmap.gsspec.
+ repeat rewrite Regmap.gsspec.
destruct (peq r0 r). subst r0. rewrite peq_true. auto.
- rewrite peq_false. auto. apply shiftpos_diff; auto.
- rewrite Regmap.gso. auto. xomega.
+ rewrite peq_false. auto. apply shiftpos_diff; auto.
+ rewrite Regmap.gso. auto. xomega.
Qed.
Lemma agree_set_reg_undef:
@@ -171,10 +171,10 @@ Lemma agree_set_reg_undef:
agree_regs F ctx (rs#r <- Vundef) (rs'#(sreg ctx r) <- v').
Proof.
unfold agree_regs; intros. destruct H. split; intros.
- repeat rewrite Regmap.gsspec.
+ repeat rewrite Regmap.gsspec.
destruct (peq r0 r). subst r0. rewrite peq_true. auto.
- rewrite peq_false. auto. apply shiftpos_diff; auto.
- rewrite Regmap.gsspec. destruct (peq r0 r); auto.
+ rewrite peq_false. auto. apply shiftpos_diff; auto.
+ rewrite Regmap.gsspec. destruct (peq r0 r); auto.
Qed.
Lemma agree_set_reg_undef':
@@ -183,9 +183,9 @@ Lemma agree_set_reg_undef':
agree_regs F ctx (rs#r <- Vundef) rs'.
Proof.
unfold agree_regs; intros. destruct H. split; intros.
- rewrite Regmap.gsspec.
+ rewrite Regmap.gsspec.
destruct (peq r0 r). subst r0. auto. auto.
- rewrite Regmap.gsspec. destruct (peq r0 r); auto.
+ rewrite Regmap.gsspec. destruct (peq r0 r); auto.
Qed.
Lemma agree_regs_invariant:
@@ -195,7 +195,7 @@ Lemma agree_regs_invariant:
agree_regs F ctx rs rs2.
Proof.
unfold agree_regs; intros. destruct H. split; intros.
- rewrite H0. auto.
+ rewrite H0. auto.
apply shiftpos_above.
eapply Plt_le_trans. apply shiftpos_below. xomega.
apply H1; auto.
@@ -207,13 +207,13 @@ Lemma agree_regs_incr:
inject_incr F F' ->
agree_regs F' ctx rs1 rs2.
Proof.
- intros. destruct H. split; intros. eauto. auto.
+ intros. destruct H. split; intros. eauto. auto.
Qed.
Remark agree_regs_init:
forall F ctx rs, agree_regs F ctx (Regmap.init Vundef) rs.
Proof.
- intros; split; intros. rewrite Regmap.gi; auto. rewrite Regmap.gi; auto.
+ intros; split; intros. rewrite Regmap.gi; auto. rewrite Regmap.gi; auto.
Qed.
Lemma agree_regs_init_regs:
@@ -225,7 +225,7 @@ Proof.
induction rl; simpl; intros.
apply agree_regs_init.
inv H. apply agree_regs_init.
- apply agree_set_reg; auto.
+ apply agree_set_reg; auto.
Qed.
(** ** Executing sequences of moves *)
@@ -246,7 +246,7 @@ Proof.
(* rdsts = nil *)
inv H0. exists rs1; split. apply star_refl. split. apply agree_regs_init. auto.
(* rdsts = a :: rdsts *)
- inv H2. inv H0.
+ inv H2. inv H0.
exists rs1; split. apply star_refl. split. apply agree_regs_init. auto.
simpl in H0. inv H0.
exploit IHrdsts; eauto. intros [rs2 [A [B C]]].
@@ -285,7 +285,7 @@ Lemma range_private_invariant:
range_private F1 m1 m1' sp lo hi.
Proof.
intros; red; intros. exploit H; eauto. intros [A B]. split; auto.
- intros; red; intros. exploit H0; eauto. omega. intros [P Q].
+ intros; red; intros. exploit H0; eauto. omega. intros [P Q].
eelim B; eauto.
Qed.
@@ -305,14 +305,14 @@ Lemma range_private_alloc_left:
(forall b, b <> sp -> F1 b = F b) ->
range_private F1 m1 m' sp' (base + Zmax sz 0) hi.
Proof.
- intros; red; intros.
+ intros; red; intros.
exploit (H ofs). generalize (Zmax2 sz 0). omega. intros [A B].
split; auto. intros; red; intros.
exploit Mem.perm_alloc_inv; eauto.
destruct (eq_block b sp); intros.
- subst b. rewrite H1 in H4; inv H4.
+ subst b. rewrite H1 in H4; inv H4.
rewrite Zmax_spec in H3. destruct (zlt 0 sz); omega.
- rewrite H2 in H4; auto. eelim B; eauto.
+ rewrite H2 in H4; auto. eelim B; eauto.
Qed.
Lemma range_private_free_left:
@@ -323,22 +323,22 @@ Lemma range_private_free_left:
Mem.inject F m m' ->
range_private F m1 m' sp base hi.
Proof.
- intros; red; intros.
+ intros; red; intros.
destruct (zlt ofs (base + Zmax sz 0)) as [z|z].
- red; split.
+ red; split.
replace ofs with ((ofs - base) + base) by omega.
eapply Mem.perm_inject; eauto.
eapply Mem.free_range_perm; eauto.
- rewrite Zmax_spec in z. destruct (zlt 0 sz); omega.
+ rewrite Zmax_spec in z. destruct (zlt 0 sz); omega.
intros; red; intros. destruct (eq_block b b0).
subst b0. rewrite H1 in H4; inv H4.
eelim Mem.perm_free_2; eauto. rewrite Zmax_spec in z. destruct (zlt 0 sz); omega.
- exploit Mem.mi_no_overlap; eauto.
+ exploit Mem.mi_no_overlap; eauto.
apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
- eapply Mem.free_range_perm. eauto.
+ eapply Mem.free_range_perm. eauto.
instantiate (1 := ofs - base). rewrite Zmax_spec in z. destruct (zlt 0 sz); omega.
- eapply Mem.perm_free_3; eauto.
- intros [A | A]. congruence. omega.
+ eapply Mem.perm_free_3; eauto.
+ intros [A | A]. congruence. omega.
exploit (H ofs). omega. intros [A B]. split. auto.
intros; red; intros. eelim B; eauto. eapply Mem.perm_free_3; eauto.
@@ -358,13 +358,13 @@ Lemma range_private_extcall:
Proof.
intros until hi; intros RP PERM UNCH INJ INCR SEP VB.
red; intros. exploit RP; eauto. intros [A B].
- split. eapply Mem.perm_unchanged_on; eauto.
+ split. eapply Mem.perm_unchanged_on; eauto.
intros. red in SEP. destruct (F b) as [[sp1 delta1] |] eqn:?.
- exploit INCR; eauto. intros EQ; rewrite H0 in EQ; inv EQ.
- red; intros; eelim B; eauto. eapply PERM; eauto.
- red. destruct (plt b (Mem.nextblock m1)); auto.
+ exploit INCR; eauto. intros EQ; rewrite H0 in EQ; inv EQ.
+ red; intros; eelim B; eauto. eapply PERM; eauto.
+ red. destruct (plt b (Mem.nextblock m1)); auto.
exploit Mem.mi_freeblocks; eauto. congruence.
- exploit SEP; eauto. tauto.
+ exploit SEP; eauto. tauto.
Qed.
(** ** Relating global environments *)
@@ -392,7 +392,7 @@ Proof.
assert (A: Val.inject F rs#r rs'#(sreg ctx r)). eapply agree_val_reg; eauto.
rewrite EQ in A; inv A.
inv H1. rewrite DOMAIN in H5. inv H5. auto.
- apply FUNCTIONS with fd.
+ apply FUNCTIONS with fd.
rewrite EQ in H; rewrite Genv.find_funct_find_funct_ptr in H. auto.
rewrite H2. eapply functions_translated; eauto.
(* symbol *)
@@ -419,24 +419,24 @@ Proof.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
-- exploit Mem.loadv_inject; eauto.
- instantiate (1 := Vptr sp' (Int.add ofs (Int.repr (dstk ctx)))).
+- exploit Mem.loadv_inject; eauto.
+ instantiate (1 := Vptr sp' (Int.add ofs (Int.repr (dstk ctx)))).
simpl. econstructor; eauto. rewrite Int.add_zero_l; auto.
- intros (v' & A & B). exists v'; split; auto. constructor. simpl. rewrite Int.add_zero_l; auto.
+ intros (v' & A & B). exists v'; split; auto. constructor. simpl. rewrite Int.add_zero_l; auto.
- econstructor; split. constructor. simpl. econstructor; eauto. rewrite ! Int.add_zero_l; auto.
- assert (Val.inject F (Senv.symbol_address ge id ofs) (Senv.symbol_address tge id ofs)).
- { unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
+ { unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
rewrite symbols_preserved. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
inv MG. econstructor. eauto. rewrite Int.add_zero; auto. }
exploit Mem.loadv_inject; eauto. intros (v' & A & B).
- exists v'; eauto with barg.
-- econstructor; split. constructor.
- unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
+ exists v'; eauto with barg.
+- econstructor; split. constructor.
+ unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
rewrite symbols_preserved. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
inv MG. econstructor. eauto. rewrite Int.add_zero; auto.
- destruct IHeval_builtin_arg1 as (v1 & A1 & B1).
destruct IHeval_builtin_arg2 as (v2 & A2 & B2).
- econstructor; split. eauto with barg. apply Val.longofwords_inject; auto.
+ econstructor; split. eauto with barg. apply Val.longofwords_inject; auto.
Qed.
Lemma tr_builtin_args:
@@ -522,7 +522,7 @@ Lemma match_stacks_globalenvs:
forall stk stk' bound,
match_stacks F m m' stk stk' bound -> exists b, match_globalenvs F b
with match_stacks_inside_globalenvs:
- forall stk stk' f ctx sp rs',
+ forall stk stk' f ctx sp rs',
match_stacks_inside F m m' stk stk' f ctx sp rs' -> exists b, match_globalenvs F b.
Proof.
induction 1; eauto.
@@ -534,13 +534,13 @@ Lemma match_globalenvs_preserves_globals:
Proof.
intros. inv H. red. split. eauto. split. eauto.
intros. symmetry. eapply IMAGE; eauto.
-Qed.
+Qed.
Lemma match_stacks_inside_globals:
- forall stk stk' f ctx sp rs',
+ forall stk stk' f ctx sp rs',
match_stacks_inside F m m' stk stk' f ctx sp rs' -> meminj_preserves_globals ge F.
Proof.
- intros. exploit match_stacks_inside_globalenvs; eauto. intros [b A].
+ intros. exploit match_stacks_inside_globalenvs; eauto. intros [b A].
eapply match_globalenvs_preserves_globals; eauto.
Qed.
@@ -551,10 +551,10 @@ Lemma match_stacks_bound:
match_stacks F m m' stk stk' bound1.
Proof.
intros. inv H.
- apply match_stacks_nil with bound0. auto. eapply Ple_trans; eauto.
- eapply match_stacks_cons; eauto. eapply Plt_le_trans; eauto.
- eapply match_stacks_untailcall; eauto. eapply Plt_le_trans; eauto.
-Qed.
+ apply match_stacks_nil with bound0. auto. eapply Ple_trans; eauto.
+ eapply match_stacks_cons; eauto. eapply Plt_le_trans; eauto.
+ eapply match_stacks_untailcall; eauto. eapply Plt_le_trans; eauto.
+Qed.
Variable F1: meminj.
Variables m1 m1': mem.
@@ -562,7 +562,7 @@ Hypothesis INCR: inject_incr F F1.
Lemma match_stacks_invariant:
forall stk stk' bound, match_stacks F m m' stk stk' bound ->
- forall (INJ: forall b1 b2 delta,
+ forall (INJ: forall b1 b2 delta,
F1 b1 = Some(b2, delta) -> Plt b2 bound -> F b1 = Some(b2, delta))
(PERM1: forall b1 b2 delta ofs,
F1 b1 = Some(b2, delta) -> Plt b2 bound ->
@@ -574,11 +574,11 @@ Lemma match_stacks_invariant:
match_stacks F1 m1 m1' stk stk' bound
with match_stacks_inside_invariant:
- forall stk stk' f' ctx sp' rs1,
+ forall stk stk' f' ctx sp' rs1,
match_stacks_inside F m m' stk stk' f' ctx sp' rs1 ->
forall rs2
(RS: forall r, Plt r ctx.(dreg) -> rs2#r = rs1#r)
- (INJ: forall b1 b2 delta,
+ (INJ: forall b1 b2 delta,
F1 b1 = Some(b2, delta) -> Ple b2 sp' -> F b1 = Some(b2, delta))
(PERM1: forall b1 b2 delta ofs,
F1 b1 = Some(b2, delta) -> Ple b2 sp' ->
@@ -593,42 +593,42 @@ Proof.
induction 1; intros.
(* nil *)
apply match_stacks_nil with (bound1 := bound1).
- inv MG. constructor; auto.
+ inv MG. constructor; auto.
intros. apply IMAGE with delta. eapply INJ; eauto. eapply Plt_le_trans; eauto.
auto. auto.
(* cons *)
apply match_stacks_cons with (ctx := ctx); auto.
eapply match_stacks_inside_invariant; eauto.
- intros; eapply INJ; eauto; xomega.
+ intros; eapply INJ; eauto; xomega.
intros; eapply PERM1; eauto; xomega.
intros; eapply PERM2; eauto; xomega.
intros; eapply PERM3; eauto; xomega.
eapply agree_regs_incr; eauto.
- eapply range_private_invariant; eauto.
+ eapply range_private_invariant; eauto.
(* untailcall *)
- apply match_stacks_untailcall with (ctx := ctx); auto.
+ apply match_stacks_untailcall with (ctx := ctx); auto.
eapply match_stacks_inside_invariant; eauto.
intros; eapply INJ; eauto; xomega.
intros; eapply PERM1; eauto; xomega.
intros; eapply PERM2; eauto; xomega.
intros; eapply PERM3; eauto; xomega.
- eapply range_private_invariant; eauto.
+ eapply range_private_invariant; eauto.
induction 1; intros.
(* base *)
eapply match_stacks_inside_base; eauto.
- eapply match_stacks_invariant; eauto.
+ eapply match_stacks_invariant; eauto.
intros; eapply INJ; eauto; xomega.
intros; eapply PERM1; eauto; xomega.
intros; eapply PERM2; eauto; xomega.
intros; eapply PERM3; eauto; xomega.
(* inlined *)
- apply match_stacks_inside_inlined with (ctx' := ctx'); auto.
+ apply match_stacks_inside_inlined with (ctx' := ctx'); auto.
apply IHmatch_stacks_inside; auto.
- intros. apply RS. red in BELOW. xomega.
- apply agree_regs_incr with F; auto.
- apply agree_regs_invariant with rs'; auto.
- intros. apply RS. red in BELOW. xomega.
+ intros. apply RS. red in BELOW. xomega.
+ apply agree_regs_incr with F; auto.
+ apply agree_regs_invariant with rs'; auto.
+ intros. apply RS. red in BELOW. xomega.
eapply range_private_invariant; eauto.
intros. split. eapply INJ; eauto. xomega. eapply PERM1; eauto. xomega.
intros. eapply PERM2; eauto. xomega.
@@ -655,33 +655,33 @@ End MATCH_STACKS.
(** Preservation by assignment to a register *)
Lemma match_stacks_inside_set_reg:
- forall F m m' stk stk' f' ctx sp' rs' r v,
+ forall F m m' stk stk' f' ctx sp' rs' r v,
match_stacks_inside F m m' stk stk' f' ctx sp' rs' ->
match_stacks_inside F m m' stk stk' f' ctx sp' (rs'#(sreg ctx r) <- v).
Proof.
- intros. eapply match_stacks_inside_invariant; eauto.
+ intros. eapply match_stacks_inside_invariant; eauto.
intros. apply Regmap.gso. zify. unfold sreg; rewrite shiftpos_eq. xomega.
Qed.
Lemma match_stacks_inside_set_res:
- forall F m m' stk stk' f' ctx sp' rs' res v,
+ forall F m m' stk stk' f' ctx sp' rs' res v,
match_stacks_inside F m m' stk stk' f' ctx sp' rs' ->
match_stacks_inside F m m' stk stk' f' ctx sp' (regmap_setres (sbuiltinres ctx res) v rs').
Proof.
- intros. destruct res; simpl; auto.
+ intros. destruct res; simpl; auto.
apply match_stacks_inside_set_reg; auto.
Qed.
(** Preservation by a memory store *)
Lemma match_stacks_inside_store:
- forall F m m' stk stk' f' ctx sp' rs' chunk b ofs v m1 chunk' b' ofs' v' m1',
+ forall F m m' stk stk' f' ctx sp' rs' chunk b ofs v m1 chunk' b' ofs' v' m1',
match_stacks_inside F m m' stk stk' f' ctx sp' rs' ->
Mem.store chunk m b ofs v = Some m1 ->
Mem.store chunk' m' b' ofs' v' = Some m1' ->
match_stacks_inside F m1 m1' stk stk' f' ctx sp' rs'.
Proof.
- intros.
+ intros.
eapply match_stacks_inside_invariant; eauto with mem.
Qed.
@@ -700,21 +700,21 @@ Lemma match_stacks_inside_alloc_left:
Proof.
induction 1; intros.
(* base *)
- eapply match_stacks_inside_base; eauto.
+ eapply match_stacks_inside_base; eauto.
eapply match_stacks_invariant; eauto.
intros. destruct (eq_block b1 b).
- subst b1. rewrite H1 in H4; inv H4. eelim Plt_strict; eauto.
- rewrite H2 in H4; auto.
+ subst b1. rewrite H1 in H4; inv H4. eelim Plt_strict; eauto.
+ rewrite H2 in H4; auto.
intros. exploit Mem.perm_alloc_inv; eauto. destruct (eq_block b1 b); intros; auto.
- subst b1. rewrite H1 in H4. inv H4. eelim Plt_strict; eauto.
+ subst b1. rewrite H1 in H4. inv H4. eelim Plt_strict; eauto.
(* inlined *)
- eapply match_stacks_inside_inlined; eauto.
- eapply IHmatch_stacks_inside; eauto. destruct SBELOW. omega.
+ eapply match_stacks_inside_inlined; eauto.
+ eapply IHmatch_stacks_inside; eauto. destruct SBELOW. omega.
eapply agree_regs_incr; eauto.
- eapply range_private_invariant; eauto.
+ eapply range_private_invariant; eauto.
intros. exploit Mem.perm_alloc_inv; eauto. destruct (eq_block b0 b); intros.
- subst b0. rewrite H2 in H5; inv H5. elimtype False; xomega.
- rewrite H3 in H5; auto.
+ subst b0. rewrite H2 in H5; inv H5. elimtype False; xomega.
+ rewrite H3 in H5; auto.
Qed.
(** Preservation by freeing *)
@@ -726,7 +726,7 @@ Lemma match_stacks_free_left:
match_stacks F m1 m' stk stk' sp.
Proof.
intros. eapply match_stacks_invariant; eauto.
- intros. eapply Mem.perm_free_3; eauto.
+ intros. eapply Mem.perm_free_3; eauto.
Qed.
Lemma match_stacks_free_right:
@@ -735,18 +735,18 @@ Lemma match_stacks_free_right:
Mem.free m' sp lo hi = Some m1' ->
match_stacks F m m1' stk stk' sp.
Proof.
- intros. eapply match_stacks_invariant; eauto.
- intros. eapply Mem.perm_free_1; eauto.
+ intros. eapply match_stacks_invariant; eauto.
+ intros. eapply Mem.perm_free_1; eauto.
intros. eapply Mem.perm_free_3; eauto.
Qed.
Lemma min_alignment_sound:
forall sz n, (min_alignment sz | n) -> Mem.inj_offset_aligned n sz.
Proof.
- intros; red; intros. unfold min_alignment in H.
+ intros; red; intros. unfold min_alignment in H.
assert (2 <= sz -> (2 | n)). intros.
destruct (zle sz 1). omegaContradiction.
- destruct (zle sz 2). auto.
+ destruct (zle sz 2). auto.
destruct (zle sz 4). apply Zdivides_trans with 4; auto. exists 2; auto.
apply Zdivides_trans with 8; auto. exists 4; auto.
assert (4 <= sz -> (4 | n)). intros.
@@ -780,7 +780,7 @@ Hypothesis INCR: inject_incr F1 F2.
Hypothesis SEP: inject_separated F1 F2 m1 m1'.
Lemma match_stacks_extcall:
- forall stk stk' bound,
+ forall stk stk' bound,
match_stacks F1 m1 m1' stk stk' bound ->
Ple bound (Mem.nextblock m1') ->
match_stacks F2 m2 m2' stk stk' bound
@@ -791,25 +791,25 @@ with match_stacks_inside_extcall:
match_stacks_inside F2 m2 m2' stk stk' f' ctx sp' rs'.
Proof.
induction 1; intros.
- apply match_stacks_nil with bound1; auto.
- inv MG. constructor; intros; eauto.
+ apply match_stacks_nil with bound1; auto.
+ inv MG. constructor; intros; eauto.
destruct (F1 b1) as [[b2' delta']|] eqn:?.
- exploit INCR; eauto. intros EQ; rewrite H0 in EQ; inv EQ. eapply IMAGE; eauto.
- exploit SEP; eauto. intros [A B]. elim B. red. xomega.
- eapply match_stacks_cons; eauto.
- eapply match_stacks_inside_extcall; eauto. xomega.
- eapply agree_regs_incr; eauto.
- eapply range_private_extcall; eauto. red; xomega.
+ exploit INCR; eauto. intros EQ; rewrite H0 in EQ; inv EQ. eapply IMAGE; eauto.
+ exploit SEP; eauto. intros [A B]. elim B. red. xomega.
+ eapply match_stacks_cons; eauto.
+ eapply match_stacks_inside_extcall; eauto. xomega.
+ eapply agree_regs_incr; eauto.
+ eapply range_private_extcall; eauto. red; xomega.
intros. apply SSZ2; auto. apply MAXPERM'; auto. red; xomega.
- eapply match_stacks_untailcall; eauto.
- eapply match_stacks_inside_extcall; eauto. xomega.
- eapply range_private_extcall; eauto. red; xomega.
+ eapply match_stacks_untailcall; eauto.
+ eapply match_stacks_inside_extcall; eauto. xomega.
+ eapply range_private_extcall; eauto. red; xomega.
intros. apply SSZ2; auto. apply MAXPERM'; auto. red; xomega.
induction 1; intros.
eapply match_stacks_inside_base; eauto.
- eapply match_stacks_extcall; eauto. xomega.
- eapply match_stacks_inside_inlined; eauto.
- eapply agree_regs_incr; eauto.
+ eapply match_stacks_extcall; eauto. xomega.
+ eapply match_stacks_inside_inlined; eauto.
+ eapply agree_regs_incr; eauto.
eapply range_private_extcall; eauto.
Qed.
@@ -820,7 +820,7 @@ End EXTCALL.
Lemma align_unchanged:
forall n amount, amount > 0 -> (amount | n) -> align n amount = n.
Proof.
- intros. destruct H0 as [p EQ]. subst n. unfold align. decEq.
+ intros. destruct H0 as [p EQ]. subst n. unfold align. decEq.
apply Zdiv_unique with (b := amount - 1). omega. omega.
Qed.
@@ -836,15 +836,15 @@ Lemma match_stacks_inside_inlined_tailcall:
Proof.
intros. inv H.
(* base *)
- eapply match_stacks_inside_base; eauto. congruence.
+ eapply match_stacks_inside_base; eauto. congruence.
rewrite H1. rewrite DSTK. apply align_unchanged. apply min_alignment_pos. apply Zdivide_0.
(* inlined *)
assert (dstk ctx <= dstk ctx'). rewrite H1. apply align_le. apply min_alignment_pos.
- eapply match_stacks_inside_inlined; eauto.
- red; intros. destruct (zlt ofs (dstk ctx)). apply PAD; omega. apply H3. inv H4. xomega.
- congruence.
+ eapply match_stacks_inside_inlined; eauto.
+ red; intros. destruct (zlt ofs (dstk ctx)). apply PAD; omega. apply H3. inv H4. xomega.
+ congruence.
unfold context_below in *. xomega.
- unfold context_stack_call in *. omega.
+ unfold context_stack_call in *. omega.
Qed.
(** ** Relating states *)
@@ -915,7 +915,7 @@ Lemma tr_funbody_inv:
forall sz cts f c pc i,
tr_funbody fenv sz cts f c -> f.(fn_code)!pc = Some i -> tr_instr fenv sz cts pc i c.
Proof.
- intros. inv H. eauto.
+ intros. inv H. eauto.
Qed.
Theorem step_simulation:
@@ -929,13 +929,13 @@ Proof.
(* nop *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
- left; econstructor; split.
+ left; econstructor; split.
eapply plus_one. eapply exec_Inop; eauto.
econstructor; eauto.
(* op *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
- exploit eval_operation_inject.
+ exploit eval_operation_inject.
eapply match_stacks_inside_globals; eauto.
eexact SP.
instantiate (2 := rs##args). instantiate (1 := rs'##(sregs ctx args)). eapply agree_val_regs; eauto.
@@ -943,14 +943,14 @@ Proof.
fold (sop ctx op). intros [v' [A B]].
left; econstructor; split.
eapply plus_one. eapply exec_Iop; eauto. erewrite eval_operation_preserved; eauto.
- exact symbols_preserved.
- econstructor; eauto.
+ exact symbols_preserved.
+ econstructor; eauto.
apply match_stacks_inside_set_reg; auto.
- apply agree_set_reg; auto.
-
+ apply agree_set_reg; auto.
+
(* load *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
- exploit eval_addressing_inject.
+ exploit eval_addressing_inject.
eapply match_stacks_inside_globals; eauto.
eexact SP.
instantiate (2 := rs##args). instantiate (1 := rs'##(sregs ctx args)). eapply agree_val_regs; eauto.
@@ -961,19 +961,19 @@ Proof.
rewrite <- P. apply eval_addressing_preserved. exact symbols_preserved.
left; econstructor; split.
eapply plus_one. eapply exec_Iload; eauto.
- econstructor; eauto.
+ econstructor; eauto.
apply match_stacks_inside_set_reg; auto.
- apply agree_set_reg; auto.
+ apply agree_set_reg; auto.
(* store *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
- exploit eval_addressing_inject.
+ exploit eval_addressing_inject.
eapply match_stacks_inside_globals; eauto.
eexact SP.
instantiate (2 := rs##args). instantiate (1 := rs'##(sregs ctx args)). eapply agree_val_regs; eauto.
eauto.
fold saddr. intros [a' [P Q]].
- exploit Mem.storev_mapped_inject; eauto. eapply agree_val_reg; eauto.
+ exploit Mem.storev_mapped_inject; eauto. eapply agree_val_reg; eauto.
intros [m1' [U V]].
assert (eval_addressing tge (Vptr sp' Int.zero) (saddr ctx addr) rs' ## (sregs ctx args) = Some a').
rewrite <- P. apply eval_addressing_preserved. exact symbols_preserved.
@@ -998,32 +998,32 @@ Proof.
eapply plus_one. eapply exec_Icall; eauto.
eapply sig_function_translated; eauto.
econstructor; eauto.
- eapply match_stacks_cons; eauto.
- eapply agree_val_regs; eauto.
+ eapply match_stacks_cons; eauto.
+ eapply agree_val_regs; eauto.
(* inlined *)
assert (fd = Internal f0).
simpl in H0. destruct (Genv.find_symbol ge id) as [b|] eqn:?; try discriminate.
- exploit (funenv_program_compat prog); eauto. intros.
+ exploit (funenv_program_compat prog); eauto. intros.
unfold ge in H0. congruence.
subst fd.
- right; split. simpl; omega. split. auto.
- econstructor; eauto.
+ right; split. simpl; omega. split. auto.
+ econstructor; eauto.
eapply match_stacks_inside_inlined; eauto.
red; intros. apply PRIV. inv H13. destruct H16. xomega.
apply agree_val_regs_gen; auto.
- red; intros; apply PRIV. destruct H16. omega.
+ red; intros; apply PRIV. destruct H16. omega.
(* tailcall *)
exploit match_stacks_inside_globalenvs; eauto. intros [bound G].
exploit find_function_agree; eauto. intros [fd' [A B]].
assert (PRIV': range_private F m' m'0 sp' (dstk ctx) f'.(fn_stacksize)).
- eapply range_private_free_left; eauto. inv FB. rewrite <- H4. auto.
+ eapply range_private_free_left; eauto. inv FB. rewrite <- H4. auto.
exploit tr_funbody_inv; eauto. intros TR; inv TR.
(* within the original function *)
inv MS0; try congruence.
assert (X: { m1' | Mem.free m'0 sp' 0 (fn_stacksize f') = Some m1'}).
apply Mem.range_perm_free. red; intros.
- destruct (zlt ofs f.(fn_stacksize)).
+ destruct (zlt ofs f.(fn_stacksize)).
replace ofs with (ofs + dstk ctx) by omega. eapply Mem.perm_inject; eauto.
eapply Mem.free_range_perm; eauto. omega.
inv FB. eapply range_private_perms; eauto. xomega.
@@ -1032,17 +1032,17 @@ Proof.
eapply plus_one. eapply exec_Itailcall; eauto.
eapply sig_function_translated; eauto.
econstructor; eauto.
- eapply match_stacks_bound with (bound := sp').
+ eapply match_stacks_bound with (bound := sp').
eapply match_stacks_invariant; eauto.
- intros. eapply Mem.perm_free_3; eauto.
- intros. eapply Mem.perm_free_1; eauto.
+ intros. eapply Mem.perm_free_3; eauto.
+ intros. eapply Mem.perm_free_1; eauto.
intros. eapply Mem.perm_free_3; eauto.
erewrite Mem.nextblock_free; eauto. red in VB; xomega.
eapply agree_val_regs; eauto.
eapply Mem.free_right_inject; eauto. eapply Mem.free_left_inject; eauto.
(* show that no valid location points into the stack block being freed *)
- intros. rewrite DSTK in PRIV'. exploit (PRIV' (ofs + delta)). omega. intros [P Q].
- eelim Q; eauto. replace (ofs + delta - delta) with ofs by omega.
+ intros. rewrite DSTK in PRIV'. exploit (PRIV' (ofs + delta)). omega. intros [P Q].
+ eelim Q; eauto. replace (ofs + delta - delta) with ofs by omega.
apply Mem.perm_max with k. apply Mem.perm_implies with p; auto with mem.
(* turned into a call *)
left; econstructor; split.
@@ -1050,68 +1050,68 @@ Proof.
eapply sig_function_translated; eauto.
econstructor; eauto.
eapply match_stacks_untailcall; eauto.
- eapply match_stacks_inside_invariant; eauto.
+ eapply match_stacks_inside_invariant; eauto.
intros. eapply Mem.perm_free_3; eauto.
eapply agree_val_regs; eauto.
eapply Mem.free_left_inject; eauto.
(* inlined *)
assert (fd = Internal f0).
simpl in H0. destruct (Genv.find_symbol ge id) as [b|] eqn:?; try discriminate.
- exploit (funenv_program_compat prog); eauto. intros.
+ exploit (funenv_program_compat prog); eauto. intros.
unfold ge in H0. congruence.
subst fd.
- right; split. simpl; omega. split. auto.
+ right; split. simpl; omega. split. auto.
econstructor; eauto.
eapply match_stacks_inside_inlined_tailcall; eauto.
eapply match_stacks_inside_invariant; eauto.
intros. eapply Mem.perm_free_3; eauto.
apply agree_val_regs_gen; auto.
eapply Mem.free_left_inject; eauto.
- red; intros; apply PRIV'.
+ red; intros; apply PRIV'.
assert (dstk ctx <= dstk ctx'). red in H14; rewrite H14. apply align_le. apply min_alignment_pos.
omega.
(* builtin *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
- exploit match_stacks_inside_globalenvs; eauto. intros [bound MG].
+ exploit match_stacks_inside_globalenvs; eauto. intros [bound MG].
exploit tr_builtin_args; eauto. intros (vargs' & P & Q).
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject; eauto.
eapply match_stacks_inside_globals; eauto.
intros [F1 [v1 [m1' [A [B [C [D [E [J K]]]]]]]]].
left; econstructor; split.
- eapply plus_one. eapply exec_Ibuiltin; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply plus_one. eapply exec_Ibuiltin; eauto.
+ eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
econstructor.
- eapply match_stacks_inside_set_res.
+ eapply match_stacks_inside_set_res.
eapply match_stacks_inside_extcall with (F1 := F) (F2 := F1) (m1 := m) (m1' := m'0); eauto.
- intros; eapply external_call_max_perm; eauto.
- intros; eapply external_call_max_perm; eauto.
+ intros; eapply external_call_max_perm; eauto.
+ intros; eapply external_call_max_perm; eauto.
auto.
destruct res; simpl; [apply agree_set_reg;auto|idtac|idtac]; eapply agree_regs_incr; eauto.
auto. auto.
- eapply external_call_valid_block; eauto.
- eapply range_private_extcall; eauto.
- intros; eapply external_call_max_perm; eauto.
- auto.
- intros. apply SSZ2. eapply external_call_max_perm; eauto.
+ eapply external_call_valid_block; eauto.
+ eapply range_private_extcall; eauto.
+ intros; eapply external_call_max_perm; eauto.
+ auto.
+ intros. apply SSZ2. eapply external_call_max_perm; eauto.
(* cond *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
assert (eval_condition cond rs'##(sregs ctx args) m' = Some b).
- eapply eval_condition_inject; eauto. eapply agree_val_regs; eauto.
+ eapply eval_condition_inject; eauto. eapply agree_val_regs; eauto.
left; econstructor; split.
- eapply plus_one. eapply exec_Icond; eauto.
- destruct b; econstructor; eauto.
+ eapply plus_one. eapply exec_Icond; eauto.
+ destruct b; econstructor; eauto.
(* jumptable *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
assert (Val.inject F rs#arg rs'#(sreg ctx arg)). eapply agree_val_reg; eauto.
- rewrite H0 in H2; inv H2.
+ rewrite H0 in H2; inv H2.
left; econstructor; split.
eapply plus_one. eapply exec_Ijumptable; eauto.
- rewrite list_nth_z_map. rewrite H1. simpl; reflexivity.
- econstructor; eauto.
+ rewrite list_nth_z_map. rewrite H1. simpl; reflexivity.
+ econstructor; eauto.
(* return *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
@@ -1119,19 +1119,19 @@ Proof.
inv MS0; try congruence.
assert (X: { m1' | Mem.free m'0 sp' 0 (fn_stacksize f') = Some m1'}).
apply Mem.range_perm_free. red; intros.
- destruct (zlt ofs f.(fn_stacksize)).
+ destruct (zlt ofs f.(fn_stacksize)).
replace ofs with (ofs + dstk ctx) by omega. eapply Mem.perm_inject; eauto.
eapply Mem.free_range_perm; eauto. omega.
inv FB. eapply range_private_perms; eauto.
generalize (Zmax_spec (fn_stacksize f) 0). destruct (zlt 0 (fn_stacksize f)); omega.
destruct X as [m1' FREE].
left; econstructor; split.
- eapply plus_one. eapply exec_Ireturn; eauto.
+ eapply plus_one. eapply exec_Ireturn; eauto.
econstructor; eauto.
- eapply match_stacks_bound with (bound := sp').
+ eapply match_stacks_bound with (bound := sp').
eapply match_stacks_invariant; eauto.
- intros. eapply Mem.perm_free_3; eauto.
- intros. eapply Mem.perm_free_1; eauto.
+ intros. eapply Mem.perm_free_3; eauto.
+ intros. eapply Mem.perm_free_1; eauto.
intros. eapply Mem.perm_free_3; eauto.
erewrite Mem.nextblock_free; eauto. red in VB; xomega.
destruct or; simpl. apply agree_val_reg; auto. auto.
@@ -1139,58 +1139,58 @@ Proof.
(* show that no valid location points into the stack block being freed *)
intros. inversion FB; subst.
assert (PRIV': range_private F m' m'0 sp' (dstk ctx) f'.(fn_stacksize)).
- rewrite H8 in PRIV. eapply range_private_free_left; eauto.
- rewrite DSTK in PRIV'. exploit (PRIV' (ofs + delta)). omega. intros [A B].
- eelim B; eauto. replace (ofs + delta - delta) with ofs by omega.
+ rewrite H8 in PRIV. eapply range_private_free_left; eauto.
+ rewrite DSTK in PRIV'. exploit (PRIV' (ofs + delta)). omega. intros [A B].
+ eelim B; eauto. replace (ofs + delta - delta) with ofs by omega.
apply Mem.perm_max with k. apply Mem.perm_implies with p; auto with mem.
(* inlined *)
- right. split. simpl. omega. split. auto.
+ right. split. simpl. omega. split. auto.
econstructor; eauto.
- eapply match_stacks_inside_invariant; eauto.
+ eapply match_stacks_inside_invariant; eauto.
intros. eapply Mem.perm_free_3; eauto.
destruct or; simpl. apply agree_val_reg; auto. auto.
eapply Mem.free_left_inject; eauto.
- inv FB. rewrite H4 in PRIV. eapply range_private_free_left; eauto.
+ inv FB. rewrite H4 in PRIV. eapply range_private_free_left; eauto.
(* internal function, not inlined *)
- assert (A: exists f', tr_function fenv f f' /\ fd' = Internal f').
- Errors.monadInv FD. exists x. split; auto. eapply transf_function_spec; eauto.
+ assert (A: exists f', tr_function fenv f f' /\ fd' = Internal f').
+ Errors.monadInv FD. exists x. split; auto. eapply transf_function_spec; eauto.
destruct A as [f' [TR EQ]]. inversion TR; subst.
- exploit Mem.alloc_parallel_inject. eauto. eauto. apply Zle_refl.
- instantiate (1 := fn_stacksize f'). inv H0. xomega.
+ exploit Mem.alloc_parallel_inject. eauto. eauto. apply Zle_refl.
+ instantiate (1 := fn_stacksize f'). inv H0. xomega.
intros [F' [m1' [sp' [A [B [C [D E]]]]]]].
left; econstructor; split.
eapply plus_one. eapply exec_function_internal; eauto.
rewrite H5. econstructor.
instantiate (1 := F'). apply match_stacks_inside_base.
assert (SP: sp' = Mem.nextblock m'0) by (eapply Mem.alloc_result; eauto).
- rewrite <- SP in MS0.
+ rewrite <- SP in MS0.
eapply match_stacks_invariant; eauto.
- intros. destruct (eq_block b1 stk).
- subst b1. rewrite D in H7; inv H7. subst b2. eelim Plt_strict; eauto.
- rewrite E in H7; auto.
- intros. exploit Mem.perm_alloc_inv. eexact H. eauto.
- destruct (eq_block b1 stk); intros; auto.
- subst b1. rewrite D in H7; inv H7. subst b2. eelim Plt_strict; eauto.
- intros. eapply Mem.perm_alloc_1; eauto.
- intros. exploit Mem.perm_alloc_inv. eexact A. eauto.
+ intros. destruct (eq_block b1 stk).
+ subst b1. rewrite D in H7; inv H7. subst b2. eelim Plt_strict; eauto.
+ rewrite E in H7; auto.
+ intros. exploit Mem.perm_alloc_inv. eexact H. eauto.
+ destruct (eq_block b1 stk); intros; auto.
+ subst b1. rewrite D in H7; inv H7. subst b2. eelim Plt_strict; eauto.
+ intros. eapply Mem.perm_alloc_1; eauto.
+ intros. exploit Mem.perm_alloc_inv. eexact A. eauto.
rewrite dec_eq_false; auto.
- auto. auto. auto.
+ auto. auto. auto.
rewrite H4. apply agree_regs_init_regs. eauto. auto. inv H0; auto. congruence. auto.
eapply Mem.valid_new_block; eauto.
red; intros. split.
eapply Mem.perm_alloc_2; eauto. inv H0; xomega.
intros; red; intros. exploit Mem.perm_alloc_inv. eexact H. eauto.
- destruct (eq_block b stk); intros.
+ destruct (eq_block b stk); intros.
subst. rewrite D in H8; inv H8. inv H0; xomega.
rewrite E in H8; auto. eelim Mem.fresh_block_alloc. eexact A. eapply Mem.mi_mappedblocks; eauto.
auto.
- intros. exploit Mem.perm_alloc_inv; eauto. rewrite dec_eq_true. omega.
+ intros. exploit Mem.perm_alloc_inv; eauto. rewrite dec_eq_true. omega.
(* internal function, inlined *)
inversion FB; subst.
- exploit Mem.alloc_left_mapped_inject.
+ exploit Mem.alloc_left_mapped_inject.
eauto.
eauto.
(* sp' is valid *)
@@ -1212,7 +1212,7 @@ Proof.
apply Mem.perm_max with k. apply Mem.perm_implies with p; auto with mem.
intros [F' [A [B [C D]]]].
exploit tr_moves_init_regs; eauto. intros [rs'' [P [Q R]]].
- left; econstructor; split.
+ left; econstructor; split.
eapply plus_left. eapply exec_Inop; eauto. eexact P. traceEq.
econstructor.
eapply match_stacks_inside_alloc_left; eauto.
@@ -1226,21 +1226,21 @@ Proof.
(* external function *)
exploit match_stacks_globalenvs; eauto. intros [bound MG].
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject; eauto.
eapply match_globalenvs_preserves_globals; eauto.
intros [F1 [v1 [m1' [A [B [C [D [E [J K]]]]]]]]].
- simpl in FD. inv FD.
+ simpl in FD. inv FD.
left; econstructor; split.
- eapply plus_one. eapply exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
+ eapply plus_one. eapply exec_function_external; eauto.
+ eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
econstructor.
eapply match_stacks_bound with (Mem.nextblock m'0).
eapply match_stacks_extcall with (F1 := F) (F2 := F1) (m1 := m) (m1' := m'0); eauto.
- intros; eapply external_call_max_perm; eauto.
+ intros; eapply external_call_max_perm; eauto.
intros; eapply external_call_max_perm; eauto.
xomega.
- eapply external_call_nextblock; eauto.
+ eapply external_call_nextblock; eauto.
auto. auto.
(* return fron noninlined function *)
@@ -1248,8 +1248,8 @@ Proof.
(* normal case *)
left; econstructor; split.
eapply plus_one. eapply exec_return.
- econstructor; eauto.
- apply match_stacks_inside_set_reg; auto.
+ econstructor; eauto.
+ apply match_stacks_inside_set_reg; auto.
apply agree_set_reg; auto.
(* untailcall case *)
inv MS; try congruence.
@@ -1261,26 +1261,26 @@ Proof.
*)
left; econstructor; split.
eapply plus_one. eapply exec_return.
- eapply match_regular_states.
+ eapply match_regular_states.
eapply match_stacks_inside_set_reg; eauto.
- auto.
+ auto.
apply agree_set_reg; auto.
auto. auto. auto.
red; intros. destruct (zlt ofs (dstk ctx)). apply PAD; omega. apply PRIV; omega.
- auto. auto.
-
+ auto. auto.
+
(* return from inlined function *)
- inv MS0; try congruence. rewrite RET0 in RET; inv RET.
- unfold inline_return in AT.
+ inv MS0; try congruence. rewrite RET0 in RET; inv RET.
+ unfold inline_return in AT.
assert (PRIV': range_private F m m' sp' (dstk ctx' + mstk ctx') f'.(fn_stacksize)).
red; intros. destruct (zlt ofs (dstk ctx)). apply PAD. omega. apply PRIV. omega.
destruct or.
(* with a result *)
- left; econstructor; split.
- eapply plus_one. eapply exec_Iop; eauto. simpl. reflexivity.
+ left; econstructor; split.
+ eapply plus_one. eapply exec_Iop; eauto. simpl. reflexivity.
econstructor; eauto. apply match_stacks_inside_set_reg; auto. apply agree_set_reg; auto.
(* without a result *)
- left; econstructor; split.
+ left; econstructor; split.
eapply plus_one. eapply exec_Inop; eauto.
econstructor; eauto. subst vres. apply agree_set_reg_undef'; auto.
Qed.
@@ -1293,29 +1293,29 @@ Proof.
exists (Callstate nil tf nil m0); split.
econstructor; eauto.
unfold transf_program in TRANSF. eapply Genv.init_mem_transf_partial; eauto.
- rewrite symbols_preserved.
+ rewrite symbols_preserved.
rewrite (transform_partial_program_main _ _ TRANSF). auto.
- rewrite <- H3. apply sig_function_translated; auto.
- econstructor; eauto.
- instantiate (1 := Mem.flat_inj (Mem.nextblock m0)).
+ rewrite <- H3. apply sig_function_translated; auto.
+ econstructor; eauto.
+ instantiate (1 := Mem.flat_inj (Mem.nextblock m0)).
apply match_stacks_nil with (Mem.nextblock m0).
- constructor; intros.
- unfold Mem.flat_inj. apply pred_dec_true; auto.
+ constructor; intros.
+ unfold Mem.flat_inj. apply pred_dec_true; auto.
unfold Mem.flat_inj in H. destruct (plt b1 (Mem.nextblock m0)); congruence.
eapply Genv.find_symbol_not_fresh; eauto.
eapply Genv.find_funct_ptr_not_fresh; eauto.
- eapply Genv.find_var_info_not_fresh; eauto.
- apply Ple_refl.
+ eapply Genv.find_var_info_not_fresh; eauto.
+ apply Ple_refl.
eapply Genv.initmem_inject; eauto.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
intros. inv H0. inv H.
exploit match_stacks_empty; eauto. intros EQ; subst. inv VINJ. constructor.
- exploit match_stacks_inside_empty; eauto. intros [A B]. congruence.
+ exploit match_stacks_inside_empty; eauto. intros [A B]. congruence.
Qed.
Theorem transf_program_correct:
@@ -1325,7 +1325,7 @@ Proof.
eexact public_preserved.
eexact transf_initial_states.
eexact transf_final_states.
- eexact step_simulation.
+ eexact step_simulation.
Qed.
End INLINING.
diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v
index 161e2a6e..ba62313f 100644
--- a/backend/Inliningspec.v
+++ b/backend/Inliningspec.v
@@ -40,24 +40,24 @@ Remark add_globdef_compat:
fenv_compat (Genv.add_global ge idg) (Inlining.add_globdef fenv idg).
Proof.
intros. destruct idg as [id gd]. red; simpl; intros.
- unfold Genv.find_symbol in H1; simpl in H1.
+ unfold Genv.find_symbol in H1; simpl in H1.
unfold Genv.find_funct_ptr; simpl.
rewrite PTree.gsspec in H1. destruct (peq id0 id).
(* same *)
- subst id0. inv H1. destruct gd. destruct f0.
+ subst id0. inv H1. destruct gd. destruct f0.
destruct (should_inline id f0).
rewrite PTree.gss in H0. rewrite PTree.gss. inv H0; auto.
rewrite PTree.grs in H0; discriminate.
rewrite PTree.grs in H0; discriminate.
rewrite PTree.grs in H0; discriminate.
(* different *)
- destruct gd. rewrite PTree.gso. eapply H; eauto.
+ destruct gd. rewrite PTree.gso. eapply H; eauto.
destruct f0. destruct (should_inline id f0).
rewrite PTree.gso in H0; auto.
rewrite PTree.gro in H0; auto.
rewrite PTree.gro in H0; auto.
red; intros; subst b. eelim Plt_strict. eapply Genv.genv_symb_range; eauto.
- rewrite PTree.gro in H0; auto. eapply H; eauto.
+ rewrite PTree.gro in H0; auto. eapply H; eauto.
Qed.
Lemma funenv_program_compat:
@@ -68,7 +68,7 @@ Proof.
assert (forall gl ge fenv,
fenv_compat ge fenv ->
fenv_compat (Genv.add_globals ge gl) (fold_left add_globdef gl fenv)).
- induction gl; simpl; intros. auto. apply IHgl. apply add_globdef_compat; auto.
+ induction gl; simpl; intros. auto. apply IHgl. apply add_globdef_compat; auto.
apply H. red; intros. rewrite PTree.gempty in H0; discriminate.
Qed.
@@ -80,12 +80,12 @@ Proof.
zify. omega.
Qed.
-Lemma shiftpos_inj:
+Lemma shiftpos_inj:
forall x y n, shiftpos x n = shiftpos y n -> x = y.
Proof.
intros.
assert (Zpos (shiftpos x n) = Zpos (shiftpos y n)) by congruence.
- rewrite ! shiftpos_eq in H0.
+ rewrite ! shiftpos_eq in H0.
assert (Z.pos x = Z.pos y) by omega.
congruence.
Qed.
@@ -99,32 +99,32 @@ Qed.
Lemma shiftpos_above:
forall x n, Ple n (shiftpos x n).
Proof.
- intros. unfold Ple; zify. rewrite shiftpos_eq. xomega.
+ intros. unfold Ple; zify. rewrite shiftpos_eq. xomega.
Qed.
Lemma shiftpos_not_below:
forall x n, Plt (shiftpos x n) n -> False.
Proof.
- intros. generalize (shiftpos_above x n). xomega.
+ intros. generalize (shiftpos_above x n). xomega.
Qed.
Lemma shiftpos_below:
forall x n, Plt (shiftpos x n) (Pplus x n).
Proof.
- intros. unfold Plt; zify. rewrite shiftpos_eq. omega.
+ intros. unfold Plt; zify. rewrite shiftpos_eq. omega.
Qed.
Lemma shiftpos_le:
forall x y n, Ple x y -> Ple (shiftpos x n) (shiftpos y n).
Proof.
- intros. unfold Ple in *; zify. rewrite ! shiftpos_eq. omega.
+ intros. unfold Ple in *; zify. rewrite ! shiftpos_eq. omega.
Qed.
(** ** Working with the state monad *)
Remark bind_inversion:
- forall (A B: Type) (f: mon A) (g: A -> mon B)
+ forall (A B: Type) (f: mon A) (g: A -> mon B)
(y: B) (s1 s3: state) (i: sincr s1 s3),
bind f g s1 = R y s3 i ->
exists x, exists s2, exists i1, exists i2,
@@ -156,7 +156,7 @@ Ltac monadInv H :=
match type of H with
| (ret _ _ = R _ _ _) => monadInv1 H
| (bind ?F ?G ?S = R ?X ?S' ?I) => monadInv1 H
- | (?F _ _ _ _ _ _ _ _ = R _ _ _) =>
+ | (?F _ _ _ _ _ _ _ _ = R _ _ _) =>
((progress simpl in H) || unfold F in H); monadInv1 H
| (?F _ _ _ _ _ _ _ = R _ _ _) =>
((progress simpl in H) || unfold F in H); monadInv1 H
@@ -189,7 +189,7 @@ 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.
+ destruct (IHl s1) as [i2 EQ]. rewrite EQ. econstructor; eauto.
Qed.
Lemma ptree_mfold_spec:
@@ -197,7 +197,7 @@ Lemma ptree_mfold_spec:
ptree_mfold f t s = R x s' i ->
exists i', mlist_iter2 f (PTree.elements t) s = R tt s' i'.
Proof.
- intros.
+ 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.
@@ -220,12 +220,12 @@ Lemma add_moves_unchanged:
Plt pc s.(st_nextnode) \/ Ple s'.(st_nextnode) pc ->
s'.(st_code)!pc = s.(st_code)!pc.
Proof.
- induction srcs; simpl; intros.
+ induction srcs; simpl; intros.
monadInv H. auto.
destruct dsts; monadInv H. auto.
- transitivity (st_code s0)!pc. eapply IHsrcs; eauto. monadInv EQ; simpl. xomega.
+ transitivity (st_code s0)!pc. eapply IHsrcs; eauto. monadInv EQ; simpl. xomega.
monadInv EQ; simpl. apply PTree.gso.
- inversion INCR0; simpl in *. xomega.
+ inversion INCR0; simpl in *. xomega.
Qed.
Lemma add_moves_spec:
@@ -234,15 +234,15 @@ Lemma add_moves_spec:
(forall pc, Ple s.(st_nextnode) pc -> Plt pc s'.(st_nextnode) -> c!pc = s'.(st_code)!pc) ->
tr_moves c pc1 srcs dsts pc2.
Proof.
- induction srcs; simpl; intros.
+ induction srcs; simpl; intros.
monadInv H. apply tr_moves_nil; auto.
- destruct dsts; monadInv H. apply tr_moves_nil; auto.
- apply tr_moves_cons with x. eapply IHsrcs; eauto.
+ destruct dsts; monadInv H. apply tr_moves_nil; auto.
+ apply tr_moves_cons with x. eapply IHsrcs; eauto.
intros. inversion INCR. apply H0; xomega.
monadInv EQ.
- rewrite H0. erewrite add_moves_unchanged; eauto.
- simpl. apply PTree.gss.
- simpl. xomega.
+ rewrite H0. erewrite add_moves_unchanged; eauto.
+ simpl. apply PTree.gss.
+ simpl. xomega.
xomega.
inversion INCR; inversion INCR0; simpl in *; xomega.
Qed.
@@ -386,25 +386,25 @@ Proof.
generalize set_instr_other; intros A.
intros. unfold expand_instr in H; destruct instr; eauto.
(* call *)
- destruct (can_inline fe s1). eauto.
+ destruct (can_inline fe s1). eauto.
monadInv H. unfold inline_function in EQ. monadInv EQ.
- transitivity (s2.(st_code)!pc'). eauto.
+ transitivity (s2.(st_code)!pc'). eauto.
transitivity (s5.(st_code)!pc'). eapply add_moves_unchanged; eauto.
- left. inversion INCR5. inversion INCR3. monadInv EQ1; simpl in *. xomega.
- transitivity (s4.(st_code)!pc'). eapply rec_unchanged; eauto.
+ left. inversion INCR5. inversion INCR3. monadInv EQ1; simpl in *. xomega.
+ transitivity (s4.(st_code)!pc'). eapply rec_unchanged; eauto.
simpl. monadInv EQ; simpl. monadInv EQ1; simpl. xomega.
- simpl. monadInv EQ1; simpl. auto.
+ simpl. monadInv EQ1; simpl. auto.
monadInv EQ; simpl. monadInv EQ1; simpl. auto.
(* tailcall *)
destruct (can_inline fe s1).
destruct (retinfo ctx) as [[rpc rreg]|]; eauto.
monadInv H. unfold inline_tail_function in EQ. monadInv EQ.
- transitivity (s2.(st_code)!pc'). eauto.
+ transitivity (s2.(st_code)!pc'). eauto.
transitivity (s5.(st_code)!pc'). eapply add_moves_unchanged; eauto.
- left. inversion INCR5. inversion INCR3. monadInv EQ1; simpl in *. xomega.
- transitivity (s4.(st_code)!pc'). eapply rec_unchanged; eauto.
+ left. inversion INCR5. inversion INCR3. monadInv EQ1; simpl in *. xomega.
+ transitivity (s4.(st_code)!pc'). eapply rec_unchanged; eauto.
simpl. monadInv EQ; simpl. monadInv EQ1; simpl. xomega.
- simpl. monadInv EQ1; simpl. auto.
+ simpl. monadInv EQ1; simpl. auto.
monadInv EQ; simpl. monadInv EQ1; simpl. auto.
(* return *)
destruct (retinfo ctx) as [[rpc rreg]|]; eauto.
@@ -425,8 +425,8 @@ Proof.
(* inductive case *)
destruct a as [pc1 instr1]; simpl in *.
monadInv H. inv H3.
- transitivity ((st_code s0)!pc).
- eapply IHl; eauto. destruct INCR; xomega. destruct INCR; xomega.
+ transitivity ((st_code s0)!pc).
+ eapply IHl; eauto. destruct INCR; xomega. destruct INCR; xomega.
eapply expand_instr_unchanged; eauto.
Qed.
@@ -439,12 +439,12 @@ 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.
+ exploit ptree_mfold_spec; eauto. intros [INCR' ITER].
+ eapply iter_expand_instr_unchanged; eauto.
+ subst s0; auto.
subst s0; simpl. xomega.
- red; intros. exploit list_in_map_inv; eauto. intros [pc1 [A B]].
- subst pc. unfold spc in H1. eapply shiftpos_not_below; eauto.
+ red; intros. exploit list_in_map_inv; eauto. intros [pc1 [A B]].
+ subst pc. unfold spc in H1. eapply shiftpos_not_below; eauto.
apply PTree.elements_keys_norepet.
subst s0; auto.
Qed.
@@ -456,7 +456,7 @@ Hypothesis rec_spec:
Ple (ctx.(dpc) + max_pc_function f) s.(st_nextnode) ->
ctx.(mreg) = max_reg_function f ->
Ple (Pplus ctx.(dreg) ctx.(mreg)) s.(st_nextreg) ->
- ctx.(mstk) >= 0 ->
+ ctx.(mstk) >= 0 ->
ctx.(mstk) = Zmax (fn_stacksize f) 0 ->
(min_alignment (fn_stacksize f) | ctx.(dstk)) ->
ctx.(dstk) >= 0 ->
@@ -496,14 +496,14 @@ Proof.
(* call *)
destruct (can_inline fe s1) as [|id f P Q].
(* not inlined *)
- eapply tr_call; eauto.
+ eapply tr_call; eauto.
(* inlined *)
subst s1.
monadInv EXP. unfold inline_function in EQ; monadInv EQ.
set (ctx' := callcontext ctx x1 x2 (max_reg_function f) (fn_stacksize f) n r).
- inversion EQ0; inversion EQ1; inversion EQ. inv_incr.
+ inversion EQ0; inversion EQ1; inversion EQ. inv_incr.
apply tr_call_inlined with (pc1 := x0) (ctx' := ctx') (f := f); auto.
- eapply BASE; eauto.
+ eapply BASE; eauto.
eapply add_moves_spec; eauto.
intros. rewrite S1. eapply set_instr_other; eauto. unfold node; xomega.
xomega. xomega.
@@ -517,24 +517,24 @@ Proof.
omega.
intros. simpl in H. rewrite S1.
transitivity s1.(st_code)!pc0. eapply set_instr_other; eauto. unfold node in *; xomega.
- eapply add_moves_unchanged; eauto. unfold node in *; xomega. xomega.
- red; simpl. subst s2; simpl in *. xomega.
+ eapply add_moves_unchanged; eauto. unfold node in *; xomega. xomega.
+ red; simpl. subst s2; simpl in *. xomega.
red; simpl. split. auto. apply align_le. apply min_alignment_pos.
(* tailcall *)
destruct (can_inline fe s1) as [|id f P Q].
(* not inlined *)
- destruct (retinfo ctx) as [[rpc rreg] | ] eqn:?.
+ destruct (retinfo ctx) as [[rpc rreg] | ] eqn:?.
(* turned into a call *)
- eapply tr_tailcall_call; eauto.
+ eapply tr_tailcall_call; eauto.
(* preserved *)
- eapply tr_tailcall; eauto.
+ eapply tr_tailcall; eauto.
(* inlined *)
subst s1.
monadInv EXP. unfold inline_function in EQ; monadInv EQ.
set (ctx' := tailcontext ctx x1 x2 (max_reg_function f) (fn_stacksize f)) in *.
- inversion EQ0; inversion EQ1; inversion EQ. inv_incr.
+ inversion EQ0; inversion EQ1; inversion EQ. inv_incr.
apply tr_tailcall_inlined with (pc1 := x0) (ctx' := ctx') (f := f); auto.
- eapply BASE; eauto.
+ eapply BASE; eauto.
eapply add_moves_spec; eauto.
intros. rewrite S1. eapply set_instr_other; eauto. unfold node; xomega. xomega. xomega.
eapply rec_spec; eauto.
@@ -547,18 +547,18 @@ Proof.
omega.
intros. simpl in H. rewrite S1.
transitivity s1.(st_code)!pc0. eapply set_instr_other; eauto. unfold node in *; xomega.
- eapply add_moves_unchanged; eauto. unfold node in *; xomega. xomega.
- red; simpl.
+ eapply add_moves_unchanged; eauto. unfold node in *; xomega. xomega.
+ red; simpl.
subst s2; simpl in *; xomega.
red; auto.
(* builtin *)
- eapply tr_builtin; eauto. destruct b; eauto.
+ eapply tr_builtin; eauto. destruct b; eauto.
(* return *)
- destruct (retinfo ctx) as [[rpc rreg] | ] eqn:?.
+ destruct (retinfo ctx) as [[rpc rreg] | ] eqn:?.
(* inlined *)
- eapply tr_return_inlined; eauto.
+ eapply tr_return_inlined; eauto.
(* unchanged *)
- eapply tr_return; eauto.
+ eapply tr_return; eauto.
Qed.
Lemma iter_expand_instr_spec:
@@ -580,29 +580,29 @@ Proof.
(* inductive case *)
destruct a as [pc1 instr1]; simpl in *. inv H0. monadInv H. inv_incr.
assert (A: Ple ctx.(dpc) s0.(st_nextnode)).
- assert (B: Plt (spc ctx pc) (st_nextnode s)) by eauto.
+ assert (B: Plt (spc ctx pc) (st_nextnode s)) by eauto.
unfold spc in B. generalize (shiftpos_above pc (dpc ctx)). xomega.
destruct H9. inv H.
(* same pc *)
eapply expand_instr_spec; eauto.
omega.
intros.
- transitivity ((st_code s')!pc').
- apply H7. auto. xomega.
- eapply iter_expand_instr_unchanged; eauto.
- red; intros. rewrite list_map_compose in H9. exploit list_in_map_inv; eauto.
- intros [[pc0 instr0] [P Q]]. simpl in P.
+ transitivity ((st_code s')!pc').
+ apply H7. auto. xomega.
+ eapply iter_expand_instr_unchanged; eauto.
+ red; intros. rewrite list_map_compose in H9. exploit list_in_map_inv; eauto.
+ intros [[pc0 instr0] [P Q]]. simpl in P.
assert (Plt (spc ctx pc0) (st_nextnode s)) by eauto. xomega.
- transitivity ((st_code s')!(spc ctx pc)).
- eapply H8; eauto.
- eapply iter_expand_instr_unchanged; eauto.
+ transitivity ((st_code s')!(spc ctx pc)).
+ eapply H8; eauto.
+ eapply iter_expand_instr_unchanged; eauto.
assert (Plt (spc ctx pc) (st_nextnode s)) by eauto. xomega.
- red; intros. rewrite list_map_compose in H. exploit list_in_map_inv; eauto.
+ red; intros. rewrite list_map_compose in H. exploit list_in_map_inv; eauto.
intros [[pc0 instr0] [P Q]]. simpl in P.
assert (pc = pc0) by (eapply shiftpos_inj; eauto). subst pc0.
elim H12. change pc with (fst (pc, instr0)). apply List.in_map; auto.
(* older pc *)
- inv_incr. eapply IHl; eauto.
+ inv_incr. eapply IHl; eauto.
intros. eapply Plt_le_trans. eapply H2. right; eauto. xomega.
intros; eapply Ple_trans; eauto.
intros. apply H7; auto. xomega.
@@ -614,7 +614,7 @@ Lemma expand_cfg_rec_spec:
Ple (ctx.(dpc) + max_pc_function f) s.(st_nextnode) ->
ctx.(mreg) = max_reg_function f ->
Ple (ctx.(dreg) + ctx.(mreg)) s.(st_nextreg) ->
- ctx.(mstk) >= 0 ->
+ ctx.(mstk) >= 0 ->
ctx.(mstk) = Zmax (fn_stacksize f) 0 ->
(min_alignment (fn_stacksize f) | ctx.(dstk)) ->
ctx.(dstk) >= 0 ->
@@ -622,13 +622,13 @@ Lemma expand_cfg_rec_spec:
(forall pc', Ple ctx.(dpc) pc' -> Plt pc' s'.(st_nextnode) -> c!pc' = s'.(st_code)!pc') ->
tr_funbody ctx f c.
Proof.
- intros. unfold expand_cfg_rec in H. monadInv H. inversion EQ.
- constructor.
- intros. rewrite H1. eapply max_reg_function_params; eauto.
+ intros. unfold expand_cfg_rec in H. monadInv H. inversion EQ.
+ constructor.
+ intros. rewrite H1. eapply max_reg_function_params; 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_reg_function_def with (i := instr); eauto.
+ eapply iter_expand_instr_spec; eauto.
+ apply PTree.elements_keys_norepet.
+ intros. rewrite H1. eapply max_reg_function_def with (i := instr); eauto.
eapply PTree.elements_complete; eauto.
intros.
assert (Ple pc0 (max_pc_function f)).
@@ -636,10 +636,10 @@ Proof.
eapply Plt_le_trans. apply shiftpos_below. subst s0; simpl; xomega.
subst s0; simpl; auto.
intros. apply H8; auto. subst s0; simpl in H11; xomega.
- intros. apply H8. apply shiftpos_above.
+ intros. apply H8. apply shiftpos_above.
assert (Ple pc0 (max_pc_function f)).
- eapply max_pc_function_sound. eapply PTree.elements_complete; eauto.
- eapply Plt_le_trans. apply shiftpos_below. inversion i; xomega.
+ eapply max_pc_function_sound. eapply PTree.elements_complete; eauto.
+ eapply Plt_le_trans. apply shiftpos_below. inversion i; xomega.
apply PTree.elements_correct; auto.
auto. auto. auto.
inversion INCR0. subst s0; simpl in STKSIZE; xomega.
@@ -657,17 +657,17 @@ Proof.
intros fe0; pattern fe0. apply well_founded_ind with (R := ltof _ size_fenv).
apply well_founded_ltof.
intros. unfold expand_cfg in H0. rewrite unroll_Fixm in H0.
- eapply expand_cfg_rec_unchanged; eauto. assumption.
+ eapply expand_cfg_rec_unchanged; eauto. assumption.
Qed.
Lemma expand_cfg_spec:
forall fe ctx f s x s' i c,
expand_cfg fe ctx f s = R x s' i ->
- fenv_agree fe ->
+ fenv_agree fe ->
Ple (ctx.(dpc) + max_pc_function f) s.(st_nextnode) ->
ctx.(mreg) = max_reg_function f ->
Ple (ctx.(dreg) + ctx.(mreg)) s.(st_nextreg) ->
- ctx.(mstk) >= 0 ->
+ ctx.(mstk) >= 0 ->
ctx.(mstk) = Zmax (fn_stacksize f) 0 ->
(min_alignment (fn_stacksize f) | ctx.(dstk)) ->
ctx.(dstk) >= 0 ->
@@ -678,7 +678,7 @@ Proof.
intros fe0; pattern fe0. apply well_founded_ind with (R := ltof _ size_fenv).
apply well_founded_ltof.
intros. unfold expand_cfg in H0. rewrite unroll_Fixm in H0.
- eapply expand_cfg_rec_spec; eauto.
+ eapply expand_cfg_rec_spec; eauto.
simpl. intros. eapply expand_cfg_unchanged; eauto. assumption.
Qed.
@@ -701,7 +701,7 @@ Lemma transf_function_spec:
forall f f', transf_function fenv f = OK f' -> tr_function f f'.
Proof.
intros. unfold transf_function in H.
- destruct (expand_function fenv f initstate) as [ctx s i] eqn:?.
+ destruct (expand_function fenv f initstate) as [ctx s i] eqn:?.
destruct (zlt (st_stksize s) Int.max_unsigned); inv H.
monadInv Heqr. set (ctx := initcontext x x0 (max_reg_function f) (fn_stacksize f)) in *.
Opaque initstate.
@@ -712,11 +712,11 @@ Opaque initstate.
unfold ctx; rewrite <- H1; rewrite <- H2; rewrite <- H3; simpl. xomega.
unfold ctx; rewrite <- H0; rewrite <- H1; simpl. xomega.
simpl. xomega.
- simpl. apply Zdivide_0.
+ simpl. apply Zdivide_0.
simpl. omega.
simpl. omega.
- simpl. split; auto. destruct INCR2. destruct INCR1. destruct INCR0. destruct INCR.
- simpl. change 0 with (st_stksize initstate). omega.
+ simpl. split; auto. destruct INCR2. destruct INCR1. destruct INCR0. destruct INCR.
+ simpl. change 0 with (st_stksize initstate). omega.
Qed.
End INLINING_SPEC.
diff --git a/backend/Kildall.v b/backend/Kildall.v
index 0d414d28..87090f5d 100644
--- a/backend/Kildall.v
+++ b/backend/Kildall.v
@@ -22,21 +22,21 @@ Local Unset Elimination Schemes.
Local Unset Case Analysis Schemes.
(** A forward dataflow problem is a set of inequations of the form
-- [X(s) >= transf n X(n)]
+- [X(s) >= transf n X(n)]
if program point [s] is a successor of program point [n]
- [X(n) >= a]
if [n] is an entry point and [a] its minimal approximation.
The unknowns are the [X(n)], indexed by program points (e.g. nodes in the
-CFG graph of a RTL function). They range over a given ordered set that
+CFG graph of a RTL function). They range over a given ordered set that
represents static approximations of the program state at each point.
-The [transf] function is the abstract transfer function: it computes an
+The [transf] function is the abstract transfer function: it computes an
approximation [transf n X(n)] of the program state after executing instruction
at point [n], as a function of the approximation [X(n)] of the program state
before executing that instruction.
Symmetrically, a backward dataflow problem is a set of inequations of the form
-- [X(n) >= transf s X(s)]
+- [X(n) >= transf s X(s)]
if program point [s] is a successor of program point [n]
- [X(n) >= a]
if [n] is an entry point and [a] its minimal approximation.
@@ -155,7 +155,7 @@ Context {A: Type} (code: PTree.t A) (successors: A -> list positive).
Inductive reachable: positive -> positive -> Prop :=
| reachable_refl: forall n, reachable n n
| reachable_left: forall n1 n2 n3 i,
- code!n1 = Some i -> In n2 (successors i) -> reachable n2 n3 ->
+ code!n1 = Some i -> In n2 (successors i) -> reachable n2 n3 ->
reachable n1 n3.
Scheme reachable_ind := Induction for reachable Sort Prop.
@@ -163,9 +163,9 @@ Scheme reachable_ind := Induction for reachable Sort Prop.
Lemma reachable_trans:
forall n1 n2, reachable n1 n2 -> forall n3, reachable n2 n3 -> reachable n1 n3.
Proof.
- induction 1; intros.
+ induction 1; intros.
- auto.
-- econstructor; eauto.
+- econstructor; eauto.
Qed.
Lemma reachable_right:
@@ -201,7 +201,7 @@ Variable transf: positive -> L.t -> L.t.
(i.e. put on the worklist at some point in the past).
Only the first two components are computationally relevant. The third
-is a ghost variable used only for stating and proving invariants.
+is a ghost variable used only for stating and proving invariants.
For this reason, [visited] is defined at sort [Prop] so that it is
erased during program extraction.
*)
@@ -266,7 +266,7 @@ Fixpoint propagate_succ_list (s: state) (out: L.t) (succs: list positive)
Definition step (s: state) : PMap.t L.t + state :=
match NS.pick s.(worklist) with
- | None =>
+ | None =>
inl _ (L.bot, s.(aval))
| Some(n, rem) =>
match code!n with
@@ -347,7 +347,7 @@ Remark optge_abstr_value:
optge st.(aval)!n st'.(aval)!n ->
L.ge (abstr_value n st) (abstr_value n st').
Proof.
- intros. unfold abstr_value. inv H. auto. apply L.ge_bot.
+ intros. unfold abstr_value. inv H. auto. apply L.ge_bot.
Qed.
Lemma propagate_succ_charact:
@@ -369,7 +369,7 @@ Proof.
- (* already there, unchanged *)
repeat split; intros.
+ rewrite E. constructor. eapply L.ge_trans. apply L.ge_refl. apply H; auto. apply L.ge_lub_right.
- + apply optge_refl.
+ + apply optge_refl.
+ right; auto.
+ auto.
+ auto.
@@ -378,29 +378,29 @@ Proof.
+ congruence.
- (* already there, updated *)
simpl; repeat split; intros.
- + rewrite PTree.gss. constructor. apply L.ge_lub_right.
+ + rewrite PTree.gss. constructor. apply L.ge_lub_right.
+ rewrite PTree.gso by auto. auto.
- + rewrite PTree.gsspec. destruct (peq s n).
+ + rewrite PTree.gsspec. destruct (peq s n).
subst s. rewrite E. constructor. apply L.ge_lub_left.
apply optge_refl.
- + rewrite NS.add_spec. auto.
+ + rewrite NS.add_spec. auto.
+ rewrite NS.add_spec. auto.
+ rewrite NS.add_spec in H0. intuition.
+ auto.
- + destruct H0; auto. subst n'. rewrite NS.add_spec; auto.
- + rewrite PTree.gsspec in H1. destruct (peq n' n). auto. congruence.
+ + destruct H0; auto. subst n'. rewrite NS.add_spec; auto.
+ + rewrite PTree.gsspec in H1. destruct (peq n' n). auto. congruence.
- (* not previously there, updated *)
simpl; repeat split; intros.
- + rewrite PTree.gss. apply optge_refl.
+ + rewrite PTree.gss. apply optge_refl.
+ rewrite PTree.gso by auto. auto.
- + rewrite PTree.gsspec. destruct (peq s n).
+ + rewrite PTree.gsspec. destruct (peq s n).
subst s. rewrite E. constructor.
apply optge_refl.
- + rewrite NS.add_spec. auto.
+ + rewrite NS.add_spec. auto.
+ rewrite NS.add_spec. auto.
+ rewrite NS.add_spec in H. intuition.
+ auto.
- + destruct H; auto. subst n'. rewrite NS.add_spec. auto.
+ + destruct H; auto. subst n'. rewrite NS.add_spec. auto.
+ rewrite PTree.gsspec in H0. destruct (peq n' n). auto. congruence.
Qed.
@@ -417,34 +417,34 @@ Lemma propagate_succ_list_charact:
/\ (forall n', st'.(visited) n' -> NS.In n' st'.(worklist) \/ st.(visited) n')
/\ (forall n', st.(aval)!n' = None -> st'.(aval)!n' <> None -> st'.(visited) n').
Proof.
- induction l; simpl; intros.
-- repeat split; intros.
+ induction l; simpl; intros.
+- repeat split; intros.
+ contradiction.
- + apply optge_refl.
+ + apply optge_refl.
+ auto.
+ auto.
+ auto.
+ auto.
+ auto.
+ congruence.
-- generalize (propagate_succ_charact st out a).
+- generalize (propagate_succ_charact st out a).
set (st1 := propagate_succ st out a).
intros (A1 & A2 & A3 & A4 & A5 & A6 & A7 & A8 & A9).
- generalize (IHl st1).
+ generalize (IHl st1).
set (st2 := propagate_succ_list st1 out l).
intros (B1 & B2 & B3 & B4 & B5 & B6 & B7 & B8 & B9). clear IHl.
repeat split; intros.
- + destruct H.
- * subst n. eapply optge_trans; eauto.
+ + destruct H.
+ * subst n. eapply optge_trans; eauto.
* auto.
+ rewrite B2 by tauto. apply A2; tauto.
+ eapply optge_trans; eauto.
+ destruct (B4 n). auto.
- destruct (peq n a).
+ destruct (peq n a).
* subst n. destruct A4. left; auto. right; congruence.
- * right. rewrite H. auto.
+ * right. rewrite H. auto.
+ eauto.
- + exploit B6; eauto. intros [P|P]. auto.
+ + exploit B6; eauto. intros [P|P]. auto.
exploit A6; eauto. intuition.
+ eauto.
+ specialize (B8 n'); specialize (A8 n'). intuition.
@@ -470,12 +470,12 @@ Proof.
eapply (PrimIter.iterate_prop _ _ step
(fun st => steps start st)
(fun res => exists st, steps start st /\ NS.pick (worklist st) = None /\ res = (L.bot, aval st))); eauto.
- intros. destruct (step a) eqn:E.
- exists a; split; auto.
+ intros. destruct (step a) eqn:E.
+ exists a; split; auto.
unfold step in E. destruct (NS.pick (worklist a)) as [[n rem]|].
destruct (code!n); discriminate.
- inv E. auto.
- eapply steps_right; eauto.
+ inv E. auto.
+ eapply steps_right; eauto.
constructor.
Qed.
@@ -492,10 +492,10 @@ Lemma step_incr:
forall n s1 s2, step s1 = inr s2 ->
optge s2.(aval)!n s1.(aval)!n /\ (s1.(visited) n -> s2.(visited) n).
Proof.
- unfold step; intros.
+ unfold step; intros.
destruct (NS.pick (worklist s1)) as [[p rem] | ]; try discriminate.
destruct (code!p) as [instr|]; inv H.
- + generalize (propagate_succ_list_charact
+ + generalize (propagate_succ_list_charact
(transf p (abstr_value p s1))
(successors instr)
{| aval := aval s1; worklist := rem; visited := visited s1 |}).
@@ -504,7 +504,7 @@ Proof.
(transf p (abstr_value p s1)) (successors instr)).
intros (A1 & A2 & A3 & A4 & A5 & A6 & A7 & A8 & A9).
auto.
- + split. apply optge_refl. auto.
+ + split. apply optge_refl. auto.
Qed.
Lemma steps_incr:
@@ -514,7 +514,7 @@ Proof.
induction 1.
- split. apply optge_refl. auto.
- destruct IHsteps. exploit (step_incr n); eauto. intros [P Q].
- split. eapply optge_trans; eauto. eauto.
+ split. eapply optge_trans; eauto. eauto.
Qed.
(** ** Correctness invariant *)
@@ -567,16 +567,16 @@ Proof.
+ (* n was on the worklist *)
rewrite PICK2 in P; destruct P.
* (* node n is our node pc *)
- subst n. fold out. right; intros.
- assert (i = instr) by congruence. subst i.
- apply A1; auto.
+ subst n. fold out. right; intros.
+ assert (i = instr) by congruence. subst i.
+ apply A1; auto.
* (* n was already on the worklist *)
left. apply A5; auto.
+ (* n was stable before, still is *)
- right; intros. apply optge_trans with st.(aval)!s; eauto.
+ right; intros. apply optge_trans with st.(aval)!s; eauto.
- (* defined *)
- destruct st.(aval)!n as [v'|] eqn:ST.
- + apply A7. eapply GOOD2; eauto.
+ destruct st.(aval)!n as [v'|] eqn:ST.
+ + apply A7. eapply GOOD2; eauto.
+ apply A9; auto. congruence.
Qed.
@@ -591,9 +591,9 @@ Proof.
generalize (NS.pick_some _ _ _ PICK); intro PICK2.
constructor; simpl; intros.
- (* stable *)
- exploit GOOD1; eauto. intros [P | P].
- + rewrite PICK2 in P. destruct P; auto.
- subst n. right; intros. congruence.
+ exploit GOOD1; eauto. intros [P | P].
+ + rewrite PICK2 in P. destruct P; auto.
+ subst n. right; intros. congruence.
+ right; exact P.
- (* defined *)
eapply GOOD2; eauto.
@@ -603,11 +603,11 @@ Lemma steps_state_good:
forall st1 st2, steps st1 st2 -> good_state st1 -> good_state st2.
Proof.
induction 1; intros.
-- auto.
+- auto.
- unfold step in e.
destruct (NS.pick (worklist s2)) as [[n rem] | ] eqn:PICK; try discriminate.
destruct (code!n) as [instr|] eqn:CODE; inv e.
- eapply step_state_good; eauto.
+ eapply step_state_good; eauto.
eapply step_state_good_2; eauto.
Qed.
@@ -616,8 +616,8 @@ Qed.
Lemma start_state_good:
forall enode eval, good_state (start_state enode eval).
Proof.
- intros. unfold start_state; constructor; simpl; intros.
-- subst n. rewrite NS.add_spec; auto.
+ intros. unfold start_state; constructor; simpl; intros.
+- subst n. rewrite NS.add_spec; auto.
- rewrite PTree.gsspec in H. rewrite PTree.gempty in H.
destruct (peq n enode). auto. discriminate.
Qed.
@@ -634,7 +634,7 @@ Lemma start_state_allnodes_good:
good_state start_state_allnodes.
Proof.
unfold start_state_allnodes; constructor; simpl; intros.
-- destruct H as [instr CODE]. left. eapply NS.all_nodes_spec; eauto.
+- destruct H as [instr CODE]. left. eapply NS.all_nodes_spec; eauto.
- rewrite PTree.gempty in H. congruence.
Qed.
@@ -645,9 +645,9 @@ Lemma reachable_visited:
forall p q, reachable code successors p q -> st.(visited) p -> st.(visited) q.
Proof.
intros st [GOOD1 GOOD2] PICK. induction 1; intros.
-- auto.
+- auto.
- eapply IHreachable; eauto.
- exploit GOOD1; eauto. intros [P | P].
+ exploit GOOD1; eauto. intros [P | P].
eelim NS.pick_none; eauto.
exploit P; eauto. intros OGE; inv OGE. eapply GOOD2; eauto.
Qed.
@@ -666,13 +666,13 @@ Theorem fixpoint_solution:
(forall n, L.eq (transf n L.bot) L.bot) ->
L.ge res!!s (transf n res!!n).
Proof.
- unfold fixpoint; intros.
+ unfold fixpoint; intros.
exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES).
exploit steps_state_good; eauto. apply start_state_good. intros [GOOD1 GOOD2].
rewrite RES; unfold PMap.get; simpl.
- destruct st.(aval)!n as [v|] eqn:STN.
+ destruct st.(aval)!n as [v|] eqn:STN.
- destruct (GOOD1 n) as [P|P]; eauto.
- eelim NS.pick_none; eauto.
+ eelim NS.pick_none; eauto.
exploit P; eauto. unfold abstr_value; rewrite STN. intros OGE; inv OGE. auto.
- apply L.ge_trans with L.bot. apply L.ge_bot. apply L.ge_refl. apply L.eq_sym. eauto.
Qed.
@@ -685,10 +685,10 @@ Theorem fixpoint_entry:
fixpoint ep ev = Some res ->
L.ge res!!ep ev.
Proof.
- unfold fixpoint; intros.
+ unfold fixpoint; intros.
exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES).
- exploit (steps_incr ep); eauto. simpl. rewrite PTree.gss. intros [P Q].
- rewrite RES; unfold PMap.get; simpl. inv P; auto.
+ exploit (steps_incr ep); eauto. simpl. rewrite PTree.gss. intros [P Q].
+ rewrite RES; unfold PMap.get; simpl. inv P; auto.
Qed.
(** For [fixpoint_allnodes], we show that the result is a solution
@@ -701,12 +701,12 @@ Theorem fixpoint_allnodes_solution:
In s (successors instr) ->
L.ge res!!s (transf n res!!n).
Proof.
- unfold fixpoint_allnodes; intros.
+ unfold fixpoint_allnodes; intros.
exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES).
exploit steps_state_good; eauto. apply start_state_allnodes_good. intros [GOOD1 GOOD2].
exploit (steps_incr n); eauto. simpl. intros [U V].
exploit (GOOD1 n). apply V. exists instr; auto. intros [P|P].
- eelim NS.pick_none; eauto.
+ eelim NS.pick_none; eauto.
exploit P; eauto. intros OGE. rewrite RES; unfold PMap.get; simpl.
inv OGE. assumption.
Qed.
@@ -723,15 +723,15 @@ Theorem fixpoint_nodeset_solution:
In s (successors instr) ->
L.ge res!!s (transf n res!!n).
Proof.
- unfold fixpoint_nodeset; intros.
+ unfold fixpoint_nodeset; intros.
exploit fixpoint_from_charact; eauto. intros (st & STEPS & PICK & RES).
exploit steps_state_good; eauto. apply start_state_nodeset_good. intros GOOD.
- exploit (steps_incr e); eauto. simpl. intros [U V].
- assert (st.(visited) n).
+ exploit (steps_incr e); eauto. simpl. intros [U V].
+ assert (st.(visited) n).
{ eapply reachable_visited; eauto. }
destruct GOOD as [GOOD1 GOOD2].
exploit (GOOD1 n); eauto. intros [P|P].
- eelim NS.pick_none; eauto.
+ eelim NS.pick_none; eauto.
exploit P; eauto. intros OGE. rewrite RES; unfold PMap.get; simpl.
inv OGE. assumption.
Qed.
@@ -754,8 +754,8 @@ Proof.
assert (inv (start_state ep ev)).
{
red; simpl; intros. unfold abstr_value, start_state; simpl.
- rewrite PTree.gsspec. rewrite PTree.gempty.
- destruct (peq x ep). auto. auto.
+ rewrite PTree.gsspec. rewrite PTree.gempty.
+ destruct (peq x ep). auto. auto.
}
assert (forall st v n, inv st -> P v -> inv (propagate_succ st v n)).
{
@@ -763,11 +763,11 @@ Proof.
destruct (aval st)!n as [oldl|] eqn:E.
destruct (L.beq oldl (L.lub oldl v)).
auto.
- unfold abstr_value. simpl. rewrite PTree.gsspec. destruct (peq x n).
- apply P_lub; auto. replace oldl with (abstr_value n st). auto.
- unfold abstr_value; rewrite E; auto.
- apply H1.
- unfold abstr_value. simpl. rewrite PTree.gsspec. destruct (peq x n).
+ unfold abstr_value. simpl. rewrite PTree.gsspec. destruct (peq x n).
+ apply P_lub; auto. replace oldl with (abstr_value n st). auto.
+ unfold abstr_value; rewrite E; auto.
+ apply H1.
+ unfold abstr_value. simpl. rewrite PTree.gsspec. destruct (peq x n).
auto.
apply H1.
}
@@ -782,13 +782,13 @@ Proof.
auto.
unfold step in e. destruct (NS.pick (worklist s2)) as [[n rem]|]; try discriminate.
destruct (code!n) as [instr|] eqn:INSTR; inv e.
- apply H2. apply IHsteps; auto. eapply P_transf; eauto. apply IHsteps; auto.
+ apply H2. apply IHsteps; auto. eapply P_transf; eauto. apply IHsteps; auto.
apply IHsteps; auto.
}
- unfold fixpoint in H. exploit fixpoint_from_charact; eauto.
- intros (st & STEPS & PICK & RES).
- replace (res!!pc) with (abstr_value pc st). eapply H3; eauto.
- rewrite RES; auto.
+ unfold fixpoint in H. exploit fixpoint_from_charact; eauto.
+ intros (st & STEPS & PICK & RES).
+ replace (res!!pc) with (abstr_value pc st). eapply H3; eauto.
+ rewrite RES; auto.
Qed.
End Kildall.
@@ -825,15 +825,15 @@ Fixpoint add_successors (pred: PTree.t (list positive))
Lemma add_successors_correct:
forall tolist from pred n s,
- In n pred!!!s \/ (n = from /\ In s tolist) ->
+ In n pred!!!s \/ (n = from /\ In s tolist) ->
In n (add_successors pred from tolist)!!!s.
Proof.
induction tolist; simpl; intros.
tauto.
apply IHtolist.
unfold successors_list at 1. rewrite PTree.gsspec. destruct (peq s a).
- subst a. destruct H. auto with coqlib.
- destruct H. subst n. auto with coqlib.
+ subst a. destruct H. auto with coqlib.
+ destruct H. subst n. auto with coqlib.
fold (successors_list pred s). intuition congruence.
Qed.
@@ -846,7 +846,7 @@ Lemma make_predecessors_correct_1:
code!n = Some instr -> In s (successors instr) ->
In n make_predecessors!!!s.
Proof.
- intros until s.
+ intros until s.
set (P := fun m p => m!n = Some instr -> In s (successors instr) ->
In n p!!!s).
unfold make_predecessors.
@@ -857,7 +857,7 @@ Proof.
rewrite PTree.gempty in H; congruence.
(* inductive case *)
apply add_successors_correct.
- rewrite PTree.gsspec in H2. destruct (peq n k).
+ rewrite PTree.gsspec in H2. destruct (peq n k).
inv H2. auto.
auto.
Qed.
@@ -867,7 +867,7 @@ Lemma make_predecessors_correct_2:
code!n = Some instr -> In s (successors instr) ->
exists l, make_predecessors!s = Some l /\ In n l.
Proof.
- intros. exploit make_predecessors_correct_1; eauto.
+ intros. exploit make_predecessors_correct_1; eauto.
unfold successors_list. destruct (make_predecessors!s); simpl; intros.
exists l; auto.
contradiction.
@@ -878,10 +878,10 @@ Lemma reachable_predecessors:
reachable code successors p q ->
reachable make_predecessors (fun l => l) q p.
Proof.
- induction 1.
+ induction 1.
- constructor.
-- exploit make_predecessors_correct_2; eauto. intros [l [P Q]].
- eapply reachable_right; eauto.
+- exploit make_predecessors_correct_2; eauto. intros [l [P Q]].
+ eapply reachable_right; eauto.
Qed.
End Predecessor.
@@ -953,9 +953,9 @@ Variable transf: positive -> L.t -> L.t.
Section Exit_points.
-(** Assuming that the nodes of the CFG [code] are numbered in reverse
+(** Assuming that the nodes of the CFG [code] are numbered in reverse
postorder (cf. pass [Renumber]), an edge from [n] to [s] is a
- normal edge if [s < n] and a back-edge otherwise.
+ normal edge if [s < n] and a back-edge otherwise.
[sequential_node] returns [true] if the given node has at least one
normal outgoing edge. It returns [false] if the given node is an exit
node (no outgoing edges) or the final node of a loop body
@@ -969,32 +969,32 @@ Definition sequential_node (pc: positive) (instr: A): bool :=
Definition exit_points : NS.t :=
PTree.fold
- (fun ep pc instr =>
+ (fun ep pc instr =>
if sequential_node pc instr
then ep
else NS.add pc ep)
code NS.empty.
Lemma exit_points_charact:
- forall n,
+ forall n,
NS.In n exit_points <-> exists i, code!n = Some i /\ sequential_node n i = false.
Proof.
- intros n. unfold exit_points. eapply PTree_Properties.fold_rec.
+ intros n. unfold exit_points. eapply PTree_Properties.fold_rec.
- (* extensionality *)
- intros. rewrite <- H. auto.
+ intros. rewrite <- H. auto.
- (* base case *)
- simpl. split; intros.
- eelim NS.empty_spec; eauto.
- destruct H as [i [P Q]]. rewrite PTree.gempty in P. congruence.
+ simpl. split; intros.
+ eelim NS.empty_spec; eauto.
+ destruct H as [i [P Q]]. rewrite PTree.gempty in P. congruence.
- (* inductive case *)
- intros. destruct (sequential_node k v) eqn:SN.
- + rewrite H1. rewrite PTree.gsspec. destruct (peq n k).
- subst. split; intros [i [P Q]]. congruence. inv P. congruence.
+ intros. destruct (sequential_node k v) eqn:SN.
+ + rewrite H1. rewrite PTree.gsspec. destruct (peq n k).
+ subst. split; intros [i [P Q]]. congruence. inv P. congruence.
tauto.
- + rewrite NS.add_spec. rewrite H1. rewrite PTree.gsspec. destruct (peq n k).
- subst. split. intros. exists v; auto. auto.
- split. intros [P | [i [P Q]]]. congruence. exists i; auto.
- intros [i [P Q]]. right; exists i; auto.
+ + rewrite NS.add_spec. rewrite H1. rewrite PTree.gsspec. destruct (peq n k).
+ subst. split. intros. exists v; auto. auto.
+ split. intros [P | [i [P Q]]]. congruence. exists i; auto.
+ intros [i [P Q]]. right; exists i; auto.
Qed.
Lemma reachable_exit_points:
@@ -1002,15 +1002,15 @@ Lemma reachable_exit_points:
code!pc = Some i -> exists x, NS.In x exit_points /\ reachable code successors pc x.
Proof.
intros pc0. pattern pc0. apply (well_founded_ind Plt_wf).
- intros pc HR i CODE.
- destruct (sequential_node pc i) eqn:SN.
+ intros pc HR i CODE.
+ destruct (sequential_node pc i) eqn:SN.
- (* at least one successor that decreases the pc *)
- unfold sequential_node in SN. rewrite existsb_exists in SN.
+ unfold sequential_node in SN. rewrite existsb_exists in SN.
destruct SN as [s [P Q]]. destruct (code!s) as [i'|] eqn:CS; try discriminate. InvBooleans.
- exploit (HR s); eauto. intros [x [U V]].
- exists x; split; auto. eapply reachable_left; eauto.
+ exploit (HR s); eauto. intros [x [U V]].
+ exists x; split; auto. eapply reachable_left; eauto.
- (* otherwise we are an exit point *)
- exists pc; split.
+ exists pc; split.
rewrite exit_points_charact. exists i; auto. constructor.
Qed.
@@ -1023,7 +1023,7 @@ Lemma reachable_exit_points_predecessor:
exists x, NS.In x exit_points /\ reachable (make_predecessors code successors) (fun l => l) x pc.
Proof.
intros. exploit reachable_exit_points; eauto. intros [x [P Q]].
- exists x; split; auto. apply reachable_predecessors. auto.
+ exists x; split; auto. apply reachable_predecessors. auto.
Qed.
End Exit_points.
@@ -1067,7 +1067,7 @@ Theorem fixpoint_allnodes_solution:
Proof.
intros.
exploit (make_predecessors_correct_2 code); eauto. intros [l [P Q]].
- unfold fixpoint_allnodes in H.
+ unfold fixpoint_allnodes in H.
eapply DS.fixpoint_allnodes_solution; eauto.
Qed.
@@ -1082,7 +1082,7 @@ End Backward_Dataflow_Solver.
In other terms, program points with multiple predecessors are mapped
to [L.top] (the greatest, or coarsest, approximation) and the other
program points are mapped to [transf p X[p]] where [p] is their unique
- predecessor.
+ predecessor.
This analysis applies to any type of approximations equipped with
an ordering and a greatest element. *)
@@ -1205,7 +1205,7 @@ Definition step (bb: bbmap) (st: state) : result + state :=
| None =>
inr _ (mkstate st.(aval) rem)
| Some instr =>
- inr _ (propagate_successors
+ inr _ (propagate_successors
bb (successors instr)
(transf pc st.(aval)!!pc)
(mkstate st.(aval) rem))
@@ -1214,7 +1214,7 @@ Definition step (bb: bbmap) (st: state) : result + state :=
(** Recognition of program points that have more than one predecessor. *)
-Definition is_basic_block_head
+Definition is_basic_block_head
(preds: PTree.t (list positive)) (pc: positive) : bool :=
if peq pc entrypoint then true else
match preds!!!pc with
@@ -1254,16 +1254,16 @@ Lemma multiple_predecessors:
n1 <> n2 ->
basic_block_map s = true.
Proof.
- intros.
+ intros.
assert (In n1 predecessors!!!s). eapply predecessors_correct; eauto.
assert (In n2 predecessors!!!s). eapply predecessors_correct; eauto.
unfold basic_block_map, is_basic_block_head.
- destruct (peq s entrypoint). auto.
+ destruct (peq s entrypoint). auto.
fold predecessors.
- destruct (predecessors!!!s).
+ destruct (predecessors!!!s).
auto.
destruct l.
- apply proj_sumbool_is_true. simpl in *. intuition congruence.
+ apply proj_sumbool_is_true. simpl in *. intuition congruence.
auto.
Qed.
@@ -1272,12 +1272,12 @@ Lemma no_self_loop:
code!n = Some instr -> In n (successors instr) -> basic_block_map n = true.
Proof.
intros. unfold basic_block_map, is_basic_block_head.
- destruct (peq n entrypoint). auto.
+ destruct (peq n entrypoint). auto.
fold predecessors.
- exploit predecessors_correct; eauto. intros.
+ exploit predecessors_correct; eauto. intros.
destruct (predecessors!!!n).
- contradiction.
- destruct l. apply proj_sumbool_is_true. simpl in H1. tauto.
+ contradiction.
+ destruct l. apply proj_sumbool_is_true. simpl in H1. tauto.
auto.
Qed.
@@ -1285,7 +1285,7 @@ Qed.
(** The invariant over the state is as follows:
- Points with several predecessors are mapped to [L.top]
-- Points not in the worklist satisfy their inequations
+- Points not in the worklist satisfy their inequations
(as in Kildall's algorithm).
*)
@@ -1294,7 +1294,7 @@ Definition state_invariant (st: state) : Prop :=
/\
(forall n,
In n st.(worklist) \/
- (forall instr s, code!n = Some instr -> In s (successors instr) ->
+ (forall instr s, code!n = Some instr -> In s (successors instr) ->
L.ge st.(aval)!!s (transf n st.(aval)!!n))).
Lemma propagate_successors_charact1:
@@ -1326,8 +1326,8 @@ Proof.
caseEq (bb a); intro.
elim (IHsuccs l st n); intros U V.
split; intros. apply U; auto.
- elim H0; intro. subst a. congruence. auto.
- apply V. tauto.
+ elim H0; intro. subst a. congruence. auto.
+ apply V. tauto.
set (st1 := mkstate (PMap.set a l (aval st)) (a :: worklist st)).
elim (IHsuccs l st1 n); intros U V.
split; intros.
@@ -1338,16 +1338,16 @@ Proof.
elim (U i H1); auto.
rewrite V. unfold st1; simpl. apply PMap.gss. tauto.
apply U; auto.
- rewrite V. unfold st1; simpl. apply PMap.gso.
+ rewrite V. unfold st1; simpl. apply PMap.gso.
red; intro; subst n. elim H0; intro. tauto. congruence.
- tauto.
+ tauto.
Qed.
Lemma propagate_successors_invariant:
forall pc instr res rem,
code!pc = Some instr ->
state_invariant (mkstate res (pc :: rem)) ->
- state_invariant
+ state_invariant
(propagate_successors basic_block_map (successors instr)
(transf pc res!!pc)
(mkstate res rem)).
@@ -1360,23 +1360,23 @@ Proof.
(successors instr) l (mkstate res rem)).
set (st1 := propagate_successors basic_block_map
(successors instr) l (mkstate res rem)).
- intros U V. simpl in U.
+ intros U V. simpl in U.
(* First part: BB entries remain at top *)
split; intros.
- elim (U n); intros C D. rewrite D. simpl. apply INV1. auto. tauto.
+ elim (U n); intros C D. rewrite D. simpl. apply INV1. auto. tauto.
(* Second part: monotonicity *)
(* Case 1: n = pc *)
- destruct (peq pc n). subst n.
+ destruct (peq pc n). subst n.
right; intros.
assert (instr0 = instr) by congruence. subst instr0.
elim (U s); intros C D.
replace (st1.(aval)!!pc) with res!!pc. fold l.
destruct (basic_block_map s) eqn:BB.
- rewrite D. simpl. rewrite INV1. apply L.top_ge. auto. tauto.
- elim (C H0 (refl_equal _)). intros X Y. rewrite Y. apply L.refl_ge.
- elim (U pc); intros E F. rewrite F. reflexivity.
+ rewrite D. simpl. rewrite INV1. apply L.top_ge. auto. tauto.
+ elim (C H0 (refl_equal _)). intros X Y. rewrite Y. apply L.refl_ge.
+ elim (U pc); intros E F. rewrite F. reflexivity.
destruct (In_dec peq pc (successors instr)).
- right. eapply no_self_loop; eauto.
+ right. eapply no_self_loop; eauto.
left; auto.
(* Case 2: n <> pc *)
elim (INV2 n); intro.
@@ -1388,7 +1388,7 @@ Proof.
they could change is if they were successors of pc as well,
but that gives them two different predecessors, so
they are basic block heads, and thus do not change! *)
- intros. elim (U s); intros C D. rewrite D. reflexivity.
+ intros. elim (U s); intros C D. rewrite D. reflexivity.
destruct (In_dec peq s (successors instr)).
right. eapply multiple_predecessors with (n1 := pc) (n2 := n); eauto.
left; auto.
@@ -1396,13 +1396,13 @@ Proof.
(* Case 2.2.1: n is a successor of pc. Either it is in the
worklist or it did not change *)
destruct (basic_block_map n) eqn:BB.
- right; intros.
+ right; intros.
elim (U n); intros C D. rewrite D. erewrite INV3; eauto.
tauto.
left. elim (U n); intros C D. elim (C i BB); intros. auto.
(* Case 2.2.2: n is not a successor of pc. It did not change. *)
right; intros.
- elim (U n); intros C D. rewrite D.
+ elim (U n); intros C D. rewrite D.
erewrite INV3; eauto.
tauto.
Qed.
@@ -1416,7 +1416,7 @@ Proof.
intros until rem. intros CODE [INV1 INV2]. simpl in INV1. simpl in INV2.
split; simpl; intros.
apply INV1; auto.
- destruct (INV2 n) as [[U | U] | U].
+ destruct (INV2 n) as [[U | U] | U].
subst n. right; intros; congruence.
auto.
auto.
@@ -1439,12 +1439,12 @@ Proof.
eapply (PrimIter.iterate_prop _ _ (step basic_block_map)
state_invariant).
- intros st INV. destruct st as [stin stwrk].
- unfold step. simpl. destruct stwrk as [ | pc rem ] eqn:WRK.
+ intros st INV. destruct st as [stin stwrk].
+ unfold step. simpl. destruct stwrk as [ | pc rem ] eqn:WRK.
auto.
destruct (code!pc) as [instr|] eqn:CODE.
- eapply propagate_successors_invariant; eauto.
- eapply propagate_successors_invariant_2; eauto.
+ eapply propagate_successors_invariant; eauto.
+ eapply propagate_successors_invariant_2; eauto.
eauto. apply initial_state_invariant.
Qed.
@@ -1457,11 +1457,11 @@ Theorem fixpoint_solution:
code!n = Some instr -> In s (successors instr) ->
L.ge res!!s (transf n res!!n).
Proof.
- intros.
+ intros.
assert (state_invariant (mkstate res nil)).
eapply analyze_invariant; eauto.
- elim H2; simpl; intros.
- elim (H4 n); intros.
+ elim H2; simpl; intros.
+ elim (H4 n); intros.
contradiction.
eauto.
Qed.
@@ -1471,13 +1471,13 @@ Theorem fixpoint_entry:
fixpoint = Some res ->
res!!entrypoint = L.top.
Proof.
- intros.
+ intros.
assert (state_invariant (mkstate res nil)).
- eapply analyze_invariant; eauto.
- elim H0; simpl; intros.
+ eapply analyze_invariant; eauto.
+ elim H0; simpl; intros.
apply H1. unfold basic_block_map, is_basic_block_head.
- fold predecessors. apply peq_true.
-Qed.
+ fold predecessors. apply peq_true.
+Qed.
(** ** Preservation of a property over solutions *)
@@ -1493,8 +1493,8 @@ Lemma propagate_successors_P:
Proof.
induction succs; simpl; intros.
auto.
- case (bb a). auto.
- apply IHsuccs. red; simpl; intros.
+ case (bb a). auto.
+ apply IHsuccs. red; simpl; intros.
rewrite PMap.gsspec. case (peq pc a); intro.
auto. apply H0.
Qed.
@@ -1502,7 +1502,7 @@ Qed.
Theorem fixpoint_invariant:
forall res pc, fixpoint = Some res -> P res!!pc.
Proof.
- unfold fixpoint; intros. pattern res.
+ unfold fixpoint; intros. pattern res.
eapply (PrimIter.iterate_prop _ _ (step basic_block_map) Pstate).
intros st PS. unfold step. destruct (st.(worklist)).
@@ -1510,10 +1510,10 @@ Proof.
assert (PS2: Pstate (mkstate st.(aval) l)).
red; intro; simpl. apply PS.
destruct (code!p) as [instr|] eqn:CODE.
- apply propagate_successors_P. eauto. auto.
+ apply propagate_successors_P. eauto. auto.
auto.
- eauto.
+ eauto.
red; intro; simpl. rewrite PMap.gi. apply Ptop.
Qed.
@@ -1532,7 +1532,7 @@ End BBlock_solver.
the enumeration [n-1], [n-2], ..., 3, 2, 1 where [n] is the
top CFG node is a reverse postorder traversal.
Therefore, for forward analysis, we will use an implementation
- of [NODE_SET] where the [pick] operation selects the
+ of [NODE_SET] where the [pick] operation selects the
greatest node in the working list. For backward analysis,
we will similarly pick the smallest node in the working list. *)
@@ -1562,7 +1562,7 @@ Module NodeSetForward <: NODE_SET.
Proof.
intros. rewrite PHeap.In_insert. unfold In. intuition.
Qed.
-
+
Lemma pick_none:
forall s n, pick s = None -> ~In n s.
Proof.
@@ -1582,14 +1582,14 @@ Module NodeSetForward <: NODE_SET.
Qed.
Lemma all_nodes_spec:
- forall A (code: PTree.t A) n instr,
+ forall A (code: PTree.t A) n instr,
code!n = Some instr -> In n (all_nodes code).
Proof.
intros A code n instr.
apply PTree_Properties.fold_rec with
(P := fun m set => m!n = Some instr -> In n set).
(* extensionality *)
- intros. apply H0. rewrite H. auto.
+ intros. apply H0. rewrite H. auto.
(* base case *)
rewrite PTree.gempty. congruence.
(* inductive case *)
@@ -1638,7 +1638,7 @@ Module NodeSetBackward <: NODE_SET.
Qed.
Lemma all_nodes_spec:
- forall A (code: PTree.t A) n instr,
+ forall A (code: PTree.t A) n instr,
code!n = Some instr -> In n (all_nodes code).
Proof NodeSetForward.all_nodes_spec.
End NodeSetBackward.
diff --git a/backend/LTL.v b/backend/LTL.v
index 67fb0197..48c5c850 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -77,8 +77,8 @@ Definition genv := Genv.t fundef unit.
Definition locset := Locmap.t.
(** Calling conventions are reflected at the level of location sets
- (environments mapping locations to values) by the following two
- functions.
+ (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.
@@ -87,7 +87,7 @@ Definition locset := Locmap.t.
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) =>
diff --git a/backend/Linearize.v b/backend/Linearize.v
index 78cdd743..68c2b32f 100644
--- a/backend/Linearize.v
+++ b/backend/Linearize.v
@@ -56,7 +56,7 @@ Open Scope error_monad_scope.
The main challenge in code linearization is therefore to pick a
``good'' order for the nodes that exploits well the
fall-through behavior. Many clever trace picking heuristics
- have been developed for this purpose.
+ have been developed for this purpose.
In this file, we present linearization in a way that clearly
separates the heuristic part (choosing an order for the basic blocks)
@@ -96,7 +96,7 @@ Definition reachable_aux (f: LTL.function) : option (PMap.t bool) :=
f.(fn_entrypoint) true.
Definition reachable (f: LTL.function) : PMap.t bool :=
- match reachable_aux f with
+ match reachable_aux f with
| None => PMap.init true
| Some rs => rs
end.
@@ -118,7 +118,7 @@ Fixpoint nodeset_of_list (l: list node) (s: Nodeset.t)
match l with
| nil => OK s
| hd :: tl =>
- if Nodeset.mem hd s
+ if Nodeset.mem hd s
then Error (msg "Linearize: duplicates in enumeration")
else nodeset_of_list tl (Nodeset.add hd s)
end.
diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml
index ef268562..71ee2e56 100644
--- a/backend/Linearizeaux.ml
+++ b/backend/Linearizeaux.ml
@@ -26,7 +26,7 @@ let enumerate_aux f reach =
(fun pc nodes ->
if PMap.get pc reach
then Coq_cons (pc, nodes)
- else nodes)
+ else nodes)
f.fn_nextpc
***)
@@ -100,7 +100,7 @@ let basic_blocks f joins =
(* end_block: record block that we just discovered *)
and end_block blk minpc =
blocks := (minpc, List.rev blk) :: !blocks
- in
+ in
start_block f.fn_entrypoint; !blocks
(* Flatten basic blocks in decreasing order of minpc *)
diff --git a/backend/Linearizeproof.v b/backend/Linearizeproof.v
index dc4d11ea..65258b2d 100644
--- a/backend/Linearizeproof.v
+++ b/backend/Linearizeproof.v
@@ -113,7 +113,7 @@ Proof.
caseEq (reachable_aux f).
unfold reachable_aux; intros reach A.
assert (LBoolean.ge reach!!(f.(fn_entrypoint)) true).
- eapply DS.fixpoint_entry. eexact A. auto.
+ eapply DS.fixpoint_entry. eexact A. auto.
unfold LBoolean.ge in H. tauto.
intros. apply PMap.gi.
Qed.
@@ -131,7 +131,7 @@ Proof.
unfold reachable_aux. intro reach; intros.
assert (LBoolean.ge reach!!pc' reach!!pc).
change (reach!!pc) with ((fun pc r => r) pc (reach!!pc)).
- eapply DS.fixpoint_solution; eauto. intros; apply DS.L.eq_refl.
+ eapply DS.fixpoint_solution; eauto. intros; apply DS.L.eq_refl.
elim H3; intro. congruence. auto.
intros. apply PMap.gi.
Qed.
@@ -152,13 +152,13 @@ Lemma nodeset_of_list_correct:
/\ (forall pc, Nodeset.In pc s' <-> Nodeset.In pc s \/ In pc l)
/\ (forall pc, In pc l -> ~Nodeset.In pc s).
Proof.
- induction l; simpl; intros.
+ induction l; simpl; intros.
inv H. split. constructor. split. intro; tauto. intros; tauto.
generalize H; clear H; caseEq (Nodeset.mem a s); intros.
inv H0.
exploit IHl; eauto. intros [A [B C]].
split. constructor; auto. red; intro. elim (C a H1). apply Nodeset.add_1. hnf. auto.
- split. intros. rewrite B. rewrite NodesetFacts.add_iff.
+ split. intros. rewrite B. rewrite NodesetFacts.add_iff.
unfold Nodeset.E.eq. unfold OrderedPositive.eq. tauto.
intros. destruct H1. subst pc. rewrite NodesetFacts.not_mem_iff. auto.
generalize (C pc H1). rewrite NodesetFacts.add_iff. tauto.
@@ -172,7 +172,7 @@ Lemma check_reachable_correct:
Nodeset.In pc s.
Proof.
intros f reach s.
- assert (forall l ok,
+ assert (forall l ok,
List.fold_left (fun a p => check_reachable_aux reach s a (fst p) (snd p)) l ok = true ->
ok = true /\
(forall pc i,
@@ -181,16 +181,16 @@ Proof.
Nodeset.In pc s)).
induction l; simpl; intros.
split. auto. intros. destruct H0.
- destruct a as [pc1 i1]. simpl in H.
+ destruct a as [pc1 i1]. simpl in H.
exploit IHl; eauto. intros [A B].
- unfold check_reachable_aux in A.
+ unfold check_reachable_aux in A.
split. destruct (reach!!pc1). elim (andb_prop _ _ A). auto. auto.
- intros. destruct H0. inv H0. rewrite H1 in A. destruct (andb_prop _ _ A).
+ intros. destruct H0. inv H0. rewrite H1 in A. destruct (andb_prop _ _ A).
apply Nodeset.mem_2; auto.
eauto.
intros pc i. unfold check_reachable. rewrite PTree.fold_spec. intros.
- exploit H; eauto. intros [A B]. eapply B; eauto.
+ exploit H; eauto. intros [A B]. eapply B; eauto.
apply PTree.elements_correct. eauto.
Qed.
@@ -201,9 +201,9 @@ Lemma enumerate_complete:
(reachable f)!!pc = true ->
In pc enum.
Proof.
- intros until i. unfold enumerate.
+ intros until i. unfold enumerate.
set (reach := reachable f).
- intros. monadInv H.
+ intros. monadInv H.
generalize EQ0; clear EQ0. caseEq (check_reachable f reach x); intros; inv EQ0.
exploit check_reachable_correct; eauto. intro.
exploit nodeset_of_list_correct; eauto. intros [A [B C]].
@@ -215,9 +215,9 @@ Lemma enumerate_norepet:
enumerate f = OK enum ->
list_norepet enum.
Proof.
- intros until enum. unfold enumerate.
+ intros until enum. unfold enumerate.
set (reach := reachable f).
- intros. monadInv H.
+ intros. monadInv H.
generalize EQ0; clear EQ0. caseEq (check_reachable f reach x); intros; inv EQ0.
exploit nodeset_of_list_correct; eauto. intros [A [B C]]. auto.
Qed.
@@ -246,9 +246,9 @@ Proof.
simpl; intros; discriminate.
intros c3 TAIL UNIQ. simpl.
generalize (is_label_correct lbl a). case (is_label lbl a); intro ISLBL.
- subst a. intro. inversion TAIL. congruence.
+ subst a. intro. inversion TAIL. congruence.
elim UNIQ; intros. elim H4. apply is_tail_in with c1; auto.
- inversion TAIL. congruence. apply IHc2. auto.
+ inversion TAIL. congruence. apply IHc2. auto.
destruct a; simpl in UNIQ; tauto.
Qed.
@@ -266,13 +266,13 @@ Proof.
induction c1.
simpl; intros; discriminate.
simpl starts_with. destruct a; try (intros; discriminate).
- intros.
+ intros.
apply plus_left with E0 (State s f sp c1 ls m) E0.
- simpl. constructor.
+ simpl. constructor.
destruct (peq lbl l).
subst l. replace c3 with c1. constructor.
apply find_label_unique with lbl c2; auto.
- apply plus_star.
+ apply plus_star.
apply IHc1 with c2; auto. eapply is_tail_cons_left; eauto.
traceEq.
Qed.
@@ -291,7 +291,7 @@ Lemma find_label_lin_block:
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. destruct a; simpl; auto.
+ induction b; simpl; auto. destruct a; simpl; auto.
case (starts_with s1 k); simpl; auto.
Qed.
@@ -303,7 +303,7 @@ Remark linearize_body_cons:
| Some b => Llabel pc :: linearize_block b (linearize_body f enum)
end.
Proof.
- intros. unfold linearize_body. rewrite list_fold_right_eq.
+ intros. unfold linearize_body. rewrite list_fold_right_eq.
unfold linearize_node. destruct (LTL.fn_code f)!pc; auto.
Qed.
@@ -315,13 +315,13 @@ Lemma find_label_lin_rec:
Proof.
induction enum; intros.
elim H.
- rewrite linearize_body_cons.
+ rewrite linearize_body_cons.
destruct (peq a pc).
subst a. exists (linearize_body f enum).
rewrite H0. simpl. rewrite peq_true. auto.
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.
+ exists k. destruct (LTL.fn_code f)!a.
simpl. rewrite peq_false. rewrite find_label_lin_block. auto. auto.
auto.
Qed.
@@ -334,7 +334,7 @@ Lemma find_label_lin:
exists k,
find_label pc (fn_code tf) = Some (linearize_block b k).
Proof.
- intros. monadInv H. simpl.
+ intros. monadInv H. simpl.
rewrite find_label_add_branch. apply find_label_lin_rec.
eapply enumerate_complete; eauto. auto.
Qed.
@@ -379,8 +379,8 @@ Lemma label_in_lin_rec:
Proof.
induction enum.
simpl; auto.
- rewrite linearize_body_cons. destruct (LTL.fn_code f)!a.
- simpl. intros [A|B]. left; congruence.
+ rewrite linearize_body_cons. destruct (LTL.fn_code f)!a.
+ simpl. intros [A|B]. left; congruence.
right. apply IHenum. eapply label_in_lin_block; eauto.
intro; right; auto.
Qed.
@@ -389,7 +389,7 @@ Lemma unique_labels_add_branch:
forall lbl k,
unique_labels k -> unique_labels (add_branch lbl k).
Proof.
- intros; unfold add_branch.
+ intros; unfold add_branch.
destruct (starts_with lbl k); simpl; intuition.
Qed.
@@ -410,9 +410,9 @@ Proof.
induction enum.
simpl; auto.
rewrite linearize_body_cons.
- intro. destruct (LTL.fn_code f)!a.
+ 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_rec with f.
apply label_in_lin_block with b. auto.
apply unique_labels_lin_block. apply IHenum. inversion H; auto.
apply IHenum. inversion H; auto.
@@ -424,7 +424,7 @@ Lemma unique_labels_transf_function:
unique_labels (fn_code tf).
Proof.
intros. monadInv H. simpl.
- apply unique_labels_add_branch.
+ apply unique_labels_add_branch.
apply unique_labels_lin_rec. eapply enumerate_norepet; eauto.
Qed.
@@ -438,7 +438,7 @@ Proof.
intros; discriminate.
case (is_label lbl a). intro. injection H; intro. subst c2.
constructor. constructor.
- intro. constructor. auto.
+ intro. constructor. auto.
Qed.
Lemma is_tail_add_branch:
@@ -454,7 +454,7 @@ Lemma is_tail_lin_block:
Proof.
induction b; simpl; intros.
auto.
- destruct a; eauto with coqlib.
+ 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.
@@ -558,7 +558,7 @@ Definition measure (S: LTL.state) : nat :=
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.
+ induction 1; simpl. auto. inv H; auto.
Qed.
Theorem transf_step_correct:
@@ -570,41 +570,41 @@ Proof.
induction 1; intros; try (inv MS).
(* start of block, at an [add_branch] *)
- exploit find_label_lin; eauto. intros [k F].
+ exploit find_label_lin; eauto. intros [k F].
left; econstructor; split.
- eapply add_branch_correct; eauto.
- econstructor; eauto.
+ 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].
+ exploit find_label_lin; eauto. intros [k F].
left; econstructor; split.
- apply plus_one. eapply exec_Lcond_true; eauto.
- econstructor; eauto.
+ 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].
+ exploit find_label_lin; eauto. intros [k F].
left; econstructor; split.
- apply plus_one. eapply exec_Ljumptable; eauto.
- econstructor; eauto.
+ 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 *)
left; econstructor; split. simpl.
- apply plus_one. econstructor; eauto.
- instantiate (1 := v); rewrite <- H; apply eval_operation_preserved.
+ apply plus_one. econstructor; eauto.
+ instantiate (1 := v); rewrite <- H; apply eval_operation_preserved.
exact symbols_preserved.
- econstructor; eauto.
+ econstructor; eauto.
(* Lload *)
left; econstructor; split. simpl.
- apply plus_one. econstructor.
- instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
- exact symbols_preserved. eauto. eauto.
+ apply plus_one. econstructor.
+ instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
+ exact symbols_preserved. eauto. eauto.
econstructor; eauto.
(* Lgetstack *)
@@ -614,14 +614,14 @@ Proof.
(* Lsetstack *)
left; econstructor; split. simpl.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
econstructor; eauto.
(* Lstore *)
left; econstructor; split. simpl.
- apply plus_one. econstructor.
- instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
- exact symbols_preserved. eauto. eauto.
+ apply plus_one. econstructor.
+ instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
+ exact symbols_preserved. eauto. eauto.
econstructor; eauto.
(* Lcall *)
@@ -629,7 +629,7 @@ Proof.
left; econstructor; split. simpl.
apply plus_one. econstructor; eauto.
symmetry; eapply sig_preserved; eauto.
- econstructor; eauto. constructor; auto. econstructor; eauto.
+ econstructor; eauto. constructor; auto. econstructor; eauto.
(* Ltailcall *)
exploit find_function_translated; eauto. intros [tfd [A B]].
@@ -637,7 +637,7 @@ Proof.
apply plus_one. econstructor; eauto.
rewrite (match_parent_locset _ _ STACKS). eauto.
symmetry; eapply sig_preserved; eauto.
- rewrite (stacksize_preserved _ _ TRF); eauto.
+ rewrite (stacksize_preserved _ _ TRF); eauto.
rewrite (match_parent_locset _ _ STACKS).
econstructor; eauto.
@@ -664,25 +664,25 @@ Proof.
destruct b.
(* cond is true: no branch *)
left; econstructor; split.
- apply plus_one. eapply exec_Lcond_false.
+ 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.
+ 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.
+ 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.
+ apply plus_one. eapply exec_Lcond_false. eauto. eauto.
econstructor; eauto.
(* Ljumptable *)
assert (REACH': (reachable f)!!pc = true).
- apply REACH. simpl. eapply list_nth_z_in; eauto.
- right; split. simpl; omega. split. auto. econstructor; eauto.
+ apply REACH. simpl. eapply list_nth_z_in; eauto.
+ right; split. simpl; omega. split. auto. econstructor; eauto.
(* Lreturn *)
left; econstructor; split.
@@ -695,9 +695,9 @@ Proof.
apply reachable_entrypoint.
monadInv H7.
left; econstructor; split.
- apply plus_one. eapply exec_function_internal; eauto.
+ apply plus_one. eapply exec_function_internal; eauto.
rewrite (stacksize_preserved _ _ EQ). eauto.
- generalize EQ; intro EQ'; monadInv EQ'. simpl.
+ generalize EQ; intro EQ'; monadInv EQ'. simpl.
econstructor; eauto. simpl. eapply is_tail_add_branch. constructor.
(* external function *)
@@ -710,8 +710,8 @@ Proof.
(* return *)
inv H3. inv H1.
left; econstructor; split.
- apply plus_one. econstructor.
- econstructor; eauto.
+ apply plus_one. econstructor.
+ econstructor; eauto.
Qed.
Lemma transf_initial_states:
@@ -719,18 +719,18 @@ Lemma transf_initial_states:
exists st2, Linear.initial_state tprog st2 /\ match_states st1 st2.
Proof.
intros. inversion H.
- exploit function_ptr_translated; eauto. intros [tf [A B]].
+ exploit function_ptr_translated; eauto. intros [tf [A B]].
exists (Callstate nil tf (Locmap.init Vundef) m0); split.
- econstructor; eauto. eapply Genv.init_mem_transf_partial; eauto.
+ 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).
+ symmetry. apply (transform_partial_program_main transf_fundef _ TRANSF).
rewrite <- H3. apply sig_preserved. auto.
constructor. constructor. auto.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> LTL.final_state st1 r -> Linear.final_state st2 r.
Proof.
intros. inv H0. inv H. inv H6. econstructor; eauto.
diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v
index 62a0c585..a52e47bb 100644
--- a/backend/Lineartyping.v
+++ b/backend/Lineartyping.v
@@ -72,7 +72,7 @@ Definition wt_instr (i: instruction) : bool :=
match is_move_operation op args with
| Some arg =>
subtype (mreg_type arg) (mreg_type res)
- | None =>
+ | None =>
let (targs, tres) := type_of_operation op in
subtype tres (mreg_type res)
end
@@ -105,7 +105,7 @@ Lemma wt_setreg:
Val.has_type v (mreg_type r) -> wt_locset ls -> wt_locset (Locmap.set (R r) v ls).
Proof.
intros; red; intros.
- unfold Locmap.set.
+ unfold Locmap.set.
destruct (Loc.eq (R r) l).
subst l; auto.
destruct (Loc.diff_dec (R r) l). auto. red. auto.
@@ -116,10 +116,10 @@ Lemma wt_setstack:
wt_locset ls -> wt_locset (Locmap.set (S sl ofs ty) v ls).
Proof.
intros; red; intros.
- unfold Locmap.set.
+ unfold Locmap.set.
destruct (Loc.eq (S sl ofs ty) l).
- subst l. simpl.
- generalize (Val.load_result_type (chunk_of_type ty) v).
+ subst l. simpl.
+ generalize (Val.load_result_type (chunk_of_type ty) v).
replace (type_of_chunk (chunk_of_type ty)) with ty. auto.
destruct ty; reflexivity.
destruct (Loc.diff_dec (S sl ofs ty) l). auto. red. auto.
@@ -164,7 +164,7 @@ Lemma wt_setlist:
Proof.
induction vl; destruct rl; simpl; intros; try contradiction.
auto.
- destruct H. apply IHvl; auto. apply wt_setreg; auto.
+ destruct H. apply IHvl; auto. apply wt_setreg; auto.
Qed.
Lemma wt_setres:
@@ -177,7 +177,7 @@ Proof.
induction res; simpl; intros.
- apply wt_setreg; auto. eapply Val.has_subtype; eauto.
- auto.
-- InvBooleans. eapply IHres2; eauto. destruct v; exact I.
+- InvBooleans. eapply IHres2; eauto. destruct v; exact I.
eapply IHres1; eauto. destruct v; exact I.
Qed.
@@ -189,7 +189,7 @@ Lemma wt_find_label:
Proof.
unfold wt_function; intros until c. generalize (fn_code f). induction c0; simpl; intros.
discriminate.
- InvBooleans. destruct (is_label lbl a).
+ InvBooleans. destruct (is_label lbl a).
congruence.
auto.
Qed.
@@ -250,15 +250,15 @@ Hypothesis wt_prog:
Lemma wt_find_function:
forall ros rs f, find_function ge ros rs = Some f -> wt_fundef f.
Proof.
- intros.
+ intros.
assert (X: exists i, In (i, Gfun f) prog.(prog_defs)).
{
destruct ros as [r | s]; simpl in H.
- eapply Genv.find_funct_inversion; eauto.
+ eapply Genv.find_funct_inversion; eauto.
destruct (Genv.find_symbol ge s) as [b|]; try discriminate.
eapply Genv.find_funct_ptr_inversion; eauto.
}
- destruct X as [i IN]. eapply wt_prog; eauto.
+ destruct X as [i IN]. eapply wt_prog; eauto.
Qed.
Theorem step_type_preservation:
@@ -266,38 +266,38 @@ Theorem step_type_preservation:
Proof.
induction 1; intros WTS; inv WTS.
- (* getstack *)
- simpl in *; InvBooleans.
+ simpl in *; InvBooleans.
econstructor; eauto.
- eapply wt_setreg; eauto. eapply Val.has_subtype; eauto. apply WTRS.
+ eapply wt_setreg; eauto. eapply Val.has_subtype; eauto. apply WTRS.
apply wt_undef_regs; auto.
- (* setstack *)
- simpl in *; InvBooleans.
+ simpl in *; InvBooleans.
econstructor; eauto.
apply wt_setstack. apply wt_undef_regs; auto.
- (* op *)
simpl in *. destruct (is_move_operation op args) as [src | ] eqn:ISMOVE.
+ (* move *)
- InvBooleans. exploit is_move_operation_correct; eauto. intros [EQ1 EQ2]; subst.
+ InvBooleans. exploit is_move_operation_correct; eauto. intros [EQ1 EQ2]; subst.
simpl in H. inv H.
- econstructor; eauto. apply wt_setreg. eapply Val.has_subtype; eauto. apply WTRS.
+ econstructor; eauto. apply wt_setreg. eapply Val.has_subtype; eauto. apply WTRS.
apply wt_undef_regs; auto.
- + (* other ops *)
+ + (* other ops *)
destruct (type_of_operation op) as [ty_args ty_res] eqn:TYOP. InvBooleans.
econstructor; eauto.
- apply wt_setreg; auto. eapply Val.has_subtype; eauto.
- change ty_res with (snd (ty_args, ty_res)). rewrite <- TYOP. eapply type_of_operation_sound; eauto.
- red; intros; subst op. simpl in ISMOVE.
- destruct args; try discriminate. destruct args; discriminate.
+ apply wt_setreg; auto. eapply Val.has_subtype; eauto.
+ change ty_res with (snd (ty_args, ty_res)). rewrite <- TYOP. eapply type_of_operation_sound; eauto.
+ red; intros; subst op. simpl in ISMOVE.
+ destruct args; try discriminate. destruct args; discriminate.
apply wt_undef_regs; auto.
- (* load *)
- simpl in *; InvBooleans.
+ simpl in *; InvBooleans.
econstructor; eauto.
- apply wt_setreg. eapply Val.has_subtype; eauto.
+ apply wt_setreg. eapply Val.has_subtype; eauto.
destruct a; simpl in H0; try discriminate. eapply Mem.load_type; eauto.
apply wt_undef_regs; auto.
- (* store *)
- simpl in *; InvBooleans.
- econstructor. eauto. eauto. eauto.
+ simpl in *; InvBooleans.
+ econstructor. eauto. eauto. eauto.
apply wt_undef_regs; auto.
- (* call *)
simpl in *; InvBooleans.
@@ -305,35 +305,35 @@ Proof.
eapply wt_find_function; eauto.
- (* tailcall *)
simpl in *; InvBooleans.
- econstructor; eauto.
+ econstructor; eauto.
eapply wt_find_function; eauto.
- apply wt_return_regs; auto. apply wt_parent_locset; auto.
+ apply wt_return_regs; auto. apply wt_parent_locset; auto.
- (* builtin *)
simpl in *; InvBooleans.
econstructor; eauto.
- eapply wt_setres; eauto. eapply external_call_well_typed; eauto.
+ eapply wt_setres; eauto. eapply external_call_well_typed; eauto.
apply wt_undef_regs; auto.
- (* label *)
simpl in *. econstructor; eauto.
- (* goto *)
- simpl in *. econstructor; eauto. eapply wt_find_label; eauto.
+ simpl in *. econstructor; eauto. eapply wt_find_label; eauto.
- (* cond branch, taken *)
simpl in *. econstructor. auto. auto. eapply wt_find_label; eauto.
apply wt_undef_regs; auto.
- (* cond branch, not taken *)
- simpl in *. econstructor. auto. auto. auto.
+ simpl in *. econstructor. auto. auto. auto.
apply wt_undef_regs; auto.
- (* jumptable *)
simpl in *. econstructor. auto. auto. eapply wt_find_label; eauto.
apply wt_undef_regs; auto.
- (* return *)
- simpl in *. InvBooleans.
+ simpl in *. InvBooleans.
econstructor; eauto.
apply wt_return_regs; auto. apply wt_parent_locset; auto.
- (* internal function *)
simpl in WTFD.
econstructor. eauto. eauto. eauto.
- apply wt_undef_regs. apply wt_call_regs. auto.
+ apply wt_undef_regs. apply wt_call_regs. auto.
- (* external function *)
econstructor. auto. apply wt_setlist; auto.
eapply Val.has_subtype_list. apply loc_result_type. eapply external_call_well_typed'; eauto.
@@ -344,7 +344,7 @@ Qed.
Theorem wt_initial_state:
forall S, initial_state prog S -> wt_state S.
Proof.
- induction 1. econstructor. constructor.
+ induction 1. econstructor. constructor.
unfold ge0 in H1. exploit Genv.find_funct_ptr_inversion; eauto.
intros [id IN]. eapply wt_prog; eauto.
apply wt_init.
@@ -383,7 +383,7 @@ Lemma wt_state_builtin:
wt_state (State s f sp (Lbuiltin ef args res :: c) rs m) ->
forallb (loc_valid f) (params_of_builtin_args args) = true.
Proof.
- intros. inv H. simpl in WTC; InvBooleans. auto.
+ intros. inv H. simpl in WTC; InvBooleans. auto.
Qed.
Lemma wt_callstate_wt_regs:
@@ -391,5 +391,5 @@ Lemma wt_callstate_wt_regs:
wt_state (Callstate s f rs m) ->
forall r, Val.has_type (rs (R r)) (mreg_type r).
Proof.
- intros. inv H. apply WTRS.
+ intros. inv H. apply WTRS.
Qed.
diff --git a/backend/Liveness.v b/backend/Liveness.v
index b8a5f965..16533158 100644
--- a/backend/Liveness.v
+++ b/backend/Liveness.v
@@ -60,7 +60,7 @@ Fixpoint reg_list_dead
an instruction is that a register is live before if either
it is one of the arguments of the instruction, or it is not the result
of the instruction and it is live after.
- However, if the result of a side-effect-free instruction is not
+ However, if the result of a side-effect-free instruction is not
live ``after'', the whole instruction will be removed later
(since it computes a useless result), thus its arguments need not
be live ``before''. *)
@@ -122,11 +122,11 @@ Lemma analyze_solution:
In s (successors_instr i) ->
Regset.Subset (transfer f s live!!s) live!!n.
Proof.
- unfold analyze; intros. eapply DS.fixpoint_solution; eauto.
- intros. unfold transfer; rewrite H2. apply DS.L.eq_refl.
+ unfold analyze; intros. eapply DS.fixpoint_solution; eauto.
+ intros. unfold transfer; rewrite H2. apply DS.L.eq_refl.
Qed.
-(** Given an RTL function, compute (for every PC) the list of
+(** Given an RTL function, compute (for every PC) the list of
pseudo-registers that are used for the last time in the instruction
at PC. These are the registers that are used or defined by the instruction
and dead afterwards. *)
@@ -145,4 +145,4 @@ Definition last_uses (f: function) : PTree.t (list reg) :=
| Some live => PTree.map (last_uses_at live) f.(fn_code)
end.
-
+
diff --git a/backend/Locations.v b/backend/Locations.v
index 439cd2dc..ea614585 100644
--- a/backend/Locations.v
+++ b/backend/Locations.v
@@ -35,13 +35,13 @@ Require Export Machregs.
(** A slot in an activation record is designated abstractly by a kind,
a type and an integer offset. Three kinds are considered:
-- [Local]: these are the slots used by register allocation for
+- [Local]: these are the slots used by register allocation for
pseudo-registers that cannot be assigned a hardware register.
- [Incoming]: used to store the parameters of the current function
- that cannot reside in hardware registers, as determined by the
+ that cannot reside in hardware registers, as determined by the
calling conventions.
-- [Outgoing]: used to store arguments to called functions that
- cannot reside in hardware registers, as determined by the
+- [Outgoing]: used to store arguments to called functions that
+ cannot reside in hardware registers, as determined by the
calling conventions. *)
Inductive slot: Type :=
@@ -111,19 +111,19 @@ Module Loc.
Defined.
(** As mentioned previously, two locations can be different (in the sense
- of the [<>] mathematical disequality), yet denote
+ of the [<>] mathematical disequality), yet denote
overlapping memory chunks within the activation record.
Given two locations, three cases are possible:
- They are equal (in the sense of the [=] equality)
- They are different and non-overlapping.
- They are different but overlapping.
- The second case (different and non-overlapping) is characterized
+ The second case (different and non-overlapping) is characterized
by the following [Loc.diff] predicate.
*)
Definition diff (l1 l2: loc) : Prop :=
match l1, l2 with
- | R r1, R r2 =>
+ | R r1, R r2 =>
r1 <> r2
| S s1 d1 t1, S s2 d2 t2 =>
s1 <> s2 \/ d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1
@@ -135,7 +135,7 @@ Module Loc.
forall l, ~(diff l l).
Proof.
destruct l; unfold diff; auto.
- red; intros. destruct H; auto. generalize (typesize_pos ty); omega.
+ red; intros. destruct H; auto. generalize (typesize_pos ty); omega.
Qed.
Lemma diff_not_eq:
@@ -162,7 +162,7 @@ Module Loc.
left; auto.
destruct (zle (pos0 + typesize ty0) pos).
left; auto.
- right; red; intros [P | [P | P]]. congruence. omega. omega.
+ right; red; intros [P | [P | P]]. congruence. omega. omega.
left; auto.
Defined.
@@ -181,9 +181,9 @@ Module Loc.
Lemma notin_iff:
forall l ll, notin l ll <-> (forall l', In l' ll -> Loc.diff l l').
Proof.
- induction ll; simpl.
+ induction ll; simpl.
tauto.
- rewrite IHll. intuition. subst a. auto.
+ rewrite IHll. intuition. subst a. auto.
Qed.
Lemma notin_not_in:
@@ -214,13 +214,13 @@ Module Loc.
forall a l1 l2,
disjoint (a :: l1) l2 -> disjoint l1 l2.
Proof.
- unfold disjoint; intros. auto with coqlib.
+ unfold disjoint; intros. auto with coqlib.
Qed.
Lemma disjoint_cons_right:
forall a l1 l2,
disjoint l1 (a :: l2) -> disjoint l1 l2.
Proof.
- unfold disjoint; intros. auto with coqlib.
+ unfold disjoint; intros. auto with coqlib.
Qed.
Lemma disjoint_sym:
@@ -232,20 +232,20 @@ Module Loc.
Lemma in_notin_diff:
forall l1 l2 ll, notin l1 ll -> In l2 ll -> diff l1 l2.
Proof.
- intros. rewrite notin_iff in H. auto.
+ intros. rewrite notin_iff in H. auto.
Qed.
Lemma notin_disjoint:
forall l1 l2,
(forall x, In x l1 -> notin x l2) -> disjoint l1 l2.
Proof.
- intros; red; intros. exploit H; eauto. rewrite notin_iff; intros. auto.
+ intros; red; intros. exploit H; eauto. rewrite notin_iff; intros. auto.
Qed.
Lemma disjoint_notin:
forall l1 l2 x, disjoint l1 l2 -> In x l1 -> notin x l2.
Proof.
- intros; rewrite notin_iff; intros. red in H. auto.
+ intros; rewrite notin_iff; intros. red in H. auto.
Qed.
(** [Loc.norepet ll] holds if the locations in list [ll] are pairwise
@@ -279,7 +279,7 @@ End Loc.
(** * Mappings from locations to values *)
(** The [Locmap] module defines mappings from locations to values,
- used as evaluation environments for the semantics of the [LTL]
+ used as evaluation environments for the semantics of the [LTL]
and [Linear] intermediate languages. *)
Set Implicit Arguments.
@@ -315,7 +315,7 @@ Module Locmap.
else Vundef.
Lemma gss: forall l v m,
- (set l v m) l =
+ (set l v m) l =
match l with R r => v | S sl ofs ty => Val.load_result (chunk_of_type ty) v end.
Proof.
intros. unfold set. apply dec_eq_true.
@@ -328,7 +328,7 @@ Module Locmap.
Lemma gss_typed: forall l v m, Val.has_type v (Loc.type l) -> (set l v m) l = v.
Proof.
- intros. rewrite gss. destruct l. auto. apply Val.load_result_same; auto.
+ intros. rewrite gss. destruct l. auto. apply Val.load_result_same; auto.
Qed.
Lemma gso: forall l v m p, Loc.diff l p -> (set l v m) p = m p.
@@ -348,19 +348,19 @@ Module Locmap.
Lemma guo: forall ll l m, Loc.notin l ll -> (undef ll m) l = m l.
Proof.
- induction ll; simpl; intros. auto.
- destruct H. rewrite IHll; auto. apply gso. apply Loc.diff_sym; auto.
+ induction ll; simpl; intros. auto.
+ destruct H. rewrite IHll; auto. apply gso. apply Loc.diff_sym; auto.
Qed.
Lemma gus: forall ll l m, In l ll -> (undef ll m) l = Vundef.
Proof.
assert (P: forall ll l m, m l = Vundef -> (undef ll m) l = Vundef).
- induction ll; simpl; intros. auto. apply IHll.
+ induction ll; simpl; intros. auto. apply IHll.
unfold set. destruct (Loc.eq a l).
- destruct a. auto. destruct ty; reflexivity.
+ destruct a. auto. destruct ty; reflexivity.
destruct (Loc.diff_dec a l); auto.
- induction ll; simpl; intros. contradiction.
- destruct H. apply P. subst a. apply gss_typed. exact I.
+ induction ll; simpl; intros. contradiction.
+ destruct H. apply P. subst a. apply gss_typed. exact I.
auto.
Qed.
@@ -372,7 +372,7 @@ Module Locmap.
Lemma gsetlisto: forall l ll vl m, Loc.notin l ll -> (setlist ll vl m) l = m l.
Proof.
- induction ll; simpl; intros.
+ induction ll; simpl; intros.
auto.
destruct vl; auto. destruct H. rewrite IHll; auto. apply gso; auto. apply Loc.diff_sym; auto.
Qed.
@@ -381,7 +381,7 @@ Module Locmap.
match res with
| BR r => set (R r) v m
| BR_none => m
- | BR_splitlong hi lo =>
+ | BR_splitlong hi lo =>
setres lo (Val.loword v) (setres hi (Val.hiword v) m)
end.
@@ -431,53 +431,53 @@ Module OrderedLoc <: OrderedType.
(ofs1 < ofs2 \/ (ofs1 = ofs2 /\ OrderedTyp.lt ty1 ty2)))
end.
Lemma eq_refl : forall x : t, eq x x.
- Proof (@refl_equal t).
+ 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.
+ unfold lt; intros.
destruct x; destruct y; destruct z; try tauto.
eapply Plt_trans; eauto.
- destruct H.
+ destruct H.
destruct H0. left; eapply OrderedSlot.lt_trans; eauto.
- destruct H0. subst sl0. auto.
+ destruct H0. subst sl0. auto.
destruct H. subst sl.
destruct H0. auto.
- destruct H.
+ destruct H.
right. split. auto.
intuition.
- right; split. congruence. eapply OrderedTyp.lt_trans; eauto.
+ 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.
+ 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 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 LT. red. auto.
+ + apply EQ. red. f_equal. apply IndexedMreg.index_inj. auto.
+ apply GT. red. auto.
- - apply LT. 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.
+ * 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 EQ. red; red in e; red in e0; red in e1. congruence.
apply GT. red; auto.
- * apply GT. red. auto.
+ * apply GT. red. auto.
+ apply GT. red; auto.
Defined.
Definition eq_dec := Loc.eq.
@@ -499,21 +499,21 @@ Module OrderedLoc <: OrderedType.
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.
+ 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.
+ + 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 \/ ty' = Tsingle \/ ty' = Tany32).
+ destruct H0. right. generalize (RANGE ty'); omega.
+ destruct H0.
+ assert (ty' = Tint \/ ty' = Tsingle \/ ty' = Tany32).
{ unfold OrderedTyp.lt in H1. destruct ty'; auto; compute in H1; congruence. }
- right. destruct H2 as [E|[E|E]]; subst ty'; simpl typesize; omega.
+ right. destruct H2 as [E|[E|E]]; subst ty'; simpl typesize; omega.
+ destruct H. left. apply OrderedSlot.lt_not_eq; auto.
destruct H. right.
destruct H0. left; omega.
@@ -523,23 +523,23 @@ Module OrderedLoc <: OrderedType.
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.
+ 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.
+ auto.
+ rewrite Pos.compare_antisym. rewrite C. auto.
- destruct (OrderedSlot.compare sl sl'); auto.
- destruct H. contradiction.
+ destruct H. contradiction.
destruct H.
- right; right; split; auto. left; omega.
+ right; right; split; auto. left; omega.
left; right; split; auto.
assert (EITHER: typesize ty' = 1 /\ OrderedTyp.lt ty' Tany64 \/ typesize ty' = 2).
{ destruct ty'; compute; auto. }
destruct (zlt ofs' (ofs - 1)). left; auto.
destruct EITHER as [[P Q] | P].
right; split; auto. omega.
- left; omega.
+ left; omega.
Qed.
End OrderedLoc.
diff --git a/backend/Mach.v b/backend/Mach.v
index 8853d9da..739c8212 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -13,7 +13,7 @@
(** The Mach intermediate language: abstract syntax.
Mach is the last intermediate language before generation of assembly
- code.
+ code.
*)
Require Import Coqlib.
@@ -34,14 +34,14 @@ Require Stacklayout.
(** Like Linear, the Mach language is organized as lists of instructions
operating over machine registers, with default fall-through behaviour
- and explicit labels and branch instructions.
+ and explicit labels and branch instructions.
The main difference with Linear lies in the instructions used to
access the activation record. Mach has three such instructions:
[Mgetstack] and [Msetstack] to read and write within the activation
record for the current function, at a given word offset and with a
given type; and [Mgetparam], to read within the activation record of
- the caller.
+ the caller.
These instructions implement a more concrete view of the activation
record than the the [Lgetstack] and [Lsetstack] instructions of
@@ -153,7 +153,7 @@ Lemma undef_regs_same:
Proof.
induction rl; simpl; intros. tauto.
destruct H. subst a. apply Regmap.gss.
- unfold Regmap.set. destruct (RegEq.eq r a); auto.
+ unfold Regmap.set. destruct (RegEq.eq r a); auto.
Qed.
Fixpoint set_regs (rl: list mreg) (vl: list val) (rs: regset) : regset :=
@@ -193,7 +193,7 @@ Lemma find_label_incl:
forall lbl c c', find_label lbl c = Some c' -> incl c' c.
Proof.
induction c; simpl; intros. discriminate.
- destruct (is_label lbl a). inv H. auto with coqlib. eauto with coqlib.
+ destruct (is_label lbl a). inv H. auto with coqlib. eauto with coqlib.
Qed.
Section RELSEM.
diff --git a/backend/NeedDomain.v b/backend/NeedDomain.v
index 770648b1..e40c1322 100644
--- a/backend/NeedDomain.v
+++ b/backend/NeedDomain.v
@@ -53,7 +53,7 @@ Fixpoint vagree (v w: val) (x: nval) {struct x}: Prop :=
match v, w with
| Vint p, Vint q => iagree p q m
| Vint p, _ => False
- | _, _ => True
+ | _, _ => True
end
| All => Val.lessdef v w
end.
@@ -143,7 +143,7 @@ Lemma nge_lub_l:
forall x y, nge (nlub x y) x.
Proof.
unfold nlub; destruct x, y; auto with na.
- constructor. intros. autorewrite with ints; auto. rewrite H0; auto.
+ constructor. intros. autorewrite with ints; auto. rewrite H0; auto.
Qed.
Lemma nge_lub_r:
@@ -171,13 +171,13 @@ Lemma iagree_and_eq:
forall x y mask,
iagree x y mask <-> Int.and x mask = Int.and y mask.
Proof.
- intros; split; intros.
-- Int.bit_solve. specialize (H i H0).
- destruct (Int.testbit mask i).
- rewrite ! andb_true_r; auto.
+ intros; split; intros.
+- Int.bit_solve. specialize (H i H0).
+ destruct (Int.testbit mask i).
+ rewrite ! andb_true_r; auto.
rewrite ! andb_false_r; auto.
- red; intros. exploit (eq_same_bits i); eauto; autorewrite with ints; auto.
- rewrite H1. rewrite ! andb_true_r; auto.
+ rewrite H1. rewrite ! andb_true_r; auto.
Qed.
Lemma iagree_mone:
@@ -203,7 +203,7 @@ Qed.
Lemma iagree_not:
forall x y m, iagree x y m -> iagree (Int.not x) (Int.not y) m.
Proof.
- intros; red; intros; autorewrite with ints; auto. f_equal; auto.
+ intros; red; intros; autorewrite with ints; auto. f_equal; auto.
Qed.
Lemma iagree_not':
@@ -217,7 +217,7 @@ Lemma iagree_or:
forall x y n m,
iagree x y (Int.and m (Int.not n)) -> iagree (Int.or x n) (Int.or y n) m.
Proof.
- intros. apply iagree_not'. rewrite ! Int.not_or_and_not. apply iagree_and.
+ intros. apply iagree_not'. rewrite ! Int.not_or_and_not. apply iagree_and.
apply iagree_not; auto.
Qed.
@@ -228,19 +228,19 @@ Lemma iagree_bitwise_binop:
forall x1 x2 y1 y2 m,
iagree x1 y1 m -> iagree x2 y2 m -> iagree (f x1 x2) (f y1 y2) m.
Proof.
- intros; red; intros. rewrite ! H by auto. f_equal; auto.
+ intros; red; intros. rewrite ! H by auto. f_equal; auto.
Qed.
Lemma iagree_shl:
forall x y m n,
iagree x y (Int.shru m n) -> iagree (Int.shl x n) (Int.shl y n) m.
Proof.
- intros; red; intros. autorewrite with ints; auto.
+ intros; red; intros. autorewrite with ints; auto.
destruct (zlt i (Int.unsigned n)).
- auto.
-- generalize (Int.unsigned_range n); intros.
- apply H. omega. rewrite Int.bits_shru by omega.
- replace (i - Int.unsigned n + Int.unsigned n) with i by omega.
+- generalize (Int.unsigned_range n); intros.
+ apply H. omega. rewrite Int.bits_shru by omega.
+ replace (i - Int.unsigned n + Int.unsigned n) with i by omega.
rewrite zlt_true by omega. auto.
Qed.
@@ -248,11 +248,11 @@ Lemma iagree_shru:
forall x y m n,
iagree x y (Int.shl m n) -> iagree (Int.shru x n) (Int.shru y n) m.
Proof.
- intros; red; intros. autorewrite with ints; auto.
+ intros; red; intros. autorewrite with ints; auto.
destruct (zlt (i + Int.unsigned n) Int.zwordsize).
-- generalize (Int.unsigned_range n); intros.
- apply H. omega. rewrite Int.bits_shl by omega.
- replace (i + Int.unsigned n - Int.unsigned n) with i by omega.
+- generalize (Int.unsigned_range n); intros.
+ apply H. omega. rewrite Int.bits_shl by omega.
+ replace (i + Int.unsigned n - Int.unsigned n) with i by omega.
rewrite zlt_false by omega. auto.
- auto.
Qed.
@@ -262,8 +262,8 @@ Lemma iagree_shr_1:
Int.shru (Int.shl m n) n = m ->
iagree x y (Int.shl m n) -> iagree (Int.shr x n) (Int.shr y n) m.
Proof.
- intros; red; intros. rewrite <- H in H2. rewrite Int.bits_shru in H2 by auto.
- rewrite ! Int.bits_shr by auto.
+ intros; red; intros. rewrite <- H in H2. rewrite Int.bits_shru in H2 by auto.
+ rewrite ! Int.bits_shr by auto.
destruct (zlt (i + Int.unsigned n) Int.zwordsize).
- apply H0; auto. generalize (Int.unsigned_range n); omega.
- discriminate.
@@ -274,17 +274,17 @@ Lemma iagree_shr:
iagree x y (Int.or (Int.shl m n) (Int.repr Int.min_signed)) ->
iagree (Int.shr x n) (Int.shr y n) m.
Proof.
- intros; red; intros. rewrite ! Int.bits_shr by auto.
- generalize (Int.unsigned_range n); intros.
+ intros; red; intros. rewrite ! Int.bits_shr by auto.
+ generalize (Int.unsigned_range n); intros.
set (j := if zlt (i + Int.unsigned n) Int.zwordsize
then i + Int.unsigned n
else Int.zwordsize - 1).
assert (0 <= j < Int.zwordsize).
{ unfold j; destruct (zlt (i + Int.unsigned n) Int.zwordsize); omega. }
- apply H; auto. autorewrite with ints; auto. apply orb_true_intro.
+ apply H; auto. autorewrite with ints; auto. apply orb_true_intro.
unfold j; destruct (zlt (i + Int.unsigned n) Int.zwordsize).
-- left. rewrite zlt_false by omega.
- replace (i + Int.unsigned n - Int.unsigned n) with i by omega.
+- left. rewrite zlt_false by omega.
+ replace (i + Int.unsigned n - Int.unsigned n) with i by omega.
auto.
- right. reflexivity.
Qed.
@@ -302,14 +302,14 @@ Proof.
mod Int.zwordsize) with i. auto.
apply Int.eqmod_small_eq with Int.zwordsize; auto.
apply Int.eqmod_trans with ((i - Int.unsigned amount) + Int.unsigned amount).
- apply Int.eqmod_refl2; omega.
+ apply Int.eqmod_refl2; omega.
eapply Int.eqmod_trans. 2: apply Int.eqmod_mod; auto.
- apply Int.eqmod_add.
- apply Int.eqmod_mod; auto.
- apply Int.eqmod_refl.
- apply Z_mod_lt; auto.
+ apply Int.eqmod_add.
+ apply Int.eqmod_mod; auto.
+ apply Int.eqmod_refl.
apply Z_mod_lt; auto.
-Qed.
+ apply Z_mod_lt; auto.
+Qed.
Lemma iagree_ror:
forall p q m amount,
@@ -317,9 +317,9 @@ Lemma iagree_ror:
iagree (Int.ror p amount) (Int.ror q amount) m.
Proof.
intros. rewrite ! Int.ror_rol_neg by apply int_wordsize_divides_modulus.
- apply iagree_rol.
+ apply iagree_rol.
rewrite Int.ror_rol_neg by apply int_wordsize_divides_modulus.
- rewrite Int.neg_involutive; auto.
+ rewrite Int.neg_involutive; auto.
Qed.
Lemma eqmod_iagree:
@@ -327,14 +327,14 @@ Lemma eqmod_iagree:
Int.eqmod (two_p (Int.size m)) x y ->
iagree (Int.repr x) (Int.repr y) m.
Proof.
- intros. set (p := nat_of_Z (Int.size m)).
+ intros. set (p := nat_of_Z (Int.size m)).
generalize (Int.size_range m); intros RANGE.
assert (EQ: Int.size m = Z_of_nat p). { symmetry; apply nat_of_Z_eq. omega. }
rewrite EQ in H; rewrite <- two_power_nat_two_p in H.
red; intros. rewrite ! Int.testbit_repr by auto.
- destruct (zlt i (Int.size m)).
+ destruct (zlt i (Int.size m)).
eapply Int.same_bits_eqmod; eauto. omega.
- assert (Int.testbit m i = false) by (eapply Int.bits_size_2; omega).
+ assert (Int.testbit m i = false) by (eapply Int.bits_size_2; omega).
congruence.
Qed.
@@ -345,12 +345,12 @@ Lemma iagree_eqmod:
iagree x y (complete_mask m) ->
Int.eqmod (two_p (Int.size m)) (Int.unsigned x) (Int.unsigned y).
Proof.
- intros. set (p := nat_of_Z (Int.size m)).
+ intros. set (p := nat_of_Z (Int.size m)).
generalize (Int.size_range m); intros RANGE.
assert (EQ: Int.size m = Z_of_nat p). { symmetry; apply nat_of_Z_eq. omega. }
- rewrite EQ; rewrite <- two_power_nat_two_p.
- apply Int.eqmod_same_bits. intros. apply H. omega.
- unfold complete_mask. rewrite Int.bits_zero_ext by omega.
+ rewrite EQ; rewrite <- two_power_nat_two_p.
+ apply Int.eqmod_same_bits. intros. apply H. omega.
+ unfold complete_mask. rewrite Int.bits_zero_ext by omega.
rewrite zlt_true by omega. rewrite Int.bits_mone by omega. auto.
Qed.
@@ -361,13 +361,13 @@ Proof.
+ subst m; reflexivity.
+ assert (Int.unsigned m <> 0).
{ red; intros; elim n. rewrite <- (Int.repr_unsigned m). rewrite H; auto. }
- assert (0 < Int.size m).
+ assert (0 < Int.size m).
{ apply Int.Zsize_pos'. generalize (Int.unsigned_range m); omega. }
generalize (Int.size_range m); intros.
- f_equal. apply Int.bits_size_4. tauto.
+ f_equal. apply Int.bits_size_4. tauto.
rewrite Int.bits_zero_ext by omega. rewrite zlt_true by omega.
apply Int.bits_mone; omega.
- intros. rewrite Int.bits_zero_ext by omega. apply zlt_false; omega.
+ intros. rewrite Int.bits_zero_ext by omega. apply zlt_false; omega.
Qed.
(** ** Abstract operations over value needs. *)
@@ -416,7 +416,7 @@ Proof.
unfold orimm; intros; destruct x; simpl in *.
- auto.
- unfold Val.or; InvAgree. apply iagree_or; auto.
-- InvAgree. simpl. apply Val.lessdef_same. f_equal. apply iagree_mone.
+- InvAgree. simpl. apply Val.lessdef_same. f_equal. apply iagree_mone.
apply iagree_or. rewrite Int.and_commut. rewrite Int.and_mone. auto.
Qed.
@@ -438,8 +438,8 @@ Lemma vagree_bitwise_binop:
x.
Proof.
unfold bitwise; intros. destruct x; simpl in *.
-- auto.
-- InvAgree.
+- auto.
+- InvAgree.
- inv H0; auto. inv H1; auto. destruct w1; auto.
Qed.
@@ -482,8 +482,8 @@ Lemma shlimm_sound:
vagree v w (shlimm x n) ->
vagree (Val.shl v (Vint n)) (Val.shl w (Vint n)) x.
Proof.
- unfold shlimm; intros. unfold Val.shl.
- destruct (Int.ltu n Int.iwordsize).
+ unfold shlimm; intros. unfold Val.shl.
+ destruct (Int.ltu n Int.iwordsize).
destruct x; simpl in *.
- auto.
- InvAgree. apply iagree_shl; auto.
@@ -504,7 +504,7 @@ Lemma shruimm_sound:
vagree (Val.shru v (Vint n)) (Val.shru w (Vint n)) x.
Proof.
unfold shruimm; intros. unfold Val.shru.
- destruct (Int.ltu n Int.iwordsize).
+ destruct (Int.ltu n Int.iwordsize).
destruct x; simpl in *.
- auto.
- InvAgree. apply iagree_shru; auto.
@@ -528,10 +528,10 @@ Lemma shrimm_sound:
vagree (Val.shr v (Vint n)) (Val.shr w (Vint n)) x.
Proof.
unfold shrimm; intros. unfold Val.shr.
- destruct (Int.ltu n Int.iwordsize).
+ destruct (Int.ltu n Int.iwordsize).
destruct x; simpl in *.
- auto.
-- InvAgree.
+- InvAgree.
destruct (Int.eq_dec (Int.shru (Int.shl m n) n) m).
apply iagree_shr_1; auto.
apply iagree_shr; auto.
@@ -553,10 +553,10 @@ Lemma rolm_sound:
Proof.
unfold rolm; intros; destruct x; simpl in *.
- auto.
-- unfold Val.rolm; InvAgree. unfold Int.rolm.
- apply iagree_and. apply iagree_rol. auto.
-- unfold Val.rolm; InvAgree. apply Val.lessdef_same. f_equal. apply iagree_mone.
- unfold Int.rolm. apply iagree_and. apply iagree_rol. rewrite Int.and_commut.
+- unfold Val.rolm; InvAgree. unfold Int.rolm.
+ apply iagree_and. apply iagree_rol. auto.
+- unfold Val.rolm; InvAgree. apply Val.lessdef_same. f_equal. apply iagree_mone.
+ unfold Int.rolm. apply iagree_and. apply iagree_rol. rewrite Int.and_commut.
rewrite Int.and_mone. auto.
Qed.
@@ -573,15 +573,15 @@ Lemma ror_sound:
vagree (Val.ror v (Vint n)) (Val.ror w (Vint n)) x.
Proof.
unfold ror; intros. unfold Val.ror.
- destruct (Int.ltu n Int.iwordsize).
+ destruct (Int.ltu n Int.iwordsize).
destruct x; simpl in *.
- auto.
-- InvAgree. apply iagree_ror; auto.
+- InvAgree. apply iagree_ror; auto.
- inv H; auto.
- destruct v; auto with na.
Qed.
-(** Modular arithmetic operations: add, mul, opposite.
+(** Modular arithmetic operations: add, mul, opposite.
(But not subtraction because of the pointer - pointer case. *)
Definition modarith (x: nval) :=
@@ -596,10 +596,10 @@ Lemma add_sound:
vagree v1 w1 (modarith x) -> vagree v2 w2 (modarith x) ->
vagree (Val.add v1 v2) (Val.add w1 w2) x.
Proof.
- unfold modarith; intros. destruct x; simpl in *.
+ unfold modarith; intros. destruct x; simpl in *.
- auto.
-- unfold Val.add; InvAgree. apply eqmod_iagree. apply Int.eqmod_add; apply iagree_eqmod; auto.
-- inv H; auto. inv H0; auto. destruct w1; auto.
+- unfold Val.add; InvAgree. apply eqmod_iagree. apply Int.eqmod_add; apply iagree_eqmod; auto.
+- inv H; auto. inv H0; auto. destruct w1; auto.
Qed.
Remark modarith_idem: forall nv, modarith (modarith nv) = modarith nv.
@@ -612,10 +612,10 @@ Lemma mul_sound:
vagree v1 w1 (modarith x) -> vagree v2 w2 (modarith x) ->
vagree (Val.mul v1 v2) (Val.mul w1 w2) x.
Proof.
- unfold mul, add; intros. destruct x; simpl in *.
+ unfold mul, add; intros. destruct x; simpl in *.
- auto.
-- unfold Val.mul; InvAgree. apply eqmod_iagree. apply Int.eqmod_mult; apply iagree_eqmod; auto.
-- inv H; auto. inv H0; auto. destruct w1; auto.
+- unfold Val.mul; InvAgree. apply eqmod_iagree. apply Int.eqmod_mult; apply iagree_eqmod; auto.
+- inv H; auto. inv H0; auto. destruct w1; auto.
Qed.
Lemma neg_sound:
@@ -625,8 +625,8 @@ Lemma neg_sound:
Proof.
intros; destruct x; simpl in *.
- auto.
-- unfold Val.neg; InvAgree.
- apply eqmod_iagree. apply Int.eqmod_neg. apply iagree_eqmod; auto.
+- unfold Val.neg; InvAgree.
+ apply eqmod_iagree. apply Int.eqmod_neg. apply iagree_eqmod; auto.
- inv H; simpl; auto.
Qed.
@@ -648,12 +648,12 @@ Proof.
unfold zero_ext; intros.
destruct x; simpl in *.
- auto.
-- unfold Val.zero_ext; InvAgree.
- red; intros. autorewrite with ints; try omega.
+- unfold Val.zero_ext; InvAgree.
+ red; intros. autorewrite with ints; try omega.
destruct (zlt i1 n); auto. apply H; auto.
autorewrite with ints; try omega. rewrite zlt_true; auto.
-- unfold Val.zero_ext; InvAgree; auto. apply Val.lessdef_same. f_equal.
- Int.bit_solve; try omega. destruct (zlt i1 n); auto. apply H; auto.
+- unfold Val.zero_ext; InvAgree; auto. apply Val.lessdef_same. f_equal.
+ Int.bit_solve; try omega. destruct (zlt i1 n); auto. apply H; auto.
autorewrite with ints; try omega. apply zlt_true; auto.
Qed.
@@ -672,25 +672,25 @@ Lemma sign_ext_sound:
Proof.
unfold sign_ext; intros. destruct x; simpl in *.
- auto.
-- unfold Val.sign_ext; InvAgree.
+- unfold Val.sign_ext; InvAgree.
red; intros. autorewrite with ints; try omega.
set (j := if zlt i1 n then i1 else n - 1).
- assert (0 <= j < Int.zwordsize).
+ assert (0 <= j < Int.zwordsize).
{ unfold j; destruct (zlt i1 n); omega. }
- apply H; auto.
- autorewrite with ints; try omega. apply orb_true_intro.
- unfold j; destruct (zlt i1 n).
- left. rewrite zlt_true; auto.
- right. rewrite Int.unsigned_repr. rewrite zlt_false by omega.
- replace (n - 1 - (n - 1)) with 0 by omega. reflexivity.
+ apply H; auto.
+ autorewrite with ints; try omega. apply orb_true_intro.
+ unfold j; destruct (zlt i1 n).
+ left. rewrite zlt_true; auto.
+ right. rewrite Int.unsigned_repr. rewrite zlt_false by omega.
+ replace (n - 1 - (n - 1)) with 0 by omega. reflexivity.
generalize Int.wordsize_max_unsigned; omega.
-- unfold Val.sign_ext; InvAgree; auto. apply Val.lessdef_same. f_equal.
+- unfold Val.sign_ext; InvAgree; auto. apply Val.lessdef_same. f_equal.
Int.bit_solve; try omega.
set (j := if zlt i1 n then i1 else n - 1).
- assert (0 <= j < Int.zwordsize).
+ assert (0 <= j < Int.zwordsize).
{ unfold j; destruct (zlt i1 n); omega. }
- apply H; auto. rewrite Int.bits_zero_ext; try omega.
- rewrite zlt_true. apply Int.bits_mone; auto.
+ apply H; auto. rewrite Int.bits_zero_ext; try omega.
+ rewrite zlt_true. apply Int.bits_mone; auto.
unfold j. destruct (zlt i1 n); omega.
Qed.
@@ -713,25 +713,25 @@ Proof.
(list_repeat (size_chunk_nat chunk) Undef)
(encode_val chunk w)).
{
- rewrite <- (encode_val_length chunk w).
+ rewrite <- (encode_val_length chunk w).
apply repeat_Undef_inject_any.
}
assert (SAME: forall vl1 vl2,
vl1 = vl2 ->
list_forall2 memval_lessdef vl1 vl2).
{
- intros. subst vl2. revert vl1. induction vl1; constructor; auto.
- apply memval_lessdef_refl.
+ intros. subst vl2. revert vl1. induction vl1; constructor; auto.
+ apply memval_lessdef_refl.
}
intros. unfold store_argument in H; destruct chunk.
-- InvAgree. apply SAME. simpl; f_equal. apply encode_int_8_mod.
+- InvAgree. apply SAME. simpl; f_equal. apply encode_int_8_mod.
change 8 with (Int.size (Int.repr 255)). apply iagree_eqmod; auto.
-- InvAgree. apply SAME. simpl; f_equal. apply encode_int_8_mod.
+- InvAgree. apply SAME. simpl; f_equal. apply encode_int_8_mod.
change 8 with (Int.size (Int.repr 255)). apply iagree_eqmod; auto.
-- InvAgree. apply SAME. simpl; f_equal. apply encode_int_16_mod.
+- InvAgree. apply SAME. simpl; f_equal. apply encode_int_16_mod.
change 16 with (Int.size (Int.repr 65535)). apply iagree_eqmod; auto.
-- InvAgree. apply SAME. simpl; f_equal. apply encode_int_16_mod.
+- InvAgree. apply SAME. simpl; f_equal. apply encode_int_16_mod.
change 16 with (Int.size (Int.repr 65535)). apply iagree_eqmod; auto.
- apply encode_val_inject. rewrite val_inject_id; auto.
- apply encode_val_inject. rewrite val_inject_id; auto.
@@ -768,7 +768,7 @@ Lemma maskzero_sound:
Val.maskzero_bool v n = Some b ->
Val.maskzero_bool w n = Some b.
Proof.
- unfold maskzero; intros.
+ unfold maskzero; intros.
unfold Val.maskzero_bool; InvAgree; try discriminate.
inv H0. rewrite iagree_and_eq in H. rewrite H. auto.
Qed.
@@ -795,9 +795,9 @@ Let valid_pointer_inj:
Mem.valid_pointer m1 b1 (Int.unsigned ofs) = true ->
Mem.valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
Proof.
- unfold inject_id; intros. inv H. rewrite Int.add_zero.
+ unfold inject_id; intros. inv H. rewrite Int.add_zero.
rewrite Mem.valid_pointer_nonempty_perm in *. eauto.
-Qed.
+Qed.
Let weak_valid_pointer_inj:
forall b1 ofs b2 delta,
@@ -805,7 +805,7 @@ Let weak_valid_pointer_inj:
Mem.weak_valid_pointer m1 b1 (Int.unsigned ofs) = true ->
Mem.weak_valid_pointer m2 b2 (Int.unsigned (Int.add ofs (Int.repr delta))) = true.
Proof.
- unfold inject_id; intros. inv H. rewrite Int.add_zero.
+ unfold inject_id; intros. inv H. rewrite Int.add_zero.
rewrite Mem.weak_valid_pointer_spec in *.
rewrite ! Mem.valid_pointer_nonempty_perm in *.
destruct H0; [left|right]; eauto.
@@ -830,7 +830,7 @@ Let valid_different_pointers_inj:
b1' <> b2' \/
Int.unsigned (Int.add ofs1 (Int.repr delta1)) <> Int.unsigned (Int.add ofs2 (Int.repr delta2)).
Proof.
- unfold inject_id; intros. left; congruence.
+ unfold inject_id; intros. left; congruence.
Qed.
Lemma default_needs_of_condition_sound:
@@ -846,7 +846,7 @@ Qed.
Lemma default_needs_of_operation_sound:
forall op args1 v1 args2 nv,
eval_operation ge (Vptr sp Int.zero) op args1 m1 = Some v1 ->
- vagree_list args1 args2 nil
+ vagree_list args1 args2 nil
\/ vagree_list args1 args2 (default nv :: nil)
\/ vagree_list args1 args2 (default nv :: default nv :: nil) ->
nv <> Nothing ->
@@ -854,12 +854,12 @@ Lemma default_needs_of_operation_sound:
eval_operation ge (Vptr sp Int.zero) op args2 m2 = Some v2
/\ vagree v1 v2 nv.
Proof.
- intros. assert (default nv = All) by (destruct nv; simpl; congruence).
+ intros. assert (default nv = All) by (destruct nv; simpl; congruence).
rewrite H2 in H0.
assert (Val.lessdef_list args1 args2).
{
- destruct H0. auto with na.
- destruct H0. inv H0; constructor; auto with na.
+ destruct H0. auto with na.
+ destruct H0. inv H0; constructor; auto with na.
inv H0; constructor; auto with na. inv H8; constructor; auto with na.
}
exploit (@eval_operation_inj _ _ _ _ ge ge inject_id).
@@ -869,7 +869,7 @@ Proof.
apply val_inject_list_lessdef; eauto.
eauto.
intros (v2 & A & B). exists v2; split; auto.
- apply vagree_lessdef. apply val_inject_lessdef. auto.
+ apply vagree_lessdef. apply val_inject_lessdef. auto.
Qed.
End DEFAULT.
@@ -890,12 +890,12 @@ Lemma andimm_redundant_sound:
vagree (Val.and v (Vint n)) w x.
Proof.
unfold andimm_redundant; intros. destruct x; try discriminate.
-- simpl; auto.
+- simpl; auto.
- InvBooleans. simpl in *; unfold Val.and; InvAgree.
- red; intros. exploit (eq_same_bits i1); eauto.
- autorewrite with ints; auto. rewrite H2; simpl; intros.
- destruct (Int.testbit n i1) eqn:N; try discriminate.
- rewrite andb_true_r. apply H0; auto. autorewrite with ints; auto.
+ red; intros. exploit (eq_same_bits i1); eauto.
+ autorewrite with ints; auto. rewrite H2; simpl; intros.
+ destruct (Int.testbit n i1) eqn:N; try discriminate.
+ rewrite andb_true_r. apply H0; auto. autorewrite with ints; auto.
rewrite H2, N; auto.
Qed.
@@ -915,7 +915,7 @@ Proof.
unfold orimm_redundant; intros. destruct x; try discriminate.
- auto.
- InvBooleans. simpl in *; unfold Val.or; InvAgree.
- apply iagree_not'. rewrite Int.not_or_and_not.
+ apply iagree_not'. rewrite Int.not_or_and_not.
apply (andimm_redundant_sound (Vint (Int.not i)) (Vint (Int.not i0)) (I m) (Int.not n)).
simpl. rewrite Int.not_involutive. apply proj_sumbool_is_true. auto.
simpl. apply iagree_not; auto.
@@ -933,9 +933,9 @@ Proof.
unfold rolm_redundant; intros; InvBooleans. subst amount. rewrite Val.rolm_zero.
apply andimm_redundant_sound; auto.
assert (forall n, Int.ror n Int.zero = n).
- { intros. rewrite Int.ror_rol_neg by apply int_wordsize_divides_modulus.
+ { intros. rewrite Int.ror_rol_neg by apply int_wordsize_divides_modulus.
rewrite Int.neg_zero. apply Int.rol_zero. }
- unfold rolm, andimm in *. destruct x; auto.
+ unfold rolm, andimm in *. destruct x; auto.
rewrite H in H0. auto.
rewrite H in H0. auto.
Qed.
@@ -956,8 +956,8 @@ Lemma zero_ext_redundant_sound:
Proof.
unfold zero_ext_redundant; intros. destruct x; try discriminate.
- auto.
-- simpl in *; InvAgree. simpl. InvBooleans. rewrite <- H.
- red; intros; autorewrite with ints; try omega.
+- simpl in *; InvAgree. simpl. InvBooleans. rewrite <- H.
+ red; intros; autorewrite with ints; try omega.
destruct (zlt i1 n). apply H0; auto.
rewrite Int.bits_zero_ext in H3 by omega. rewrite zlt_false in H3 by auto. discriminate.
Qed.
@@ -978,10 +978,10 @@ Lemma sign_ext_redundant_sound:
Proof.
unfold sign_ext_redundant; intros. destruct x; try discriminate.
- auto.
-- simpl in *; InvAgree. simpl. InvBooleans. rewrite <- H.
- red; intros; autorewrite with ints; try omega.
+- simpl in *; InvAgree. simpl. InvBooleans. rewrite <- H.
+ red; intros; autorewrite with ints; try omega.
destruct (zlt i1 n). apply H0; auto.
- rewrite Int.bits_or; auto. rewrite H3; auto.
+ rewrite Int.bits_or; auto. rewrite H3; auto.
rewrite Int.bits_zero_ext in H3 by omega. rewrite zlt_false in H3 by auto. discriminate.
Qed.
@@ -1014,7 +1014,7 @@ End NVal.
Module NE := LPMap1(NVal).
-Definition nenv := NE.t.
+Definition nenv := NE.t.
Definition nreg (ne: nenv) (r: reg) := NE.get r ne.
@@ -1024,7 +1024,7 @@ Definition eagree (e1 e2: regset) (ne: nenv) : Prop :=
Lemma nreg_agree:
forall rs1 rs2 ne r, eagree rs1 rs2 ne -> vagree rs1#r rs2#r (nreg ne r).
Proof.
- intros. apply H.
+ intros. apply H.
Qed.
Hint Resolve nreg_agree: na.
@@ -1033,7 +1033,7 @@ Lemma eagree_ge:
forall e1 e2 ne ne',
eagree e1 e2 ne -> NE.ge ne ne' -> eagree e1 e2 ne'.
Proof.
- intros; red; intros. apply nge_agree with (NE.get r ne); auto. apply H0.
+ intros; red; intros. apply nge_agree with (NE.get r ne); auto. apply H0.
Qed.
Lemma eagree_bot:
@@ -1045,15 +1045,15 @@ Qed.
Lemma eagree_same:
forall e ne, eagree e e ne.
Proof.
- intros; red; intros. apply vagree_same.
+ intros; red; intros. apply vagree_same.
Qed.
Lemma eagree_update_1:
forall e1 e2 ne v1 v2 nv r,
eagree e1 e2 ne -> vagree v1 v2 nv -> eagree (e1#r <- v1) (e2#r <- v2) (NE.set r nv ne).
Proof.
- intros; red; intros. rewrite NE.gsspec. rewrite ! PMap.gsspec.
- destruct (peq r0 r); auto.
+ intros; red; intros. rewrite NE.gsspec. rewrite ! PMap.gsspec.
+ destruct (peq r0 r); auto.
Qed.
Lemma eagree_update:
@@ -1062,7 +1062,7 @@ Lemma eagree_update:
eagree e1 e2 (NE.set r Nothing ne) ->
eagree (e1#r <- v1) (e2#r <- v2) ne.
Proof.
- intros; red; intros. specialize (H0 r0). rewrite NE.gsspec in H0.
+ intros; red; intros. specialize (H0 r0). rewrite NE.gsspec in H0.
rewrite ! PMap.gsspec. destruct (peq r0 r).
subst r0. auto.
auto.
@@ -1073,8 +1073,8 @@ Lemma eagree_update_dead:
nreg ne r = Nothing ->
eagree e1 e2 ne -> eagree (e1#r <- v1) e2 ne.
Proof.
- intros; red; intros. rewrite PMap.gsspec.
- destruct (peq r0 r); auto. subst. unfold nreg in H. rewrite H. red; auto.
+ intros; red; intros. rewrite PMap.gsspec.
+ destruct (peq r0 r); auto. subst. unfold nreg in H. rewrite H. red; auto.
Qed.
(** * Neededness for memory locations *)
@@ -1146,12 +1146,12 @@ Lemma nlive_add:
Int.unsigned ofs <= i < Int.unsigned ofs + sz ->
nlive (nmem_add nm p sz) b i.
Proof.
- intros. unfold nmem_add. destruct nm. apply nlive_all.
- inv H1; try (apply nlive_all).
+ intros. unfold nmem_add. destruct nm. apply nlive_all.
+ inv H1; try (apply nlive_all).
- (* Gl id ofs *)
- assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
+ assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
destruct gl!id as [iv|] eqn:NG.
- + constructor; simpl; intros.
+ + constructor; simpl; intros.
congruence.
assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0.
rewrite PTree.gss in H5. inv H5. rewrite ISet.In_remove.
@@ -1161,13 +1161,13 @@ Proof.
assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0.
congruence.
- (* Glo id *)
- assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
+ assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
constructor; simpl; intros.
congruence.
- assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0.
+ assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0.
rewrite PTree.grs in H5. congruence.
- (* Stk ofs *)
- constructor; simpl; intros.
+ constructor; simpl; intros.
rewrite ISet.In_remove. intros [A B]. elim A; auto.
assert (bc b = BCglob id) by (eapply H; eauto). congruence.
- (* Stack *)
@@ -1183,19 +1183,19 @@ Proof.
intros. inversion H; subst. unfold nmem_add; destruct p; try (apply nlive_all).
- (* Gl id ofs *)
destruct gl!id as [iv|] eqn:NG.
- + split; simpl; intros. auto.
- rewrite PTree.gsspec in H1. destruct (peq id0 id); eauto. inv H1.
+ + split; simpl; intros. auto.
+ rewrite PTree.gsspec in H1. destruct (peq id0 id); eauto. inv H1.
rewrite ISet.In_remove. intros [P Q]. eelim GL; eauto.
- + auto.
+ + auto.
- (* Glo id *)
- split; simpl; intros. auto.
+ split; simpl; intros. auto.
rewrite PTree.grspec in H1. destruct (PTree.elt_eq id0 id). congruence. eauto.
- (* Stk ofs *)
- split; simpl; intros.
+ split; simpl; intros.
rewrite ISet.In_remove. intros [P Q]. eelim STK; eauto.
eauto.
- (* Stack *)
- split; simpl; intros.
+ split; simpl; intros.
apply ISet.In_empty.
eauto.
Qed.
@@ -1243,13 +1243,13 @@ Proof.
split; simpl; auto; intros.
rewrite PTree.gsspec in H6. destruct (peq id0 id).
+ inv H6. destruct H3. congruence. destruct gl!id as [iv0|] eqn:NG.
- rewrite ISet.In_add. intros [P|P]. omega. eelim GL; eauto.
- rewrite ISet.In_interval. omega.
-+ eauto.
+ rewrite ISet.In_add. intros [P|P]. omega. eelim GL; eauto.
+ rewrite ISet.In_interval. omega.
++ eauto.
- (* Stk ofs *)
- split; simpl; auto; intros. destruct H3.
- elim H3. subst b'. eapply bc_stack; eauto.
- rewrite ISet.In_add. intros [P|P]. omega. eapply STK; eauto.
+ split; simpl; auto; intros. destruct H3.
+ elim H3. subst b'. eapply bc_stack; eauto.
+ rewrite ISet.In_add. intros [P|P]. omega. eapply STK; eauto.
Qed.
(** Test (conservatively) whether some locations in the range delimited
@@ -1284,12 +1284,12 @@ Proof.
inv H1; try discriminate.
- (* Gl id ofs *)
assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto).
- destruct gl!id as [iv|] eqn:HG; inv H2.
+ destruct gl!id as [iv|] eqn:HG; inv H2.
destruct (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) iv) eqn:IC; try discriminate.
rewrite ISet.contains_spec in IC. eelim GL; eauto.
-- (* Stk ofs *)
+- (* Stk ofs *)
destruct (ISet.contains (Int.unsigned ofs) (Int.unsigned ofs + sz) stk) eqn:IC; try discriminate.
- rewrite ISet.contains_spec in IC. eelim STK; eauto. eapply bc_stack; eauto.
+ rewrite ISet.contains_spec in IC. eelim STK; eauto. eapply bc_stack; eauto.
Qed.
(** Kill all stack locations between 0 and [sz], and mark everything else
@@ -1303,7 +1303,7 @@ Lemma nlive_dead_stack:
forall sz b' i, b' <> sp \/ ~(0 <= i < sz) -> nlive (nmem_dead_stack sz) b' i.
Proof.
intros; constructor; simpl; intros.
-- rewrite ISet.In_interval. intuition.
+- rewrite ISet.In_interval. intuition.
- rewrite PTree.gempty in H1; discriminate.
Qed.
@@ -1330,10 +1330,10 @@ Proof.
intros. inversion H; subst. destruct nm2; simpl. auto.
constructor; simpl; intros.
- rewrite ISet.In_inter. intros [P Q]. eelim STK; eauto.
-- rewrite PTree.gcombine in H1 by auto.
+- rewrite PTree.gcombine in H1 by auto.
destruct gl!id as [iv1|] eqn:NG1; try discriminate;
destruct gl0!id as [iv2|] eqn:NG2; inv H1.
- rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto.
+ rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto.
Qed.
Lemma nlive_lub_r:
@@ -1342,10 +1342,10 @@ Proof.
intros. inversion H; subst. destruct nm1; simpl. auto.
constructor; simpl; intros.
- rewrite ISet.In_inter. intros [P Q]. eelim STK; eauto.
-- rewrite PTree.gcombine in H1 by auto.
+- rewrite PTree.gcombine in H1 by auto.
destruct gl0!id as [iv1|] eqn:NG1; try discriminate;
destruct gl!id as [iv2|] eqn:NG2; inv H1.
- rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto.
+ rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto.
Qed.
(** Boolean-valued equality test *)
@@ -1362,18 +1362,18 @@ Lemma nmem_beq_sound:
nmem_beq nm1 nm2 = true ->
(nlive nm1 b ofs <-> nlive nm2 b ofs).
Proof.
- unfold nmem_beq; intros.
+ unfold nmem_beq; intros.
destruct nm1 as [ | stk1 gl1]; destruct nm2 as [ | stk2 gl2]; try discriminate.
- split; intros L; inv L.
- InvBooleans. rewrite ISet.beq_spec in H0. rewrite PTree.beq_correct in H1.
split; intros L; inv L; constructor; intros.
-+ rewrite <- H0. eauto.
++ rewrite <- H0. eauto.
+ specialize (H1 id). rewrite H2 in H1. destruct gl1!id as [iv1|] eqn: NG; try contradiction.
- rewrite ISet.beq_spec in H1. rewrite <- H1. eauto.
-+ rewrite H0. eauto.
+ rewrite ISet.beq_spec in H1. rewrite <- H1. eauto.
++ rewrite H0. eauto.
+ specialize (H1 id). rewrite H2 in H1. destruct gl2!id as [iv2|] eqn: NG; try contradiction.
rewrite ISet.beq_spec in H1. rewrite H1. eauto.
-Qed.
+Qed.
End LOCATIONS.
@@ -1390,11 +1390,11 @@ Module NA <: SEMILATTICE.
Lemma eq_refl: forall x, eq x x.
Proof.
- unfold eq; destruct x; simpl; split. apply NE.eq_refl. tauto.
+ unfold eq; destruct x; simpl; split. apply NE.eq_refl. tauto.
Qed.
Lemma eq_sym: forall x y, eq x y -> eq y x.
Proof.
- unfold eq; destruct x, y; simpl. intros [A B].
+ unfold eq; destruct x, y; simpl. intros [A B].
split. apply NE.eq_sym; auto.
intros. rewrite B. tauto.
Qed.
@@ -1407,10 +1407,10 @@ Module NA <: SEMILATTICE.
Definition beq (x y: t) : bool :=
NE.beq (fst x) (fst y) && nmem_beq (snd x) (snd y).
-
+
Lemma beq_correct: forall x y, beq x y = true -> eq x y.
Proof.
- unfold beq, eq; destruct x, y; simpl; intros. InvBooleans. split.
+ unfold beq, eq; destruct x, y; simpl; intros. InvBooleans. split.
apply NE.beq_correct; auto.
intros. apply nmem_beq_sound; auto.
Qed.
@@ -1438,7 +1438,7 @@ Module NA <: SEMILATTICE.
Proof.
unfold ge, bot; destruct x; simpl. split.
apply NE.ge_bot.
- intros. inv H.
+ intros. inv H.
Qed.
Definition lub (x y: t) : t :=
@@ -1446,13 +1446,13 @@ Module NA <: SEMILATTICE.
Lemma ge_lub_left: forall x y, ge (lub x y) x.
Proof.
- unfold ge; destruct x, y; simpl; split.
+ unfold ge; destruct x, y; simpl; split.
apply NE.ge_lub_left.
intros; apply nlive_lub_l; auto.
Qed.
Lemma ge_lub_right: forall x y, ge (lub x y) y.
Proof.
- unfold ge; destruct x, y; simpl; split.
+ unfold ge; destruct x, y; simpl; split.
apply NE.ge_lub_right.
intros; apply nlive_lub_r; auto.
Qed.
diff --git a/backend/PrintAsm.ml b/backend/PrintAsm.ml
index 594b43b7..e7c945e3 100644
--- a/backend/PrintAsm.ml
+++ b/backend/PrintAsm.ml
@@ -38,7 +38,7 @@ module Printer(Target:TARGET) =
let print_location oc loc =
if loc <> Cutil.no_loc then Target.print_file_line oc (fst loc) (snd loc)
-
+
let print_function oc name fn =
Hashtbl.clear current_function_labels;
Target.reset_constants ();
@@ -66,7 +66,7 @@ module Printer(Target:TARGET) =
Target.print_jumptable oc jmptbl;
if !Clflags.option_g then
Hashtbl.iter (fun p i -> Debug.add_label name p i) current_function_labels
-
+
let print_init_data oc name id =
if Str.string_match PrintCsyntax.re_string_literal (extern_atom name) 0
&& List.for_all (function Init_int8 _ -> true | _ -> false) id
@@ -74,7 +74,7 @@ module Printer(Target:TARGET) =
fprintf oc " .ascii \"%s\"\n" (PrintCsyntax.string_of_init id)
else
List.iter (Target.print_init oc) id
-
+
let print_var oc name v =
match v.gvar_init with
| [] -> ()
@@ -101,7 +101,7 @@ module Printer(Target:TARGET) =
let sz =
match v.gvar_init with [Init_space sz] -> sz | _ -> assert false in
Target.print_comm_symb oc sz name align
-
+
let print_globdef oc (name,gdef) =
match gdef with
| Gfun (Internal code) -> print_function oc name code
@@ -116,7 +116,7 @@ module Printer(Target:TARGET) =
let symbol = Target.symbol
end
- module DebugPrinter = DwarfPrinter (DwarfTarget)
+ module DebugPrinter = DwarfPrinter (DwarfTarget)
end
let print_program oc p db =
diff --git a/backend/PrintAsmaux.ml b/backend/PrintAsmaux.ml
index 78399c04..4a612c26 100644
--- a/backend/PrintAsmaux.ml
+++ b/backend/PrintAsmaux.ml
@@ -98,7 +98,7 @@ let elf_symbol_offset oc (symb, ofs) =
let elf_print_fun_info oc name =
fprintf oc " .type %a, @function\n" elf_symbol name;
fprintf oc " .size %a, . - %a\n" elf_symbol name elf_symbol name
-
+
let elf_print_var_info oc name =
fprintf oc " .type %a, @object\n" elf_symbol name;
fprintf oc " .size %a, . - %a\n" elf_symbol name elf_symbol name
@@ -109,20 +109,20 @@ let cfi_startproc =
(fun oc -> fprintf oc " .cfi_startproc\n")
else
(fun _ -> ())
-
+
let cfi_endproc =
if Configuration.asm_supports_cfi then
(fun oc -> fprintf oc " .cfi_endproc\n")
else
(fun _ -> ())
-
-
+
+
let cfi_adjust =
if Configuration.asm_supports_cfi then
(fun oc delta -> fprintf oc " .cfi_adjust_cfa_offset %ld\n" delta)
else
(fun _ _ -> ())
-
+
let cfi_rel_offset =
if Configuration.asm_supports_cfi then
(fun oc reg ofs -> fprintf oc " .cfi_rel_offset %s, %ld\n" reg ofs)
@@ -211,7 +211,7 @@ let print_debug_info comment print_line print_preg sp_name oc kind txt args =
comment print_debug_args args;
| _ ->
()
-
+
(** Inline assembly *)
let print_asm_argument print_preg oc modifier = function
@@ -256,7 +256,7 @@ let print_inline_asm print_preg oc txt sg args res =
(** Print CompCert version and command-line as asm comment *)
let print_version_and_options oc comment =
- let version_string =
+ let version_string =
if Version.buildnr <> "" && Version.tag <> "" then
sprintf "%s, Build: %s, Tag: %s" Version.version Version.buildnr Version.tag
else
diff --git a/backend/PrintCminor.ml b/backend/PrintCminor.ml
index 19f4c839..9b6b1488 100644
--- a/backend/PrintCminor.ml
+++ b/backend/PrintCminor.ml
@@ -139,7 +139,7 @@ let rec expr p (prec, e) =
if assoc = LtoR
then (prec', prec' + 1)
else (prec' + 1, prec') in
- if prec' < prec
+ if prec' < prec
then fprintf p "@[<hov 2>("
else fprintf p "@[<hov 2>";
begin match e with
@@ -238,14 +238,14 @@ let rec print_stmt p s =
print_expr_list (true, el)
print_sig (ef_sig ef)
| Sbuiltin(Some id, ef, el) ->
- fprintf p "@[<hv 2>%s =@ builtin %s@,(@[<hov 0>%a@]) : @[<hov 0>%a@];@]"
+ fprintf p "@[<hv 2>%s =@ builtin %s@,(@[<hov 0>%a@]) : @[<hov 0>%a@];@]"
(ident_name id)
(name_of_external ef)
print_expr_list (true, el)
print_sig (ef_sig ef)
| Sseq(s1,s2) when just_skips s1 && just_skips s2 ->
()
- | Sseq(s1, s2) when just_skips s1 ->
+ | Sseq(s1, s2) when just_skips s1 ->
print_stmt p s2
| Sseq(s1, s2) when just_skips s2 ->
print_stmt p s1
@@ -277,7 +277,7 @@ let rec print_stmt p s =
(if long then "l" else "") print_expr e;
List.iter
(fun (n, x) ->
- fprintf p "@ case %s%s: exit %d;"
+ fprintf p "@ case %s%s: exit %d;"
(Z.to_string n)
(if long then "LL" else "")
(Nat.to_int x))
@@ -334,12 +334,12 @@ let print_init_data p = function
let rec print_init_data_list p = function
| [] -> ()
| [item] -> print_init_data p item
- | item::rest ->
+ | item::rest ->
(print_init_data p item;
fprintf p ",";
print_init_data_list p rest)
-let print_globvar p gv =
+let print_globvar p gv =
if (gv.gvar_readonly) then
fprintf p "readonly ";
if (gv.gvar_volatile) then
diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml
index 78ce1816..f2242c13 100644
--- a/backend/PrintRTL.ml
+++ b/backend/PrintRTL.ml
@@ -99,7 +99,7 @@ let print_function pp id f =
(List.rev_map
(fun (pc, i) -> (P.to_int pc, i))
(PTree.elements f.fn_code)) in
- print_succ pp f.fn_entrypoint
+ print_succ pp f.fn_entrypoint
(match instrs with (pc1, _) :: _ -> pc1 | [] -> -1);
List.iter (print_instruction pp) instrs;
fprintf pp "}\n\n"
diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml
index bb67dc96..dd8434da 100644
--- a/backend/PrintXTL.ml
+++ b/backend/PrintXTL.ml
@@ -71,7 +71,7 @@ let liveset pp lv =
fprintf pp "{";
VSet.iter (function V(r, ty) -> fprintf pp " x%d" (P.to_int r)
| L l -> ())
- lv;
+ lv;
fprintf pp " }"
let print_succ pp s dfl =
@@ -145,7 +145,7 @@ let print_function pp ?alloc ?live f =
(List.map
(fun (pc, i) -> (P.to_int pc, i))
(PTree.elements f.fn_code)) in
- print_succ pp f.fn_entrypoint
+ print_succ pp f.fn_entrypoint
(match instrs with (pc1, _) :: _ -> pc1 | [] -> -1);
List.iter (print_block pp) instrs;
fprintf pp "}\n\n";
diff --git a/backend/RTL.v b/backend/RTL.v
index 3cd4335d..a39d37cb 100644
--- a/backend/RTL.v
+++ b/backend/RTL.v
@@ -64,7 +64,7 @@ Inductive instruction: Type :=
| Icall: signature -> reg + ident -> list reg -> reg -> node -> instruction
(** [Icall sig fn args dest succ] invokes the function determined by
[fn] (either a function pointer found in a register or a
- function name), giving it the values of registers [args]
+ function name), giving it the values of registers [args]
as arguments. It stores the return value in [dest] and branches
to [succ]. *)
| Itailcall: signature -> reg + ident -> list reg -> instruction
@@ -127,7 +127,7 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
| _, _ => Regmap.init Vundef
end.
-(** The dynamic semantics of RTL is given in small-step style, as a
+(** The dynamic semantics of RTL is given in small-step style, as a
set of transitions between states. A state captures the current
point in the execution. Three kinds of states appear in the transitions:
@@ -149,7 +149,7 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
In all three kinds of states, the [cs] parameter represents the call stack.
It is a list of frames [Stackframe res f sp pc rs]. Each frame represents
-a function call in progress.
+a function call in progress.
[res] is the pseudo-register that will receive the result of the call.
[f] is the calling function.
[sp] is its stack pointer.
@@ -355,9 +355,9 @@ Proof.
assert (t1 = E0 -> exists s2, step (Genv.globalenv p) s t2 s2).
intros. subst. inv H0. exists s1; auto.
inversion H; subst; auto.
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (State s0 f sp pc' (regmap_setres res vres2 rs) m2). eapply exec_Ibuiltin; eauto.
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate s0 vres2 m2). econstructor; eauto.
(* trace length *)
red; intros; inv H; simpl; try omega.
@@ -450,15 +450,15 @@ Definition max_pc_function (f: function) :=
Lemma max_pc_function_sound:
forall f pc i, f.(fn_code)!pc = Some i -> Ple pc (max_pc_function f).
Proof.
- intros until i. unfold max_pc_function.
+ intros until i. unfold max_pc_function.
apply PTree_Properties.fold_rec with (P := fun c m => c!pc = Some i -> Ple pc m).
(* extensionality *)
- intros. apply H0. rewrite H; auto.
+ intros. apply H0. rewrite H; auto.
(* base case *)
rewrite PTree.gempty. congruence.
(* inductive case *)
- intros. rewrite PTree.gsspec in H2. destruct (peq pc k).
- inv H2. xomega.
+ intros. rewrite PTree.gsspec in H2. destruct (peq pc k).
+ inv H2. xomega.
apply Ple_trans with a. auto. xomega.
Qed.
@@ -493,9 +493,9 @@ Definition max_reg_function (f: function) :=
Remark max_reg_instr_ge:
forall m pc i, Ple m (max_reg_instr m pc i).
Proof.
- intros.
+ intros.
assert (X: forall l n, Ple m n -> Ple m (fold_left Pmax l n)).
- { induction l; simpl; intros.
+ { induction l; simpl; intros.
auto.
apply IHl. xomega. }
destruct i; simpl; try (destruct s0); repeat (apply X); try xomega.
@@ -518,9 +518,9 @@ Qed.
Remark max_reg_instr_uses:
forall m pc i r, In r (instr_uses i) -> Ple r (max_reg_instr m pc i).
Proof.
- intros.
+ intros.
assert (X: forall l n, In r l \/ Ple r n -> Ple r (fold_left Pmax l n)).
- { induction l; simpl; intros.
+ { induction l; simpl; intros.
tauto.
apply IHl. destruct H0 as [[A|A]|A]. right; subst; xomega. auto. right; xomega. }
destruct i; simpl in *; try (destruct s0); try (apply X; auto).
@@ -536,36 +536,36 @@ Lemma max_reg_function_def:
forall f pc i r,
f.(fn_code)!pc = Some i -> instr_defs i = Some r -> Ple r (max_reg_function f).
Proof.
- intros.
+ intros.
assert (Ple r (PTree.fold max_reg_instr f.(fn_code) 1%positive)).
- { revert H.
+ { revert H.
apply PTree_Properties.fold_rec with
(P := fun c m => c!pc = Some i -> Ple r m).
- intros. rewrite H in H1; auto.
- rewrite PTree.gempty; congruence.
- - intros. rewrite PTree.gsspec in H3. destruct (peq pc k).
- + inv H3. eapply max_reg_instr_def; eauto.
+ - intros. rewrite PTree.gsspec in H3. destruct (peq pc k).
+ + inv H3. eapply max_reg_instr_def; eauto.
+ apply Ple_trans with a. auto. apply max_reg_instr_ge.
}
- unfold max_reg_function. xomega.
+ unfold max_reg_function. xomega.
Qed.
Lemma max_reg_function_use:
forall f pc i r,
f.(fn_code)!pc = Some i -> In r (instr_uses i) -> Ple r (max_reg_function f).
Proof.
- intros.
+ intros.
assert (Ple r (PTree.fold max_reg_instr f.(fn_code) 1%positive)).
- { revert H.
+ { revert H.
apply PTree_Properties.fold_rec with
(P := fun c m => c!pc = Some i -> Ple r m).
- intros. rewrite H in H1; auto.
- rewrite PTree.gempty; congruence.
- - intros. rewrite PTree.gsspec in H3. destruct (peq pc k).
- + inv H3. eapply max_reg_instr_uses; eauto.
+ - intros. rewrite PTree.gsspec in H3. destruct (peq pc k).
+ + inv H3. eapply max_reg_instr_uses; eauto.
+ apply Ple_trans with a. auto. apply max_reg_instr_ge.
}
- unfold max_reg_function. xomega.
+ unfold max_reg_function. xomega.
Qed.
Lemma max_reg_function_params:
@@ -573,7 +573,7 @@ Lemma max_reg_function_params:
Proof.
intros.
assert (X: forall l n, In r l \/ Ple r n -> Ple r (fold_left Pmax l n)).
- { induction l; simpl; intros.
+ { induction l; simpl; intros.
tauto.
apply IHl. destruct H0 as [[A|A]|A]. right; subst; xomega. auto. right; xomega. }
assert (Y: Ple r (fold_left Pmax f.(fn_params) 1%positive)).
diff --git a/backend/RTLgen.v b/backend/RTLgen.v
index 3da961c6..49d79fb2 100644
--- a/backend/RTLgen.v
+++ b/backend/RTLgen.v
@@ -27,7 +27,7 @@ Open Local Scope string_scope.
(** * Translation environments and state *)
-(** The translation functions are parameterized by the following
+(** The translation functions are parameterized by the following
compile-time environment, which maps CminorSel local variables and
let-bound variables to RTL registers. The mapping for local variables
is computed from the CminorSel variable declarations at the beginning of
@@ -78,7 +78,7 @@ Lemma state_incr_trans:
forall s1 s2 s3, state_incr s1 s2 -> state_incr s2 s3 -> state_incr s1 s3.
Proof.
intros. inv H; inv H0. apply state_incr_intro.
- apply Ple_trans with (st_nextnode s2); assumption.
+ apply Ple_trans with (st_nextnode s2); assumption.
apply Ple_trans with (st_nextreg s2); assumption.
intros. generalize (H3 pc) (H5 pc). intuition congruence.
Qed.
@@ -93,7 +93,7 @@ Qed.
to modify the global state. These luxuries are not available in Coq,
however. Instead, we use a monadic encoding of the translation:
translation functions take the current global state as argument,
- and return either [Error msg] to denote an error,
+ and return either [Error msg] to denote an error,
or [OK r s incr] to denote success. [s] is the modified state, [r]
the result value of the translation function. and [incr] a proof
that the final state is in the [state_incr] relation with the
@@ -198,7 +198,7 @@ Definition add_instr (i: instruction) : mon node :=
fun s =>
let n := s.(st_nextnode) in
OK n
- (mkstate s.(st_nextreg) (Psucc n) (PTree.set n i s.(st_code))
+ (mkstate s.(st_nextreg) (Psucc n) (PTree.set n i s.(st_code))
(add_instr_wf s i))
(add_instr_incr s i).
@@ -306,7 +306,7 @@ Definition add_var (map: mapping) (name: ident) : mon (reg * mapping) :=
ret (r, mkmapping (PTree.set name r map.(map_vars))
map.(map_letvars)).
-Fixpoint add_vars (map: mapping) (names: list ident)
+Fixpoint add_vars (map: mapping) (names: list ident)
{struct names} : mon (list reg * mapping) :=
match names with
| nil => ret (nil, map)
@@ -339,7 +339,7 @@ Definition find_letvar (map: mapping) (idx: nat) : mon reg :=
fresh temporary register. Exception: if [a] is a let-bound variable
or a local variable, we return the RTL register associated
with that variable instead. Returning a fresh temporary in all cases
- would be semantically correct, but would generate less efficient
+ would be semantically correct, but would generate less efficient
RTL code. *)
Definition alloc_reg (map: mapping) (a: expr) : mon reg :=
diff --git a/backend/RTLgenaux.ml b/backend/RTLgenaux.ml
index e3373bf9..045299d4 100644
--- a/backend/RTLgenaux.ml
+++ b/backend/RTLgenaux.ml
@@ -72,7 +72,7 @@ let size_eos = function
let rec size_stmt = function
| Sskip -> 0
| Sassign(id, a) -> size_expr a
- | Sstore(chunk, addr, args, src) -> 1 + size_exprs args + size_expr src
+ | Sstore(chunk, addr, args, src) -> 1 + size_exprs args + size_expr src
| Scall(optid, sg, eos, args) ->
3 + size_eos eos + size_exprs args + length_exprs args
| Stailcall(sg, eos, args) ->
@@ -91,6 +91,6 @@ let rec size_stmt = function
| Slabel(lbl, s) -> size_stmt s
| Sgoto lbl -> 1
-let more_likely (c: condexpr) (ifso: stmt) (ifnot: stmt) =
+let more_likely (c: condexpr) (ifso: stmt) (ifnot: stmt) =
size_stmt ifso > size_stmt ifnot
diff --git a/backend/RTLgenproof.v b/backend/RTLgenproof.v
index 19f6f1f4..f458de8b 100644
--- a/backend/RTLgenproof.v
+++ b/backend/RTLgenproof.v
@@ -59,25 +59,25 @@ Qed.
Lemma add_var_wf:
forall s1 s2 map name r map' i,
- add_var map name s1 = OK (r,map') s2 i ->
+ add_var map name s1 = OK (r,map') s2 i ->
map_wf map -> map_valid map s1 -> map_wf map'.
Proof.
- intros. monadInv H.
+ intros. monadInv H.
apply mk_map_wf; simpl.
intros until r0. repeat rewrite PTree.gsspec.
destruct (peq id1 name); destruct (peq id2 name).
congruence.
- intros. inv H. elimtype False.
- apply valid_fresh_absurd with r0 s1.
+ intros. inv H. elimtype False.
+ apply valid_fresh_absurd with r0 s1.
apply H1. left; exists id2; auto.
eauto with rtlg.
- intros. inv H2. elimtype False.
- apply valid_fresh_absurd with r0 s1.
+ intros. inv H2. elimtype False.
+ apply valid_fresh_absurd with r0 s1.
apply H1. left; exists id1; auto.
eauto with rtlg.
inv H0. eauto.
intros until r0. rewrite PTree.gsspec.
- destruct (peq id name).
+ destruct (peq id name).
intros. inv H.
apply valid_fresh_absurd with r0 s1.
apply H1. right; auto.
@@ -90,7 +90,7 @@ Lemma add_vars_wf:
add_vars map names s1 = OK (rl,map') s2 i ->
map_wf map -> map_valid map s1 -> map_wf map'.
Proof.
- induction names; simpl; intros; monadInv H.
+ induction names; simpl; intros; monadInv H.
auto.
exploit add_vars_valid; eauto. intros [A B].
eapply add_var_wf; eauto.
@@ -174,7 +174,7 @@ Lemma match_env_update_temp:
match_env map e le (rs#r <- v).
Proof.
intros. apply match_env_invariant with rs; auto.
- intros. case (Reg.eq r r0); intro.
+ intros. case (Reg.eq r r0); intro.
subst r0; contradiction.
apply Regmap.gso; auto.
Qed.
@@ -200,7 +200,7 @@ Proof.
exists r'; split. auto. rewrite PMap.gso; auto.
red; intros. subst r'. elim n. eauto.
erewrite list_map_exten. eauto.
- intros. symmetry. apply PMap.gso. red; intros. subst x. eauto.
+ intros. symmetry. apply PMap.gso. red; intros. subst x. eauto.
Qed.
(** A variant of [match_env_update_var] where a variable is optionally
@@ -214,8 +214,8 @@ Lemma match_env_update_dest:
match_env map e le rs ->
match_env map (set_optvar dst v e) le (rs#r <- tv).
Proof.
- intros. inv H1; simpl.
- eapply match_env_update_temp; eauto.
+ intros. inv H1; simpl.
+ eapply match_env_update_temp; eauto.
eapply match_env_update_var; eauto.
Qed.
Hint Resolve match_env_update_dest: rtlg.
@@ -253,7 +253,7 @@ Lemma match_env_unbind_letvar:
match_env (add_letvar map r) e (v :: le) rs ->
match_env map e le rs.
Proof.
- unfold add_letvar; intros. inv H. simpl in *.
+ unfold add_letvar; intros. inv H. simpl in *.
constructor. auto. inversion me_letvars0. auto.
Qed.
@@ -282,13 +282,13 @@ Lemma match_set_params_init_regs:
Proof.
induction il; intros.
- inv H. split. apply match_env_empty. auto. intros.
+ inv H. split. apply match_env_empty. auto. intros.
simpl. apply Regmap.gi.
monadInv H. simpl.
exploit add_vars_valid; eauto. apply init_mapping_valid. intros [A B].
exploit add_var_valid; eauto. intros [A' B']. clear B'.
- monadInv EQ1.
+ monadInv EQ1.
destruct H0 as [ | v1 tv1 vs tvs].
(* vl = nil *)
destruct (IHil _ _ _ _ nil nil _ EQ) as [ME UNDEF]. constructor. inv ME. split.
@@ -306,13 +306,13 @@ Proof.
intros id v. repeat rewrite PTree.gsspec. destruct (peq id a); intros.
subst a. inv H. inv H1. exists x1; split. auto. rewrite Regmap.gss. constructor.
inv H1. eexists; eauto.
- exploit me_vars0; eauto. intros [r' [C D]].
+ exploit me_vars0; eauto. intros [r' [C D]].
exists r'; split. auto. rewrite Regmap.gso. auto.
apply valid_fresh_different with s.
apply B. left; exists id; auto.
- eauto with rtlg.
+ eauto with rtlg.
destruct (map_letvars x0). auto. simpl in me_letvars0. inversion me_letvars0.
- intros. rewrite Regmap.gso. apply UNDEF.
+ intros. rewrite Regmap.gso. apply UNDEF.
apply reg_fresh_decr with s2; eauto with rtlg.
apply sym_not_equal. apply valid_fresh_different with s2; auto.
Qed.
@@ -330,15 +330,15 @@ Proof.
inv H2. auto.
- monadInv H2.
- exploit IHil; eauto. intro.
+ monadInv H2.
+ exploit IHil; eauto. intro.
monadInv EQ1.
constructor.
- intros id v. simpl. repeat rewrite PTree.gsspec.
- destruct (peq id a). subst a. intro.
+ intros id v. simpl. repeat rewrite PTree.gsspec.
+ destruct (peq id a). subst a. intro.
exists x1. split. auto. inv H3. constructor.
eauto with rtlg.
- intros. eapply me_vars; eauto.
+ intros. eapply me_vars; eauto.
simpl. eapply me_letvars; eauto.
Qed.
@@ -406,7 +406,7 @@ Lemma sig_transl_function:
Proof.
intros until tf. unfold transl_fundef, transf_partial_fundef.
case f; intro.
- unfold transl_function.
+ unfold transl_function.
destruct (reserve_labels (fn_body f0) (PTree.empty node, init_state)) as [ngoto s0].
case (transl_fun f0 ngoto s0); simpl; intros.
discriminate.
@@ -429,10 +429,10 @@ Lemma tr_move_correct:
rs'#r2 = rs#r1 /\
(forall r, r <> r2 -> rs'#r = rs#r).
Proof.
- intros. inv H.
+ intros. inv H.
exists rs; split. constructor. auto.
- exists (rs#r2 <- (rs#r1)); split.
- apply star_one. eapply exec_Iop. eauto. auto.
+ exists (rs#r2 <- (rs#r1)); split.
+ apply star_one. eapply exec_Iop. eauto. auto.
split. apply Regmap.gss. intros; apply Regmap.gso; auto.
Qed.
@@ -475,7 +475,7 @@ Variable m: mem.
We formalize this simulation property by the following predicate
parameterized by the CminorSel evaluation (left arrow). *)
-Definition transl_expr_prop
+Definition transl_expr_prop
(le: letenv) (a: expr) (v: val) : Prop :=
forall tm cs f map pr ns nd rd rs dst
(MWF: map_wf map)
@@ -489,7 +489,7 @@ Definition transl_expr_prop
/\ (forall r, In r pr -> rs'#r = rs#r)
/\ Mem.extends m tm'.
-Definition transl_exprlist_prop
+Definition transl_exprlist_prop
(le: letenv) (al: exprlist) (vl: list val) : Prop :=
forall tm cs f map pr ns nd rl rs
(MWF: map_wf map)
@@ -503,7 +503,7 @@ Definition transl_exprlist_prop
/\ (forall r, In r pr -> rs'#r = rs#r)
/\ Mem.extends m tm'.
-Definition transl_condexpr_prop
+Definition transl_condexpr_prop
(le: letenv) (a: condexpr) (v: bool) : Prop :=
forall tm cs f map pr ns ntrue nfalse rs
(MWF: map_wf map)
@@ -531,22 +531,22 @@ Lemma transl_expr_Evar_correct:
Proof.
intros; red; intros. inv TE.
exploit match_env_find_var; eauto. intro EQ.
- exploit tr_move_correct; eauto. intros [rs' [A [B C]]].
+ exploit tr_move_correct; eauto. intros [rs' [A [B C]]].
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.
+ intros. destruct (Reg.eq r rd). subst r. auto. auto.
split. eapply match_env_invariant; eauto.
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.
+ apply match_env_update_dest; auto.
+ intros. rewrite Regmap.gsspec. destruct (peq r0 rd). congruence. auto.
+ split. congruence.
split. intros. apply C. intuition congruence.
auto.
Qed.
@@ -560,7 +560,7 @@ Lemma transl_expr_Eop_correct:
transl_expr_prop le (Eop op args) v.
Proof.
intros; red; intros. inv TE.
-(* normal case *)
+(* normal case *)
exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RR1 [RO1 EXT1]]]]]].
edestruct eval_operation_lessdef as [v' []]; eauto.
exists (rs1#rd <- v'); exists tm1.
@@ -599,13 +599,13 @@ Proof.
apply eval_addressing_preserved. exact symbols_preserved.
auto. traceEq.
(* Match-env *)
- split. eauto with rtlg.
+ split. eauto with rtlg.
(* Result *)
split. rewrite Regmap.gss. auto.
(* Other regs *)
split. intros. rewrite Regmap.gso. auto. intuition congruence.
(* Mem *)
- auto.
+ auto.
Qed.
Lemma transl_expr_Econdition_correct:
@@ -618,7 +618,7 @@ Lemma transl_expr_Econdition_correct:
transl_expr_prop le (Econdition a ifso ifnot) v.
Proof.
intros; red; intros; inv TE.
- exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [OTHER1 EXT1]]]]].
+ exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [OTHER1 EXT1]]]]].
assert (tr_expr f.(fn_code) map pr (if va then ifso else ifnot) (if va then ntrue else nfalse) nd rd dst).
destruct va; auto.
exploit H2; eauto. intros [rs2 [tm2 [EX2 [ME2 [RES2 [OTHER2 EXT2]]]]]].
@@ -643,10 +643,10 @@ Lemma transl_expr_Elet_correct:
transl_expr_prop (v1 :: le) a2 v2 ->
transl_expr_prop le (Elet a1 a2) v2.
Proof.
- intros; red; intros; inv TE.
+ intros; red; intros; inv TE.
exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RES1 [OTHER1 EXT1]]]]]].
assert (map_wf (add_letvar map r)).
- eapply add_letvar_wf; eauto.
+ eapply add_letvar_wf; eauto.
exploit H2; eauto. eapply match_env_bind_letvar; eauto.
intros [rs2 [tm2 [EX2 [ME3 [RES2 [OTHER2 EXT2]]]]]].
exists rs2; exists tm2.
@@ -673,9 +673,9 @@ Proof.
(* Exec *)
split. eexact EX1.
(* Match-env *)
- split.
+ split.
destruct H2 as [[A B] | [A B]].
- subst r dst; simpl.
+ subst r dst; simpl.
apply match_env_invariant with rs. auto.
intros. destruct (Reg.eq r rd). subst r. auto. auto.
apply match_env_invariant with (rs#rd <- (rs#r)).
@@ -684,9 +684,9 @@ Proof.
intros. rewrite Regmap.gsspec. destruct (peq r0 rd); auto.
congruence.
(* Result *)
- split. rewrite RES1. eapply match_env_find_letvar; eauto.
+ split. rewrite RES1. eapply match_env_find_letvar; eauto.
(* Other regs *)
- split. intros.
+ split. intros.
destruct H2 as [[A B] | [A B]].
destruct (Reg.eq r0 rd); subst; auto.
apply OTHER1. intuition congruence.
@@ -712,7 +712,7 @@ Lemma transl_expr_Ebuiltin_correct:
Proof.
intros; red; intros. inv TE.
exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RR1 [RO1 EXT1]]]]]].
- exploit external_call_mem_extends; eauto.
+ exploit external_call_mem_extends; eauto.
intros [v' [tm2 [A [B [C [D E]]]]]].
exists (rs1#rd <- v'); exists tm2.
(* Exec *)
@@ -720,7 +720,7 @@ Proof.
change (rs1#rd <- v') with (regmap_setres (BR rd) v' rs1).
eapply exec_Ibuiltin; eauto.
eapply eval_builtin_args_trivial.
- eapply external_call_symbols_preserved; eauto.
+ eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
reflexivity.
(* Match-env *)
@@ -745,9 +745,9 @@ Lemma transl_expr_Eexternal_correct:
Proof.
intros; red; intros. inv TE.
exploit H3; eauto. intros [rs1 [tm1 [EX1 [ME1 [RR1 [RO1 EXT1]]]]]].
- exploit external_call_mem_extends; eauto.
+ 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.
+ 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.
@@ -756,7 +756,7 @@ Proof.
eapply star_left. eapply exec_function_external.
eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- apply star_one. apply exec_return.
+ apply star_one. apply exec_return.
reflexivity. reflexivity. reflexivity.
(* Match-env *)
split. eauto with rtlg.
@@ -794,7 +794,7 @@ Proof.
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.
+ split. eapply star_trans. eexact EX1. eexact EX2. auto.
(* Match-env *)
split. assumption.
(* Results *)
@@ -803,7 +803,7 @@ Proof.
auto.
(* Other regs *)
split. intros. transitivity (rs1#r).
- apply OTHER2; auto. simpl; tauto.
+ apply OTHER2; auto. simpl; tauto.
apply OTHER1; auto.
(* Mem *)
auto.
@@ -816,16 +816,16 @@ Lemma transl_condexpr_CEcond_correct:
eval_condition cond vl m = Some vb ->
transl_condexpr_prop le (CEcond cond al) vb.
Proof.
- intros; red; intros. inv TE.
+ intros; red; intros. inv TE.
exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [RES1 [OTHER1 EXT1]]]]]].
exists rs1; exists tm1.
(* Exec *)
- split. eapply plus_right. eexact EX1. eapply exec_Icond. eauto.
+ split. eapply plus_right. eexact EX1. eapply exec_Icond. eauto.
eapply eval_condition_lessdef; eauto. auto. traceEq.
(* Match-env *)
split. assumption.
(* Other regs *)
- split. assumption.
+ split. assumption.
(* Mem *)
auto.
Qed.
@@ -838,7 +838,7 @@ Lemma transl_condexpr_CEcondition_correct:
transl_condexpr_prop le (if va then b else c) v ->
transl_condexpr_prop le (CEcondition a b c) v.
Proof.
- intros; red; intros. inv TE.
+ intros; red; intros. inv TE.
exploit H0; eauto. intros [rs1 [tm1 [EX1 [ME1 [OTHER1 EXT1]]]]].
assert (tr_condition (fn_code f) map pr (if va then b else c) (if va then n2 else n3) ntrue nfalse).
destruct va; auto.
@@ -849,7 +849,7 @@ Proof.
(* Match-env *)
split. assumption.
(* Other regs *)
- split. intros. rewrite OTHER2; auto.
+ split. intros. rewrite OTHER2; auto.
(* Mem *)
auto.
Qed.
@@ -862,11 +862,11 @@ Lemma transl_condexpr_CElet_correct:
transl_condexpr_prop (v1 :: le) b v2 ->
transl_condexpr_prop le (CElet a b) v2.
Proof.
- intros; red; intros. inv TE.
+ intros; red; intros. inv TE.
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.
+ eapply add_letvar_wf; eauto.
+ exploit H2; eauto. eapply match_env_bind_letvar; eauto.
intros [rs2 [tm2 [EX2 [ME3 [OTHER2 EXT2]]]]].
exists rs2; exists tm2.
(* Exec *)
@@ -874,7 +874,7 @@ Proof.
(* Match-env *)
split. eapply match_env_unbind_letvar; eauto.
(* Other regs *)
- split. intros. rewrite OTHER2; auto.
+ split. intros. rewrite OTHER2; auto.
(* Mem *)
auto.
Qed.
@@ -950,7 +950,7 @@ Proof
(** Exit expressions. *)
-Definition transl_exitexpr_prop
+Definition transl_exitexpr_prop
(le: letenv) (a: exitexpr) (x: nat) : Prop :=
forall tm cs f map ns nexits rs
(MWF: map_wf map)
@@ -981,21 +981,21 @@ Proof.
auto.
- (* XEcondition *)
exploit transl_condexpr_correct; eauto. intros (rs1 & tm1 & EXEC1 & ME1 & RES1 & EXT1).
- exploit IHeval_exitexpr; eauto.
+ exploit IHeval_exitexpr; eauto.
instantiate (2 := if va then n2 else n3). destruct va; eauto.
intros (nd & rs2 & tm2 & EXEC2 & EXIT2 & ME2 & EXT2).
- exists nd, rs2, tm2.
+ exists nd, rs2, tm2.
split. eapply star_trans. apply plus_star. eexact EXEC1. eexact EXEC2. traceEq.
auto.
- (* XElet *)
exploit transl_expr_correct; eauto. intros (rs1 & tm1 & EXEC1 & ME1 & RES1 & PRES1 & EXT1).
assert (map_wf (add_letvar map r)).
- eapply add_letvar_wf; eauto.
+ eapply add_letvar_wf; eauto.
exploit IHeval_exitexpr; eauto. eapply match_env_bind_letvar; eauto.
intros (nd & rs2 & tm2 & EXEC2 & EXIT2 & ME2 & EXT2).
exists nd, rs2, tm2.
- split. eapply star_trans. eexact EXEC1. eexact EXEC2. traceEq.
- split. auto.
+ split. eapply star_trans. eexact EXEC1. eexact EXEC2. traceEq.
+ split. auto.
split. eapply match_env_unbind_letvar; eauto.
auto.
Qed.
@@ -1010,20 +1010,20 @@ Lemma eval_exprlist_append:
Proof.
induction al1; simpl; intros vl1 al2 vl2 E1 E2; inv E1.
- auto.
-- simpl. constructor; eauto.
+- simpl. constructor; eauto.
Qed.
Lemma invert_eval_builtin_arg:
forall a v,
eval_builtin_arg ge sp e m a v ->
- exists vl,
+ exists vl,
eval_exprlist ge sp e m nil (exprlist_of_expr_list (params_of_builtin_arg a)) vl
/\ Events.eval_builtin_arg ge (fun v => v) sp m (fst (convert_builtin_arg a vl)) v
/\ (forall vl', convert_builtin_arg a (vl ++ vl') = (fst (convert_builtin_arg a vl), vl')).
Proof.
induction 1; simpl; econstructor; intuition eauto with evalexpr barg.
- constructor.
- constructor.
+ constructor.
+ constructor.
repeat constructor.
Qed.
@@ -1040,7 +1040,7 @@ Proof.
destruct IHlist_forall2 as (vl2 & D & E).
exists (vl1 ++ vl2); split.
apply eval_exprlist_append; auto.
- rewrite C; simpl. constructor; auto.
+ rewrite C; simpl. constructor; auto.
Qed.
Lemma transl_eval_builtin_arg:
@@ -1055,18 +1055,18 @@ Proof.
induction a; simpl; intros until v; intros LD EV;
try (now (inv EV; econstructor; eauto with barg)).
- destruct rl; simpl in LD; inv LD; inv EV; simpl.
- econstructor; eauto with barg.
+ econstructor; eauto with barg.
exists (rs#p); intuition auto. constructor.
- destruct (convert_builtin_arg a1 vl) as [a1' vl1] eqn:CV1; simpl in *.
destruct (convert_builtin_arg a2 vl1) as [a2' vl2] eqn:CV2; simpl in *.
destruct (convert_builtin_arg a1 rl) as [a1'' rl1] eqn:CV3; simpl in *.
destruct (convert_builtin_arg a2 rl1) as [a2'' rl2] eqn:CV4; simpl in *.
- inv EV.
- exploit IHa1; eauto. rewrite CV1; simpl; eauto.
+ inv EV.
+ exploit IHa1; eauto. rewrite CV1; simpl; eauto.
rewrite CV1, CV3; simpl. intros (v1' & A1 & B1 & C1).
exploit IHa2. eexact C1. rewrite CV2; simpl; eauto.
rewrite CV2, CV4; simpl. intros (v2' & A2 & B2 & C2).
- exists (Val.longofwords v1' v2'); split. constructor; auto.
+ exists (Val.longofwords v1' v2'); split. constructor; auto.
split; auto. apply Val.longofwords_lessdef; auto.
Qed.
@@ -1081,8 +1081,8 @@ Proof.
induction al; simpl; intros until vl; intros LD EV.
- inv EV. exists (@nil val); split; constructor.
- destruct (convert_builtin_arg a vl1) as [a1' vl2] eqn:CV1; simpl in *.
- inv EV.
- exploit transl_eval_builtin_arg. eauto. instantiate (2 := a). rewrite CV1; simpl; eauto.
+ inv EV.
+ exploit transl_eval_builtin_arg. eauto. instantiate (2 := a). rewrite CV1; simpl; eauto.
rewrite CV1; simpl. intros (v1' & A1 & B1 & C1).
exploit IHal. eexact C1. eauto. intros (vl' & A2 & B2).
destruct (convert_builtin_arg a rl) as [a1'' rl2]; simpl in *.
@@ -1145,10 +1145,10 @@ Lemma lt_state_wf:
well_founded lt_state.
Proof.
unfold lt_state. apply wf_inverse_image with (f := measure_state).
- apply wf_lex_ord. apply lt_wf. apply lt_wf.
+ apply wf_lex_ord. apply lt_wf. apply lt_wf.
Qed.
-(** ** Semantic preservation for the translation of statements *)
+(** ** Semantic preservation for the translation of statements *)
(** The simulation diagram for the translation of statements
and functions is a "star" diagram of the form:
@@ -1180,7 +1180,7 @@ Inductive tr_fun (tf: function) (map: mapping) (f: CminorSel.function)
tf.(fn_stacksize) = f.(fn_stackspace) ->
tr_fun tf map f ngoto nret rret.
-Inductive tr_cont: RTL.code -> mapping ->
+Inductive tr_cont: RTL.code -> mapping ->
CminorSel.cont -> node -> list node -> labelmap -> node -> option reg ->
list RTL.stackframe -> Prop :=
| tr_Kseq: forall c map s k nd nexits ngoto nret rret cs n,
@@ -1269,7 +1269,7 @@ Proof.
(* seq *)
caseEq (find_label lbl s1 (Kseq s2 k)); intros.
inv H1. inv H2. eapply IHs1; eauto. econstructor; eauto.
- inv H2. eapply IHs2; eauto.
+ inv H2. eapply IHs2; eauto.
(* ifthenelse *)
caseEq (find_label lbl s1 k); intros.
inv H1. inv H2. eapply IHs1; eauto.
@@ -1308,22 +1308,22 @@ Proof.
econstructor; eauto. constructor.
(* skip return *)
- inv TS.
+ inv TS.
assert ((fn_code tf)!ncont = Some(Ireturn rret)
/\ match_stacks k cs).
- inv TK; simpl in H; try contradiction; auto.
+ inv TK; simpl in H; try contradiction; auto.
destruct H1.
assert (fn_stacksize tf = fn_stackspace f).
- inv TF. auto.
+ inv TF. auto.
edestruct Mem.free_parallel_extends as [tm' []]; eauto.
econstructor; split.
left; apply plus_one. eapply exec_Ireturn. eauto.
rewrite H3. eauto.
constructor; auto.
-
+
(* assign *)
inv TS.
- exploit transl_expr_correct; eauto.
+ exploit transl_expr_correct; eauto.
intros [rs' [tm' [A [B [C [D E]]]]]].
econstructor; split.
right; split. eauto. Lt_state.
@@ -1367,8 +1367,8 @@ Proof.
exploit functions_translated; eauto. intros [tf' [P Q]].
econstructor; split.
left; eapply plus_right. eexact E.
- eapply exec_Icall; eauto. simpl. rewrite symbols_preserved. rewrite H4.
- rewrite Genv.find_funct_find_funct_ptr in P. eauto.
+ eapply exec_Icall; eauto. simpl. rewrite symbols_preserved. rewrite H4.
+ rewrite Genv.find_funct_find_funct_ptr in P. eauto.
apply sig_transl_function; auto.
traceEq.
constructor; auto. econstructor; eauto.
@@ -1400,19 +1400,19 @@ Proof.
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.
- rewrite Genv.find_funct_find_funct_ptr in P. eauto.
+ eapply exec_Itailcall; eauto. simpl. rewrite symbols_preserved. rewrite H5.
+ rewrite Genv.find_funct_find_funct_ptr in P. eauto.
apply sig_transl_function; auto.
rewrite H; eauto.
traceEq.
constructor; auto.
(* builtin *)
- inv TS.
+ inv TS.
exploit invert_eval_builtin_args; eauto. intros (vparams & P & Q).
exploit transl_exprlist_correct; eauto.
intros [rs' [tm' [E [F [G [J K]]]]]].
- exploit transl_eval_builtin_args; eauto.
+ exploit transl_eval_builtin_args; eauto.
intros (vargs' & U & V).
exploit (@eval_builtin_args_lessdef _ ge (fun r => rs'#r) (fun r => rs'#r)); eauto.
intros (vargs'' & X & Y).
@@ -1421,31 +1421,31 @@ Proof.
econstructor; split.
left. eapply plus_right. eexact E.
eapply exec_Ibuiltin. eauto.
- eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
+ eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
eapply external_call_symbols_preserved. eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- traceEq.
+ traceEq.
econstructor; eauto. constructor.
eapply match_env_update_res; eauto.
-
+
(* seq *)
- inv TS.
+ inv TS.
econstructor; split.
right; split. apply star_refl. Lt_state.
- econstructor; eauto. econstructor; eauto.
+ econstructor; eauto. econstructor; eauto.
(* ifthenelse *)
- inv TS.
+ inv TS.
exploit transl_condexpr_correct; eauto. intros [rs' [tm' [A [B [C D]]]]].
econstructor; split.
left. eexact A.
- destruct b; econstructor; eauto.
+ destruct b; econstructor; eauto.
(* loop *)
inversion TS; subst.
econstructor; split.
- left. apply plus_one. eapply exec_Inop; eauto.
- econstructor; eauto.
+ left. apply plus_one. eapply exec_Inop; eauto.
+ econstructor; eauto.
econstructor; eauto.
econstructor; eauto.
@@ -1456,7 +1456,7 @@ Proof.
econstructor; eauto. econstructor; eauto.
(* exit seq *)
- inv TS. inv TK.
+ inv TS. inv TK.
econstructor; split.
right; split. apply star_refl. Lt_state.
econstructor; eauto. econstructor; eauto.
@@ -1475,11 +1475,11 @@ Proof.
(* switch *)
inv TS.
- exploit transl_exitexpr_correct; eauto.
- intros (nd & rs' & tm' & A & B & C & D).
+ exploit transl_exitexpr_correct; eauto.
+ intros (nd & rs' & tm' & A & B & C & D).
econstructor; split.
- right; split. eexact A. Lt_state.
- econstructor; eauto. constructor; auto.
+ right; split. eexact A. Lt_state.
+ econstructor; eauto. constructor; auto.
(* return none *)
inv TS.
@@ -1511,11 +1511,11 @@ Proof.
(* goto *)
inv TS. inversion TF; subst.
- exploit tr_find_label; eauto. eapply tr_cont_call_cont; eauto.
+ exploit tr_find_label; eauto. eapply tr_cont_call_cont; eauto.
intros [ns2 [nd2 [nexits2 [A [B C]]]]].
econstructor; split.
left; apply plus_one. eapply exec_Inop; eauto.
- econstructor; eauto.
+ econstructor; eauto.
(* internal call *)
monadInv TF. exploit transl_function_charact; eauto. intro TRF.
@@ -1536,19 +1536,19 @@ Proof.
inversion MS; subst; econstructor; eauto.
(* external call *)
- monadInv TF.
+ monadInv TF.
edestruct external_call_mem_extends as [tvres [tm' [A [B [C D]]]]]; eauto.
econstructor; split.
- left; apply plus_one. eapply exec_function_external; eauto.
+ left; apply plus_one. eapply exec_function_external; eauto.
eapply external_call_symbols_preserved. eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
constructor; auto.
(* return *)
inv MS.
- econstructor; split.
- left; apply plus_one; constructor.
- econstructor; eauto. constructor.
+ econstructor; split.
+ left; apply plus_one; constructor.
+ econstructor; eauto. constructor.
eapply match_env_update_dest; eauto.
Qed.
@@ -1582,8 +1582,8 @@ Proof.
eexact public_preserved.
eexact transl_initial_states.
eexact transl_final_states.
- apply lt_state_wf.
- exact transl_step_correct.
+ apply lt_state_wf.
+ exact transl_step_correct.
Qed.
End CORRECTNESS.
diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v
index 1e665002..17022a7d 100644
--- a/backend/RTLgenspec.v
+++ b/backend/RTLgenspec.v
@@ -12,7 +12,7 @@
(** Abstract specification of RTL generation *)
-(** In this module, we define inductive predicates that specify the
+(** In this module, we define inductive predicates that specify the
translations from Cminor to RTL performed by the functions in module
[RTLgen]. We then show that these functions satisfy these relational
specifications. The relational specifications will then be used
@@ -43,7 +43,7 @@ Require Import RTLgen.
*)
Remark bind_inversion:
- forall (A B: Type) (f: mon A) (g: A -> mon B)
+ forall (A B: Type) (f: mon A) (g: A -> mon B)
(y: B) (s1 s3: state) (i: state_incr s1 s3),
bind f g s1 = OK y s3 i ->
exists x, exists s2, exists i1, exists i2,
@@ -64,7 +64,7 @@ Remark bind2_inversion:
f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2.
Proof.
unfold bind2; intros.
- exploit bind_inversion; eauto.
+ exploit bind_inversion; eauto.
intros [[x y] [s2 [i1 [i2 [P Q]]]]]. simpl in Q.
exists x; exists y; exists s2; exists i1; exists i2; auto.
Qed.
@@ -108,7 +108,7 @@ Ltac monadInv H :=
| (error _ _ = OK _ _ _) => monadInv1 H
| (bind ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H
| (bind2 ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H
- | (?F _ _ _ _ _ _ _ _ = OK _ _ _) =>
+ | (?F _ _ _ _ _ _ _ _ = OK _ _ _) =>
((progress simpl in H) || unfold F in H); monadInv1 H
| (?F _ _ _ _ _ _ _ = OK _ _ _) =>
((progress simpl in H) || unfold F in H); monadInv1 H
@@ -163,7 +163,7 @@ Ltac saturateTrans :=
earlier, that is, it is less than the next fresh register of the state.
Otherwise, the pseudo-register is said to be fresh. *)
-Definition reg_valid (r: reg) (s: state) : Prop :=
+Definition reg_valid (r: reg) (s: state) : Prop :=
Plt r s.(st_nextreg).
Definition reg_fresh (r: reg) (s: state) : Prop :=
@@ -183,7 +183,7 @@ Proof.
Qed.
Hint Resolve valid_fresh_different: rtlg.
-Lemma reg_valid_incr:
+Lemma reg_valid_incr:
forall r s1 s2, state_incr s1 s2 -> reg_valid r s1 -> reg_valid r s2.
Proof.
intros r s1 s2 INCR.
@@ -198,7 +198,7 @@ Proof.
intros r s1 s2 INCR. inversion INCR.
unfold reg_fresh; unfold not; intros.
apply H4. apply Plt_Ple_trans with (st_nextreg s1); auto.
-Qed.
+Qed.
Hint Resolve reg_fresh_decr: rtlg.
(** Validity of a list of registers. *)
@@ -206,7 +206,7 @@ Hint Resolve reg_fresh_decr: rtlg.
Definition regs_valid (rl: list reg) (s: state) : Prop :=
forall r, In r rl -> reg_valid r s.
-Lemma regs_valid_nil:
+Lemma regs_valid_nil:
forall s, regs_valid nil s.
Proof.
intros; red; intros. elim H.
@@ -224,7 +224,7 @@ Lemma regs_valid_app:
forall rl1 rl2 s,
regs_valid rl1 s -> regs_valid rl2 s -> regs_valid (rl1 ++ rl2) s.
Proof.
- intros; red; intros. apply in_app_iff in H1. destruct H1; auto.
+ intros; red; intros. apply in_app_iff in H1. destruct H1; auto.
Qed.
Lemma regs_valid_incr:
@@ -273,7 +273,7 @@ Lemma update_instr_at:
forall n i s1 s2 incr u,
update_instr n i s1 = OK u s2 incr -> s2.(st_code)!n = Some i.
Proof.
- intros. unfold update_instr in H.
+ intros. unfold update_instr in H.
destruct (plt n (st_nextnode s1)); try discriminate.
destruct (check_empty_node s1 n); try discriminate.
inv H. simpl. apply PTree.gss.
@@ -306,7 +306,7 @@ Lemma new_reg_not_in_map:
new_reg s1 = OK r s2 i -> map_valid m s1 -> ~(reg_in_map m r).
Proof.
unfold not; intros; eauto with rtlg.
-Qed.
+Qed.
Hint Resolve new_reg_not_in_map: rtlg.
(** * Properties of operations over compilation environments *)
@@ -315,10 +315,10 @@ Lemma init_mapping_valid:
forall s, map_valid init_mapping s.
Proof.
unfold map_valid, init_mapping.
- intros s r [[id A] | B].
+ intros s r [[id A] | B].
simpl in A. rewrite PTree.gempty in A; discriminate.
simpl in B. tauto.
-Qed.
+Qed.
(** Properties of [find_var]. *)
@@ -363,17 +363,17 @@ Hint Resolve find_letvar_valid: rtlg.
(** Properties of [add_var]. *)
Lemma add_var_valid:
- forall s1 s2 map1 map2 name r i,
+ forall s1 s2 map1 map2 name r i,
add_var map1 name s1 = OK (r, map2) s2 i ->
map_valid map1 s1 ->
reg_valid r s2 /\ map_valid map2 s2.
Proof.
- intros. monadInv H.
+ intros. monadInv H.
split. eauto with rtlg.
inversion EQ. subst. red. intros r' [[id A] | B].
simpl in A. rewrite PTree.gsspec in A. destruct (peq id name).
inv A. eauto with rtlg.
- apply reg_valid_incr with s1. eauto with rtlg.
+ apply reg_valid_incr with s1. eauto with rtlg.
apply H0. left; exists id; auto.
simpl in B. apply reg_valid_incr with s1. eauto with rtlg.
apply H0. right; auto.
@@ -383,11 +383,11 @@ Lemma add_var_find:
forall s1 s2 map name r map' i,
add_var map name s1 = OK (r,map') s2 i -> map'.(map_vars)!name = Some r.
Proof.
- intros. monadInv H. simpl. apply PTree.gss.
+ intros. monadInv H. simpl. apply PTree.gss.
Qed.
Lemma add_vars_valid:
- forall namel s1 s2 map1 map2 rl i,
+ forall namel s1 s2 map1 map2 rl i,
add_vars map1 namel s1 = OK (rl, map2) s2 i ->
map_valid map1 s1 ->
regs_valid rl s2 /\ map_valid map2 s2.
@@ -428,10 +428,10 @@ Lemma add_letvar_valid:
reg_valid r s ->
map_valid (add_letvar map r) s.
Proof.
- intros; red; intros.
- destruct H1 as [[id A]|B].
+ intros; red; intros.
+ destruct H1 as [[id A]|B].
simpl in A. apply H. left; exists id; auto.
- simpl in B. elim B; intro.
+ simpl in B. elim B; intro.
subst r0; auto. apply H. right; auto.
Qed.
@@ -467,7 +467,7 @@ Lemma alloc_regs_valid:
Proof.
induction al; simpl; intros; monadInv H0.
apply regs_valid_nil.
- apply regs_valid_cons. eauto with rtlg. eauto with rtlg.
+ apply regs_valid_cons. eauto with rtlg. eauto with rtlg.
Qed.
Hint Resolve alloc_regs_valid: rtlg.
@@ -479,7 +479,7 @@ Lemma alloc_regs_fresh_or_in_map:
Proof.
induction al; simpl; intros; monadInv H0.
elim H1.
- elim H1; intro.
+ elim H1; intro.
subst r.
eapply alloc_reg_fresh_or_in_map; eauto.
exploit IHal. 2: eauto. apply map_valid_incr with s; eauto with rtlg. eauto.
@@ -502,7 +502,7 @@ Lemma alloc_optreg_fresh_or_in_map:
alloc_optreg map dest s = OK r s' i ->
reg_in_map map r \/ reg_fresh r s.
Proof.
- intros until s'. unfold alloc_optreg. destruct dest; intros.
+ intros until s'. unfold alloc_optreg. destruct dest; intros.
left; eauto with rtlg.
right; eauto with rtlg.
Qed.
@@ -546,8 +546,8 @@ Proof.
induction 1; intros.
constructor; auto.
constructor; auto.
- constructor; auto. red; intros.
- elim (in_app_or _ _ _ H2); intro.
+ constructor; auto. red; intros.
+ elim (in_app_or _ _ _ H2); intro.
generalize (H1 _ H3). tauto. tauto.
Qed.
@@ -559,7 +559,7 @@ Lemma target_reg_ok_cons:
target_reg_ok map (r' :: pr) a r.
Proof.
intros. change (r' :: pr) with ((r' :: nil) ++ pr).
- apply target_reg_ok_append; auto.
+ apply target_reg_ok_append; auto.
intros r'' [A|B]. subst r''; auto. contradiction.
Qed.
@@ -570,7 +570,7 @@ Lemma new_reg_target_ok:
new_reg s1 = OK r s2 i ->
target_reg_ok map pr a r.
Proof.
- intros. constructor.
+ intros. constructor.
red; intro. apply valid_fresh_absurd with r s1.
eauto with rtlg. eauto with rtlg.
red; intro. apply valid_fresh_absurd with r s1.
@@ -604,13 +604,13 @@ Proof.
induction al; intros; monadInv H1.
constructor.
constructor.
- eapply alloc_reg_target_ok; eauto.
- apply IHal with s s2 INCR1; eauto with rtlg.
+ eapply alloc_reg_target_ok; eauto.
+ apply IHal with s s2 INCR1; eauto with rtlg.
apply regs_valid_cons; eauto with rtlg.
Qed.
-Hint Resolve new_reg_target_ok alloc_reg_target_ok
- alloc_regs_target_ok: rtlg.
+Hint Resolve new_reg_target_ok alloc_reg_target_ok
+ alloc_regs_target_ok: rtlg.
(** The following predicate is a variant of [target_reg_ok] used
to characterize registers that are adequate for holding the return
@@ -640,7 +640,7 @@ Lemma new_reg_return_ok:
return_reg_ok s2 map (ret_reg sig r).
Proof.
intros. unfold ret_reg. destruct (sig_res sig); constructor.
- eauto with rtlg. eauto with rtlg.
+ eauto with rtlg. eauto with rtlg.
Qed.
(** * Relational specification of the translation *)
@@ -693,7 +693,7 @@ Hint Resolve reg_map_ok_novar: rtlg.
and moreover that they satisfy the [reg_map_ok] predicate.
*)
-Inductive tr_expr (c: code):
+Inductive tr_expr (c: code):
mapping -> list reg -> expr -> node -> node -> reg -> option ident -> Prop :=
| tr_Evar: forall map pr id ns nd r rd dst,
map.(map_vars)!id = Some r ->
@@ -742,7 +742,7 @@ Inductive tr_expr (c: code):
value of the CminorSel conditional expression [a] and terminate
on node [ntrue] if the condition holds and on node [nfalse] otherwise. *)
-with tr_condition (c: code):
+with tr_condition (c: code):
mapping -> list reg -> condexpr -> node -> node -> node -> Prop :=
| tr_CEcond: forall map pr cond bl ns ntrue nfalse n1 rl,
tr_exprlist c map pr bl ns n1 rl ->
@@ -764,7 +764,7 @@ with tr_condition (c: code):
of the list of CminorSel expression [exprlist] and deposit these values
in registers [rds]. *)
-with tr_exprlist (c: code):
+with tr_exprlist (c: code):
mapping -> list reg -> exprlist -> node -> node -> list reg -> Prop :=
| tr_Enil: forall map pr n,
tr_exprlist c map pr Enil n n nil
@@ -786,7 +786,7 @@ Definition tr_jumptable (nexits: list node) (tbl: list nat) (ttbl: list node) :
on the node corresponding to this exit number according to the
mapping [nexits]. *)
-Inductive tr_exitexpr (c: code):
+Inductive tr_exitexpr (c: code):
mapping -> exitexpr -> node -> list node -> Prop :=
| tr_XEcond: forall map x n nexits,
nth_error nexits x = Some n ->
@@ -903,7 +903,7 @@ Inductive tr_stmt (c: code) (map: mapping):
ngoto!lbl = Some ns ->
tr_stmt c map (Sgoto lbl) ns nd nexits ngoto nret rret.
-(** [tr_function f tf] specifies the RTL function [tf] that
+(** [tr_function f tf] specifies the RTL function [tf] that
[RTLgen.transl_function] returns. *)
Inductive tr_function: CminorSel.function -> RTL.function -> Prop :=
@@ -913,7 +913,7 @@ Inductive tr_function: CminorSel.function -> RTL.function -> Prop :=
add_vars map1 f.(CminorSel.fn_vars) s1 = OK (rvars, map2) s2 i2 ->
orret = ret_reg f.(CminorSel.fn_sig) rret ->
tr_stmt code map2 f.(CminorSel.fn_body) nentry nret nil ngoto nret orret ->
- code!nret = Some(Ireturn orret) ->
+ code!nret = Some(Ireturn orret) ->
tr_function f (RTL.mkfunction
f.(CminorSel.fn_sig)
rparams
@@ -951,22 +951,22 @@ with tr_exprlist_incr:
tr_exprlist s1.(st_code) map pr al ns nd rl ->
tr_exprlist s2.(st_code) map pr al ns nd rl.
Proof.
- intros s1 s2 EXT.
+ intros s1 s2 EXT.
pose (AT := fun pc i => instr_at_incr s1 s2 pc i EXT).
induction 1; econstructor; eauto.
eapply tr_move_incr; eauto.
eapply tr_move_incr; eauto.
- intros s1 s2 EXT.
+ intros s1 s2 EXT.
pose (AT := fun pc i => instr_at_incr s1 s2 pc i EXT).
induction 1; econstructor; eauto.
- intros s1 s2 EXT.
+ intros s1 s2 EXT.
pose (AT := fun pc i => instr_at_incr s1 s2 pc i EXT).
induction 1; econstructor; eauto.
Qed.
Lemma add_move_charact:
forall s ns rs nd rd s' i,
- add_move rs rd nd s = OK ns s' i ->
+ add_move rs rd nd s = OK ns s' i ->
tr_move s'.(st_code) ns rs nd rd.
Proof.
intros. unfold add_move in H. destruct (Reg.eq rs rd).
@@ -1013,23 +1013,23 @@ Proof.
(* Eload *)
inv OK.
econstructor; eauto with rtlg.
- eapply transl_exprlist_charact; eauto with rtlg.
+ eapply transl_exprlist_charact; eauto with rtlg.
(* Econdition *)
inv OK.
econstructor.
eauto with rtlg.
- apply tr_expr_incr with s1; auto.
+ apply tr_expr_incr with s1; auto.
eapply transl_expr_charact; eauto 2 with rtlg. constructor; auto.
- apply tr_expr_incr with s0; auto.
+ apply tr_expr_incr with s0; auto.
eapply transl_expr_charact; eauto 2 with rtlg. constructor; auto.
(* Elet *)
inv OK.
- econstructor. eapply new_reg_not_in_map; eauto with rtlg.
+ 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.
- eapply transl_expr_charact. eauto.
- apply add_letvar_valid; eauto with rtlg.
- constructor; auto.
+ eapply transl_expr_charact. eauto.
+ apply add_letvar_valid; eauto with rtlg.
+ constructor; auto.
red; unfold reg_in_map. simpl. intros [[id A] | [B | C]].
elim H. left; exists id; auto.
subst x. apply valid_fresh_absurd with rd s. auto. eauto with rtlg.
@@ -1038,7 +1038,7 @@ Proof.
(* Eletvar *)
generalize EQ; unfold find_letvar. caseEq (nth_error (map_letvars map) n); intros; inv EQ0.
monadInv EQ1.
- econstructor; eauto with rtlg.
+ econstructor; eauto with rtlg.
inv OK. left; split; congruence. right; eauto with rtlg.
eapply add_move_charact; eauto.
monadInv EQ1.
@@ -1064,7 +1064,7 @@ Proof.
generalize (VALID2 r (in_eq _ _)). eauto with rtlg.
apply tr_exprlist_incr with s0; auto.
eapply transl_exprlist_charact; eauto with rtlg.
- apply regs_valid_cons. apply VALID2. auto with coqlib. auto.
+ apply regs_valid_cons. apply VALID2. auto with coqlib. auto.
red; intros; apply VALID2; auto with coqlib.
(* Conditional expressions *)
@@ -1073,15 +1073,15 @@ Proof.
(* CEcond *)
econstructor; eauto with rtlg. eapply transl_exprlist_charact; eauto with rtlg.
(* CEcondition *)
- econstructor; eauto with rtlg.
+ econstructor; eauto with rtlg.
apply tr_condition_incr with s1; eauto with rtlg.
apply tr_condition_incr with s0; eauto with rtlg.
(* CElet *)
econstructor; eauto with rtlg.
- eapply transl_expr_charact; eauto with rtlg.
+ eapply transl_expr_charact; eauto with rtlg.
apply tr_condition_incr with s1; eauto with rtlg.
eapply transl_condexpr_charact; eauto with rtlg.
- apply add_letvar_valid; eauto with rtlg.
+ apply add_letvar_valid; eauto with rtlg.
Qed.
(** A variant of [transl_expr_charact], for use when the destination
@@ -1100,20 +1100,20 @@ Proof.
econstructor; eauto.
eapply add_move_charact; eauto.
(* Eop *)
- econstructor; eauto with rtlg.
+ econstructor; eauto with rtlg.
eapply transl_exprlist_charact; eauto with rtlg.
(* Eload *)
econstructor; eauto with rtlg.
- eapply transl_exprlist_charact; eauto with rtlg.
+ eapply transl_exprlist_charact; eauto with rtlg.
(* Econdition *)
econstructor; eauto with rtlg.
eapply transl_condexpr_charact; eauto with rtlg.
- apply tr_expr_incr with s1; auto.
- eapply IHa1; eauto 2 with rtlg.
- apply tr_expr_incr with s0; auto.
+ apply tr_expr_incr with s1; auto.
+ eapply IHa1; eauto 2 with rtlg.
+ apply tr_expr_incr with s0; auto.
eapply IHa2; eauto 2 with rtlg.
(* Elet *)
- econstructor. eapply new_reg_not_in_map; eauto with rtlg.
+ 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.
eapply IHa2; eauto.
@@ -1122,14 +1122,14 @@ Proof.
(* Eletvar *)
generalize EQ; unfold find_letvar. caseEq (nth_error (map_letvars map) n); intros; inv EQ0.
monadInv EQ1.
- econstructor; eauto with rtlg.
+ econstructor; eauto with rtlg.
eapply add_move_charact; eauto.
monadInv EQ1.
(* Ebuiltin *)
- econstructor; eauto with rtlg.
+ econstructor; eauto with rtlg.
eapply transl_exprlist_charact; eauto with rtlg.
(* Eexternal *)
- econstructor; eauto with rtlg.
+ econstructor; eauto with rtlg.
eapply transl_exprlist_charact; eauto with rtlg.
Qed.
@@ -1172,7 +1172,7 @@ Lemma transl_exit_charact:
transl_exit nexits n s = OK ne s' incr ->
nth_error nexits n = Some ne /\ s' = s.
Proof.
- intros until incr. unfold transl_exit.
+ intros until incr. unfold transl_exit.
destruct (nth_error nexits n); intro; monadInv H. auto.
Qed.
@@ -1183,7 +1183,7 @@ Lemma transl_jumptable_charact:
Proof.
induction tbl; intros.
monadInv H. split. red. simpl. intros. discriminate. auto.
- monadInv H. exploit transl_exit_charact; eauto. intros [A B].
+ monadInv H. exploit transl_exit_charact; eauto. intros [A B].
exploit IHtbl; eauto. intros [C D].
split. red. simpl. intros. destruct (zeq v 0). inv H. exists x; auto. auto.
congruence.
@@ -1198,16 +1198,16 @@ Proof.
induction a; simpl; intros; try (monadInv TR); saturateTrans.
- (* XEexit *)
exploit transl_exit_charact; eauto. intros [A B].
- econstructor; eauto.
+ econstructor; eauto.
- (* XEjumptable *)
exploit transl_jumptable_charact; eauto. intros [A B].
- econstructor; eauto.
- eapply transl_expr_charact; eauto with rtlg.
+ econstructor; eauto.
+ eapply transl_expr_charact; eauto with rtlg.
eauto with rtlg.
- (* XEcondition *)
- econstructor.
+ econstructor.
eapply transl_condexpr_charact; eauto with rtlg.
- apply tr_exitexpr_incr with s1; eauto with rtlg.
+ apply tr_exitexpr_incr with s1; eauto with rtlg.
apply tr_exitexpr_incr with s0; eauto with rtlg.
- (* XElet *)
econstructor; eauto with rtlg.
@@ -1225,7 +1225,7 @@ Proof.
destruct res; simpl; intros.
- monadInv TR. constructor. unfold find_var in EQ. destruct (map_vars map)!x; inv EQ; auto.
- destruct oty; monadInv TR.
-+ constructor. eauto with rtlg.
++ constructor. eauto with rtlg.
+ constructor.
- monadInv TR.
Qed.
@@ -1255,7 +1255,7 @@ Proof.
(* indirect *)
econstructor; eauto 4 with rtlg.
eapply transl_expr_charact; eauto 3 with rtlg.
- apply tr_exprlist_incr with s5. auto.
+ apply tr_exprlist_incr with s5. auto.
eapply transl_exprlist_charact; eauto 3 with rtlg.
eapply alloc_regs_target_ok with (s1 := s0); eauto 3 with rtlg.
apply regs_valid_cons; eauto 3 with rtlg.
@@ -1271,7 +1271,7 @@ Proof.
destruct s0 as [b | id]; monadInv TR; saturateTrans.
(* indirect *)
assert (RV: regs_valid (x :: nil) s0).
- apply regs_valid_cons; eauto 3 with rtlg.
+ apply regs_valid_cons; eauto 3 with rtlg.
econstructor; eauto 3 with rtlg.
eapply transl_expr_charact; eauto 3 with rtlg.
apply tr_exprlist_incr with s4; auto.
@@ -1284,47 +1284,47 @@ Proof.
eapply transl_exprlist_charact; eauto 3 with rtlg.
eapply convert_builtin_res_charact; eauto with rtlg.
(* Sseq *)
- econstructor.
- apply tr_stmt_incr with s0; auto.
+ econstructor.
+ apply tr_stmt_incr with s0; auto.
eapply IHstmt2; eauto with rtlg.
eapply IHstmt1; eauto with rtlg.
(* Sifthenelse *)
destruct (more_likely c stmt1 stmt2); monadInv TR.
econstructor.
- apply tr_stmt_incr with s1; auto.
+ apply tr_stmt_incr with s1; auto.
eapply IHstmt1; eauto with rtlg.
apply tr_stmt_incr with s0; auto.
eapply IHstmt2; eauto with rtlg.
eapply transl_condexpr_charact; eauto with rtlg.
econstructor.
- apply tr_stmt_incr with s0; auto.
+ apply tr_stmt_incr with s0; auto.
eapply IHstmt1; eauto with rtlg.
apply tr_stmt_incr with s1; auto.
eapply IHstmt2; eauto with rtlg.
eapply transl_condexpr_charact; eauto with rtlg.
(* Sloop *)
- econstructor.
- apply tr_stmt_incr with s1; auto.
+ econstructor.
+ apply tr_stmt_incr with s1; auto.
eapply IHstmt; eauto with rtlg.
- eauto with rtlg. eauto with rtlg.
+ eauto with rtlg. eauto with rtlg.
(* Sblock *)
- econstructor.
+ econstructor.
eapply IHstmt; eauto with rtlg.
(* Sexit *)
exploit transl_exit_charact; eauto. intros [A B].
econstructor. eauto.
(* Sswitch *)
- econstructor. eapply transl_exitexpr_charact; eauto.
+ econstructor. eapply transl_exitexpr_charact; eauto.
(* Sreturn *)
- destruct o.
- destruct rret; inv TR. inv OK.
- econstructor; eauto with rtlg.
+ destruct o.
+ destruct rret; inv TR. inv OK.
+ econstructor; eauto with rtlg.
eapply transl_expr_charact; eauto with rtlg.
- constructor. auto. simpl; tauto.
+ constructor. auto. simpl; tauto.
monadInv TR. constructor.
(* Slabel *)
generalize EQ0; clear EQ0. case_eq (ngoto!l); intros; monadInv EQ0.
- generalize EQ1; clear EQ1. unfold handle_error.
+ generalize EQ1; clear EQ1. unfold handle_error.
case_eq (update_instr n (Inop ns) s0); intros; inv EQ1.
econstructor. eauto. eauto with rtlg.
eapply tr_stmt_incr with s0; eauto with rtlg.
@@ -1339,17 +1339,17 @@ Lemma transl_function_charact:
tr_function f tf.
Proof.
intros until tf. unfold transl_function.
- caseEq (reserve_labels (fn_body f) (PTree.empty node, init_state)).
+ caseEq (reserve_labels (fn_body f) (PTree.empty node, init_state)).
intros ngoto s0 RESERVE.
- caseEq (transl_fun f ngoto s0). congruence.
- intros [nentry rparams] sfinal INCR TR E. inv E.
+ caseEq (transl_fun f ngoto s0). congruence.
+ intros [nentry rparams] sfinal INCR TR E. inv E.
monadInv TR.
exploit add_vars_valid. eexact EQ. apply init_mapping_valid.
- intros [A B].
- exploit add_vars_valid. eexact EQ1. auto.
+ intros [A B].
+ exploit add_vars_valid. eexact EQ1. auto.
intros [C D].
eapply tr_function_intro; eauto with rtlg.
- eapply transl_stmt_charact; eauto with rtlg.
+ eapply transl_stmt_charact; eauto with rtlg.
unfold ret_reg. destruct (sig_res (CminorSel.fn_sig f)).
constructor. eauto with rtlg. eauto with rtlg.
constructor.
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index effb0c7d..57fc8b86 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -172,7 +172,7 @@ Record wt_function (f: function) (env: regenv): Prop :=
wt_norepet:
list_norepet f.(fn_params);
wt_instrs:
- forall pc instr,
+ forall pc instr,
f.(fn_code)!pc = Some instr -> wt_instr f env instr;
wt_entrypoint:
valid_successor f f.(fn_entrypoint)
@@ -304,7 +304,7 @@ Definition type_instr (e: S.typenv) (i: instruction) : res S.typenv :=
| Ibuiltin ef args res s =>
let sig := ef_sig ef in
do x <- check_successor s;
- do e1 <-
+ do e1 <-
match ef with
| EF_annot _ _ | EF_debug _ _ _ => OK e
| _ => type_builtin_args e args sig.(sig_args)
@@ -342,7 +342,7 @@ Definition type_code (e: S.typenv): res S.typenv :=
(** Solve remaining constraints *)
-Definition check_params_norepet (params: list reg): res unit :=
+Definition check_params_norepet (params: list reg): res unit :=
if list_norepet_dec Reg.eq params
then OK tt
else Error(msg "duplicate parameters").
@@ -369,7 +369,7 @@ Lemma type_ros_sound:
forall e ros e' te, type_ros e ros = OK e' -> S.satisf te e' ->
match ros with inl r => te r = Tint | inr s => True end.
Proof.
- unfold type_ros; intros. destruct ros.
+ unfold type_ros; intros. destruct ros.
eapply S.set_sound; eauto.
auto.
Qed.
@@ -377,7 +377,7 @@ Qed.
Lemma check_successor_sound:
forall s x, check_successor s = OK x -> valid_successor f s.
Proof.
- unfold check_successor, valid_successor; intros.
+ unfold check_successor, valid_successor; intros.
destruct (fn_code f)!s; inv H. exists i; auto.
Qed.
@@ -386,9 +386,9 @@ Hint Resolve check_successor_sound: ty.
Lemma check_successors_sound:
forall sl x, check_successors sl = OK x -> forall s, In s sl -> valid_successor f s.
Proof.
- induction sl; simpl; intros.
+ induction sl; simpl; intros.
contradiction.
- monadInv H. destruct H0. subst a; eauto with ty. eauto.
+ monadInv H. destruct H0. subst a; eauto with ty. eauto.
Qed.
Remark type_expect_incr:
@@ -416,7 +416,7 @@ Lemma type_builtin_args_incr:
Proof.
induction a; destruct ty; simpl; intros; try discriminate.
inv H; auto.
- monadInv H. eapply type_builtin_arg_incr; eauto.
+ monadInv H. eapply type_builtin_arg_incr; eauto.
Qed.
Lemma type_builtin_res_incr:
@@ -450,7 +450,7 @@ Lemma type_builtin_res_sound:
forall e a ty e' te,
type_builtin_res e a ty = OK e' -> S.satisf te e' -> type_of_builtin_res te a = ty.
Proof.
- intros. destruct a; simpl in *.
+ intros. destruct a; simpl in *.
eapply S.set_sound; eauto.
symmetry; eapply type_expect_sound; eauto.
symmetry; eapply type_expect_sound; eauto.
@@ -495,7 +495,7 @@ Proof.
destruct l; try discriminate. destruct l; monadInv EQ0.
constructor. eapply S.move_sound; eauto. eauto with ty.
+ destruct (type_of_operation o) as [targs tres] eqn:TYOP. monadInv EQ0.
- apply wt_Iop.
+ apply wt_Iop.
unfold is_move in ISMOVE; destruct o; congruence.
rewrite TYOP. eapply S.set_list_sound; eauto with ty.
rewrite TYOP. eapply S.set_sound; eauto with ty.
@@ -511,7 +511,7 @@ Proof.
eapply S.set_sound; eauto with ty.
eauto with ty.
- (* call *)
- constructor.
+ constructor.
eapply type_ros_sound; eauto with ty.
eapply S.set_list_sound; eauto with ty.
eapply S.set_sound; eauto with ty.
@@ -520,7 +520,7 @@ Proof.
destruct (opt_typ_eq (sig_res s) (sig_res (fn_sig f))); try discriminate.
destruct (tailcall_is_possible s) eqn:TCIP; inv EQ2.
constructor.
- eapply type_ros_sound; eauto with ty.
+ eapply type_ros_sound; eauto with ty.
eapply S.set_list_sound; eauto with ty.
auto.
apply tailcall_is_possible_correct; auto.
@@ -538,12 +538,12 @@ Proof.
destruct (zle (list_length_z l * 4) Int.max_unsigned); inv EQ2.
constructor.
eapply S.set_sound; eauto.
- eapply check_successors_sound; eauto.
+ eapply check_successors_sound; eauto.
auto.
- (* return *)
simpl in H. destruct o as [r|] eqn: RET; destruct (sig_res (fn_sig f)) as [t|] eqn: RES; try discriminate.
econstructor. eauto. eapply S.set_sound; eauto with ty.
- inv H. constructor. auto.
+ inv H. constructor. auto.
Qed.
Lemma type_code_sound:
@@ -558,16 +558,16 @@ Proof.
| OK e' => c!pc = Some i -> S.satisf te e' -> wt_instr f te i
end).
change (P f.(fn_code) (OK e1)).
- rewrite <- TCODE. unfold type_code. apply PTree_Properties.fold_rec; unfold P; intros.
+ rewrite <- TCODE. unfold type_code. apply PTree_Properties.fold_rec; unfold P; intros.
- (* extensionality *)
- destruct a; auto; intros. rewrite <- H in H1. eapply H0; eauto.
+ destruct a; auto; intros. rewrite <- H in H1. eapply H0; eauto.
- (* base case *)
rewrite PTree.gempty in H; discriminate.
- (* inductive case *)
- destruct a as [e|?]; auto.
+ destruct a as [e|?]; auto.
destruct (type_instr e v) as [e'|?] eqn:TYINSTR; auto.
- intros. rewrite PTree.gsspec in H2. destruct (peq pc k).
- inv H2. eapply type_instr_sound; eauto.
+ intros. rewrite PTree.gsspec in H2. destruct (peq pc k).
+ inv H2. eapply type_instr_sound; eauto.
eapply H1; eauto. eapply type_instr_incr; eauto.
Qed.
@@ -581,12 +581,12 @@ Proof.
- (* type of parameters *)
eapply S.set_list_sound; eauto.
- (* parameters are unique *)
- unfold check_params_norepet in EQ2.
- destruct (list_norepet_dec Reg.eq (fn_params f)); inv EQ2; auto.
+ unfold check_params_norepet in EQ2.
+ destruct (list_norepet_dec Reg.eq (fn_params f)); inv EQ2; auto.
- (* instructions are well typed *)
- intros. eapply type_code_sound; eauto.
+ intros. eapply type_code_sound; eauto.
- (* entry point is valid *)
- eauto with ty.
+ eauto with ty.
Qed.
(** ** Completeness proof *)
@@ -597,7 +597,7 @@ Lemma type_ros_complete:
match ros with inl r => te r = Tint | inr s => True end ->
exists e', type_ros e ros = OK e' /\ S.satisf te e'.
Proof.
- intros; destruct ros; simpl.
+ intros; destruct ros; simpl.
eapply S.set_complete; eauto.
exists e; auto.
Qed.
@@ -605,14 +605,14 @@ Qed.
Lemma check_successor_complete:
forall s, valid_successor f s -> check_successor s = OK tt.
Proof.
- unfold valid_successor, check_successor; intros.
+ unfold valid_successor, check_successor; intros.
destruct H as [i EQ]; rewrite EQ; auto.
Qed.
Lemma type_expect_complete:
forall e ty, type_expect e ty ty = OK e.
Proof.
- unfold type_expect; intros. rewrite dec_eq_true; auto.
+ unfold type_expect; intros. rewrite dec_eq_true; auto.
Qed.
Lemma type_builtin_arg_complete:
@@ -620,7 +620,7 @@ Lemma type_builtin_arg_complete:
S.satisf te e ->
exists e', type_builtin_arg e a (type_of_builtin_arg te a) = OK e' /\ S.satisf te e'.
Proof.
- intros. destruct a; simpl; try (exists e; split; [apply type_expect_complete|assumption]).
+ intros. destruct a; simpl; try (exists e; split; [apply type_expect_complete|assumption]).
apply S.set_complete; auto.
Qed.
@@ -629,11 +629,11 @@ Lemma type_builtin_args_complete:
S.satisf te e ->
exists e', type_builtin_args e al (List.map (type_of_builtin_arg te) al) = OK e' /\ S.satisf te e'.
Proof.
- induction al; simpl; intros.
+ induction al; simpl; intros.
- exists e; auto.
-- destruct (type_builtin_arg_complete te a e) as (e1 & A & B); auto.
+- destruct (type_builtin_arg_complete te a e) as (e1 & A & B); auto.
destruct (IHal e1) as (e2 & C & D); auto.
- exists e2; split; auto. rewrite A. auto.
+ exists e2; split; auto. rewrite A. auto.
Qed.
Lemma type_builtin_res_complete:
@@ -641,7 +641,7 @@ Lemma type_builtin_res_complete:
S.satisf te e ->
exists e', type_builtin_res e a (type_of_builtin_res te a) = OK e' /\ S.satisf te e'.
Proof.
- intros. destruct a; simpl.
+ intros. destruct a; simpl.
apply S.set_complete; auto.
exists e; auto.
exists e; auto.
@@ -664,60 +664,60 @@ Proof.
exploit S.set_list_complete. eauto. eauto. intros [e1 [A B]].
exploit S.set_complete. eexact B. eauto. intros [e2 [C D]].
exists e2; split; auto.
- rewrite check_successor_complete by auto; simpl.
+ rewrite check_successor_complete by auto; simpl.
replace (is_move op) with false. rewrite A; simpl; rewrite C; auto.
destruct op; reflexivity || congruence.
- (* load *)
exploit S.set_list_complete. eauto. eauto. intros [e1 [A B]].
exploit S.set_complete. eexact B. eauto. intros [e2 [C D]].
exists e2; split; auto.
- rewrite check_successor_complete by auto; simpl.
+ rewrite check_successor_complete by auto; simpl.
rewrite A; simpl; rewrite C; auto.
- (* store *)
exploit S.set_list_complete. eauto. eauto. intros [e1 [A B]].
exploit S.set_complete. eexact B. eauto. intros [e2 [C D]].
exists e2; split; auto.
- rewrite check_successor_complete by auto; simpl.
+ rewrite check_successor_complete by auto; simpl.
rewrite A; simpl; rewrite C; auto.
- (* call *)
exploit type_ros_complete. eauto. eauto. intros [e1 [A B]].
exploit S.set_list_complete. eauto. eauto. intros [e2 [C D]].
exploit S.set_complete. eexact D. eauto. intros [e3 [E F]].
- exists e3; split; auto.
- rewrite check_successor_complete by auto; simpl.
+ exists e3; split; auto.
+ rewrite check_successor_complete by auto; simpl.
rewrite A; simpl; rewrite C; simpl; rewrite E; auto.
- (* tailcall *)
exploit type_ros_complete. eauto. eauto. intros [e1 [A B]].
exploit S.set_list_complete. eauto. eauto. intros [e2 [C D]].
- exists e2; split; auto.
- rewrite A; simpl; rewrite C; simpl.
- rewrite H2; rewrite dec_eq_true.
- replace (tailcall_is_possible sig) with true; auto.
- revert H3. unfold tailcall_possible, tailcall_is_possible. generalize (loc_arguments sig).
+ exists e2; split; auto.
+ rewrite A; simpl; rewrite C; simpl.
+ rewrite H2; rewrite dec_eq_true.
+ replace (tailcall_is_possible sig) with true; auto.
+ revert H3. unfold tailcall_possible, tailcall_is_possible. generalize (loc_arguments sig).
induction l; simpl; intros. auto.
exploit (H3 a); auto. intros. destruct a; try contradiction. apply IHl.
- intros; apply H3; auto.
+ intros; apply H3; auto.
- (* builtin *)
exploit type_builtin_args_complete; eauto. instantiate (1 := args). intros [e1 [A B]].
exploit type_builtin_res_complete; eauto. instantiate (1 := res). intros [e2 [C D]].
exploit type_builtin_res_complete. eexact H. instantiate (1 := res). intros [e3 [E F]].
rewrite check_successor_complete by auto. simpl.
exists (match ef with EF_annot _ _ | EF_debug _ _ _ => e3 | _ => e2 end); split.
- rewrite H1 in C, E.
+ rewrite H1 in C, E.
destruct ef; try (rewrite <- H0; rewrite A); simpl; auto.
destruct ef; auto.
- (* cond *)
exploit S.set_list_complete. eauto. eauto. intros [e1 [A B]].
exists e1; split; auto.
- rewrite check_successor_complete by auto; simpl.
+ rewrite check_successor_complete by auto; simpl.
rewrite check_successor_complete by auto; simpl.
auto.
- (* jumptbl *)
exploit S.set_complete. eauto. eauto. intros [e1 [A B]].
exists e1; split; auto.
- replace (check_successors tbl) with (OK tt). simpl.
- rewrite A; simpl. apply zle_true; auto.
- revert H1. generalize tbl. induction tbl0; simpl; intros. auto.
+ replace (check_successors tbl) with (OK tt). simpl.
+ rewrite A; simpl. apply zle_true; auto.
+ revert H1. generalize tbl. induction tbl0; simpl; intros. auto.
rewrite check_successor_complete by auto; simpl.
apply IHtbl0; intros; auto.
- (* return none *)
@@ -739,14 +739,14 @@ Proof.
assert (P f.(fn_code) (type_code e0)).
{
unfold type_code. apply PTree_Properties.fold_rec; unfold P; intros.
- - apply H0. intros. apply H1 with pc. rewrite <- H; auto.
- - exists e0; auto.
- - destruct H1 as [e [A B]].
+ - apply H0. intros. apply H1 with pc. rewrite <- H; auto.
+ - exists e0; auto.
+ - destruct H1 as [e [A B]].
intros. apply H2 with pc. rewrite PTree.gso; auto. congruence.
- subst a.
+ subst a.
destruct (type_instr_complete te e v) as [e' [C D]].
- auto. apply H2 with k. apply PTree.gss.
- exists e'; split; auto. rewrite C; auto.
+ auto. apply H2 with k. apply PTree.gss.
+ exists e'; split; auto. rewrite C; auto.
}
apply H; auto.
Qed.
@@ -754,15 +754,15 @@ Qed.
Theorem type_function_complete:
forall te, wt_function f te -> exists te, type_function = OK te.
Proof.
- intros. destruct H.
+ intros. destruct H.
destruct (type_code_complete te S.initial) as (e1 & A & B).
- auto. apply S.satisf_initial.
+ auto. apply S.satisf_initial.
destruct (S.set_list_complete te f.(fn_params) f.(fn_sig).(sig_args) e1) as (e2 & C & D); auto.
destruct (S.solve_complete te e2) as (te' & E); auto.
exists te'; unfold type_function.
- rewrite A; simpl. rewrite C; simpl. rewrite E; simpl.
- unfold check_params_norepet. rewrite pred_dec_true; auto. simpl.
- rewrite check_successor_complete by auto. auto.
+ rewrite A; simpl. rewrite C; simpl. rewrite E; simpl.
+ unfold check_params_norepet. rewrite pred_dec_true; auto. simpl.
+ rewrite check_successor_complete by auto. auto.
Qed.
End INFERENCE.
@@ -790,7 +790,7 @@ Lemma wt_regset_assign:
Val.has_type v (env r) ->
wt_regset env (rs#r <- v).
Proof.
- intros; red; intros.
+ intros; red; intros.
rewrite Regmap.gsspec.
case (peq r0 r); intro.
subst r0. assumption.
@@ -805,7 +805,7 @@ Proof.
induction rl; simpl.
auto.
split. apply H. apply IHrl.
-Qed.
+Qed.
Lemma wt_regset_setres:
forall env rs v res,
@@ -813,8 +813,8 @@ Lemma wt_regset_setres:
Val.has_type v (type_of_builtin_res env res) ->
wt_regset env (regmap_setres res v rs).
Proof.
- intros. destruct res; simpl in *; auto. apply wt_regset_assign; auto.
-Qed.
+ intros. destruct res; simpl in *; auto. apply wt_regset_assign; auto.
+Qed.
Lemma wt_init_regs:
forall env rl args,
@@ -822,7 +822,7 @@ Lemma wt_init_regs:
wt_regset env (init_regs args rl).
Proof.
induction rl; destruct args; simpl; intuition.
- red; intros. rewrite Regmap.gi. simpl; auto.
+ red; intros. rewrite Regmap.gi. simpl; auto.
apply wt_regset_assign; auto.
Qed.
@@ -833,7 +833,7 @@ Lemma wt_exec_Iop:
wt_regset env rs ->
wt_regset env (rs#res <- v).
Proof.
- intros. inv H.
+ intros. inv H.
simpl in H0. inv H0. apply wt_regset_assign; auto.
rewrite H4; auto.
eapply wt_regset_assign; auto.
@@ -858,7 +858,7 @@ Lemma wt_exec_Ibuiltin:
wt_regset env rs ->
wt_regset env (regmap_setres res vres rs).
Proof.
- intros. inv H.
+ intros. inv H.
eapply wt_regset_setres; eauto.
rewrite H7. eapply external_call_well_typed; eauto.
Qed.
@@ -867,7 +867,7 @@ Lemma wt_instr_at:
forall f env pc i,
wt_function f env -> f.(fn_code)!pc = Some i -> wt_instr f env i.
Proof.
- intros. inv H. eauto.
+ intros. inv H. eauto.
Qed.
Inductive wt_stackframes: list stackframe -> signature -> Prop :=
@@ -905,9 +905,9 @@ Remark wt_stackframes_change_sig:
forall s sg1 sg2,
sg1.(sig_res) = sg2.(sig_res) -> wt_stackframes s sg1 -> wt_stackframes s sg2.
Proof.
- intros. inv H0.
+ intros. inv H0.
- constructor; congruence.
-- econstructor; eauto. rewrite H3. unfold proj_sig_res. rewrite H. auto.
+- econstructor; eauto. rewrite H3. unfold proj_sig_res. rewrite H. auto.
Qed.
Section SUBJECT_REDUCTION.
@@ -936,19 +936,19 @@ Proof.
assert (wt_fundef fd).
destruct ros; simpl in H0.
pattern fd. apply Genv.find_funct_prop with fundef unit p (rs#r).
- exact wt_p. exact H0.
+ exact wt_p. exact H0.
caseEq (Genv.find_symbol ge i); intros; rewrite H1 in H0.
pattern fd. apply Genv.find_funct_ptr_prop with fundef unit p b.
exact wt_p. exact H0.
discriminate.
econstructor; eauto.
- econstructor; eauto. inv WTI; auto.
+ econstructor; eauto. inv WTI; auto.
inv WTI. rewrite <- H8. apply wt_regset_list. auto.
(* Itailcall *)
assert (wt_fundef fd).
destruct ros; simpl in H0.
pattern fd. apply Genv.find_funct_prop with fundef unit p (rs#r).
- exact wt_p. exact H0.
+ exact wt_p. exact H0.
caseEq (Genv.find_symbol ge i); intros; rewrite H1 in H0.
pattern fd. apply Genv.find_funct_ptr_prop with fundef unit p b.
exact wt_p. exact H0.
@@ -963,24 +963,24 @@ Proof.
(* Ijumptable *)
econstructor; eauto.
(* Ireturn *)
- econstructor; eauto.
- inv WTI; simpl. auto. unfold proj_sig_res; rewrite H2. auto.
+ econstructor; eauto.
+ inv WTI; simpl. auto. unfold proj_sig_res; rewrite H2. auto.
(* internal function *)
simpl in *. inv H5.
econstructor; eauto.
- inv H1. apply wt_init_regs; auto. rewrite wt_params0. auto.
+ inv H1. apply wt_init_regs; auto. rewrite wt_params0. auto.
(* external function *)
- econstructor; eauto. simpl.
+ econstructor; eauto. simpl.
eapply external_call_well_typed; eauto.
(* return *)
inv H1. econstructor; eauto.
- apply wt_regset_assign; auto. rewrite H10; auto.
+ apply wt_regset_assign; auto. rewrite H10; auto.
Qed.
Lemma wt_initial_state:
forall S, initial_state p S -> wt_state S.
Proof.
- intros. inv H. constructor. constructor. rewrite H3; auto.
+ intros. inv H. constructor. constructor. rewrite H3; auto.
pattern f. apply Genv.find_funct_ptr_prop with fundef unit p b.
exact wt_p. exact H2.
rewrite H3. constructor.
@@ -992,10 +992,10 @@ Lemma wt_instr_inv:
f.(fn_code)!pc = Some i ->
exists env, wt_instr f env i /\ wt_regset env rs.
Proof.
- intros. inv H. exists env; split; auto.
- inv WT_FN. eauto.
+ intros. inv H. exists env; split; auto.
+ inv WT_FN. eauto.
Qed.
End SUBJECT_REDUCTION.
-
+
diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml
index 76288fb5..a5fa8cd7 100644
--- a/backend/Regalloc.ml
+++ b/backend/Regalloc.ml
@@ -234,7 +234,7 @@ let block_of_RTL_instr funsig tyenv = function
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])
+ xparmove res' (expand_regs tyenv [res])
[Xbranch s])
| RTL.Itailcall(sg, ros, args) ->
let args' = vlocs (loc_arguments sg) in
@@ -273,7 +273,7 @@ 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 =
+ let b_entrypoint =
xparmove (vlocs (loc_parameters f.RTL.fn_sig))
(expand_regs tyenv f.RTL.fn_params)
[Xbranch f.RTL.fn_entrypoint] in
@@ -465,7 +465,7 @@ let spill_costs f =
let c1 = st.cost + amount in
let c2 = if c1 >= 0 then c1 else max_int (* overflow *) in
st.cost <- c2
- end;
+ end;
st.usedefs <- st.usedefs + uses in
let charge_list amount uses vl =
@@ -624,7 +624,7 @@ let add_interfs_instr g instr live =
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)
+ 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)
@@ -655,7 +655,7 @@ let add_interfs_instr g instr live =
| Some mr ->
add_interfs_list_mreg g vargs mr;
add_interfs_list_mreg g vres mr)
- clob
+ clob
| _ -> ()
end
| Xbranch s ->
@@ -776,7 +776,7 @@ 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)
@@ -845,7 +845,7 @@ let rec trim count eqs =
if count <= 0 then [] else
match eqs with
| [] -> []
- | (v, t, date) :: eqs' ->
+ | (v, t, date) :: eqs' ->
if date <= !max_age
then (v, t, date + 1) :: trim (count - 1) eqs'
else []
@@ -882,7 +882,7 @@ let spill_instr tospill eqs instr =
| 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)],
+ (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
@@ -890,13 +890,13 @@ let spill_instr tospill eqs instr =
(Xreload(arg1, res) :: c1 @ [Xop(op, res :: argl', res)],
kill res eqs2)
| true, true ->
- let tmp = new_temp (typeof res) in
+ 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
+ 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
@@ -1047,7 +1047,7 @@ let transl_instr alloc instr k =
if rarg1 = rres then
LTL.Lop(op, rargs, rres) :: k
else
- LTL.Lop(Omove, [rarg1], rres) ::
+ LTL.Lop(Omove, [rarg1], rres) ::
LTL.Lop(op, rres :: rargl, rres) :: k
end
| Xload(chunk, addr, args, dst) ->
@@ -1079,7 +1079,7 @@ 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
+ LTL.fn_code = PTree.map1 (transl_block alloc) fn.fn_code
}
@@ -1113,7 +1113,7 @@ and more_rounds f ts count =
fprintf !pp "--- Remain to be spilled:\n";
VSet.iter (fun v -> fprintf !pp "%a " PrintXTL.var v) ts';
fprintf !pp "\n\n"
- end;
+ end;
more_rounds f (VSet.union ts ts') (count + 1)
end
@@ -1148,7 +1148,7 @@ let regalloc f =
| 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 ");
+ 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/Registers.v b/backend/Registers.v
index 20532e8c..cfe8427b 100644
--- a/backend/Registers.v
+++ b/backend/Registers.v
@@ -81,7 +81,7 @@ Lemma set_reg_lessdef:
forall r v1 v2 rs1 rs2,
Val.lessdef v1 v2 -> regs_lessdef rs1 rs2 -> regs_lessdef (rs1#r <- v1) (rs2#r <- v2).
Proof.
- intros; red; intros. repeat rewrite Regmap.gsspec.
+ intros; red; intros. repeat rewrite Regmap.gsspec.
destruct (peq r0 r); auto.
Qed.
diff --git a/backend/Renumber.v b/backend/Renumber.v
index 0a2c2f12..10f58251 100644
--- a/backend/Renumber.v
+++ b/backend/Renumber.v
@@ -22,7 +22,7 @@ Require Import RTL.
enumeration of the nodes of the control-flow graph. This property
can be guaranteed when generating the CFG (module [RTLgen]), but
is, however, invalidated by further RTL optimization passes such as
- [Inlining].
+ [Inlining].
In this module, we renumber the nodes of RTL control-flow graphs
to restore the postorder property given above. In passing,
diff --git a/backend/Renumberproof.v b/backend/Renumberproof.v
index 33d6aafa..f4d9cca3 100644
--- a/backend/Renumberproof.v
+++ b/backend/Renumberproof.v
@@ -67,7 +67,7 @@ Lemma find_function_translated:
find_function ge ros rs = Some fd ->
find_function tge ros rs = Some (transf_fundef fd).
Proof.
- unfold find_function; intros. destruct ros as [r|id].
+ unfold find_function; intros. destruct ros as [r|id].
eapply functions_translated; eauto.
rewrite symbols_preserved. destruct (Genv.find_symbol ge id); try congruence.
eapply function_ptr_translated; eauto.
@@ -87,18 +87,18 @@ Lemma renum_cfg_nodes:
Proof.
set (P := fun (c c': code) =>
forall x y i, c!x = Some i -> f!x = Some y -> c'!y = Some(renum_instr f i)).
- intros c0. change (P c0 (renum_cfg f c0)). unfold renum_cfg.
+ intros c0. change (P c0 (renum_cfg f c0)). unfold renum_cfg.
apply PTree_Properties.fold_rec; unfold P; intros.
(* extensionality *)
- eapply H0; eauto. rewrite H; auto.
+ eapply H0; eauto. rewrite H; auto.
(* base *)
rewrite PTree.gempty in H; congruence.
(* induction *)
- rewrite PTree.gsspec in H2. unfold renum_node. destruct (peq x k).
- inv H2. rewrite H3. apply PTree.gss.
- destruct f!k as [y'|] eqn:?.
- rewrite PTree.gso. eauto. red; intros; subst y'. elim n. eapply f_inj; eauto.
- eauto.
+ rewrite PTree.gsspec in H2. unfold renum_node. destruct (peq x k).
+ inv H2. rewrite H3. apply PTree.gss.
+ destruct f!k as [y'|] eqn:?.
+ rewrite PTree.gso. eauto. red; intros; subst y'. elim n. eapply f_inj; eauto.
+ eauto.
Qed.
End RENUMBER.
@@ -113,9 +113,9 @@ Lemma transf_function_at:
reach f pc ->
(transf_function f).(fn_code)!(renum_pc (pnum f) pc) = Some(renum_instr (pnum f) i).
Proof.
- intros.
+ intros.
destruct (postorder_correct (successors_map f) f.(fn_entrypoint)) as [A B].
- fold (pnum f) in *.
+ fold (pnum f) in *.
unfold renum_pc. destruct (pnum f)! pc as [pc'|] eqn:?.
simpl. eapply renum_cfg_nodes; eauto.
elim (B pc); auto. unfold successors_map. rewrite PTree.gmap1. rewrite H. simpl. congruence.
@@ -132,10 +132,10 @@ Lemma reach_succ:
f.(fn_code)!pc = Some i -> In s (successors_instr i) ->
reach f pc -> reach f s.
Proof.
- unfold reach; intros. econstructor; eauto.
- unfold successors_map. rewrite PTree.gmap1. rewrite H. auto.
+ unfold reach; intros. econstructor; eauto.
+ unfold successors_map. rewrite PTree.gmap1. rewrite H. auto.
Qed.
-
+
Inductive match_frames: RTL.stackframe -> RTL.stackframe -> Prop :=
| match_frames_intro: forall res f sp pc rs
(REACH: reach f pc),
@@ -164,23 +164,23 @@ Lemma step_simulation:
Proof.
induction 1; intros S1' MS; inv MS; try TR_AT.
(* nop *)
- econstructor; split. eapply exec_Inop; eauto.
- constructor; auto. eapply reach_succ; eauto. simpl; auto.
+ econstructor; split. eapply exec_Inop; eauto.
+ constructor; auto. eapply reach_succ; eauto. simpl; auto.
(* op *)
econstructor; split.
eapply exec_Iop; eauto.
- instantiate (1 := v). rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved.
+ instantiate (1 := v). rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved.
constructor; auto. eapply reach_succ; eauto. simpl; auto.
(* load *)
econstructor; split.
assert (eval_addressing tge sp addr rs ## args = Some a).
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
eapply exec_Iload; eauto.
constructor; auto. eapply reach_succ; eauto. simpl; auto.
(* store *)
econstructor; split.
assert (eval_addressing tge sp addr rs ## args = Some a).
- rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
eapply exec_Istore; eauto.
constructor; auto. eapply reach_succ; eauto. simpl; auto.
(* call *)
@@ -204,23 +204,23 @@ Proof.
constructor; auto. eapply reach_succ; eauto. simpl; auto.
(* cond *)
econstructor; split.
- eapply exec_Icond; eauto.
+ eapply exec_Icond; eauto.
replace (if b then renum_pc (pnum f) ifso else renum_pc (pnum f) ifnot)
with (renum_pc (pnum f) (if b then ifso else ifnot)).
- constructor; auto. eapply reach_succ; eauto. simpl. destruct b; auto.
+ constructor; auto. eapply reach_succ; eauto. simpl. destruct b; auto.
destruct b; auto.
(* jumptbl *)
econstructor; split.
- eapply exec_Ijumptable; eauto. rewrite list_nth_z_map. rewrite H1. simpl; eauto.
- constructor; auto. eapply reach_succ; eauto. simpl. eapply list_nth_z_in; eauto.
+ eapply exec_Ijumptable; eauto. rewrite list_nth_z_map. rewrite H1. simpl; eauto.
+ constructor; auto. eapply reach_succ; eauto. simpl. eapply list_nth_z_in; eauto.
(* return *)
econstructor; split.
- eapply exec_Ireturn; eauto.
+ eapply exec_Ireturn; eauto.
constructor; auto.
(* internal function *)
simpl. econstructor; split.
- eapply exec_function_internal; eauto.
- constructor; auto. unfold reach. constructor.
+ eapply exec_function_internal; eauto.
+ constructor; auto. unfold reach. constructor.
(* external function *)
econstructor; split.
eapply exec_function_external; eauto.
@@ -229,8 +229,8 @@ Proof.
constructor; auto.
(* return *)
inv STACKS. inv H1.
- econstructor; split.
- eapply exec_return; eauto.
+ econstructor; split.
+ eapply exec_return; eauto.
constructor; auto.
Qed.
@@ -239,10 +239,10 @@ Lemma transf_initial_states:
exists S2, RTL.initial_state tprog S2 /\ match_states S1 S2.
Proof.
intros. inv H. econstructor; split.
- econstructor.
- eapply Genv.init_mem_transf; eauto.
- simpl. rewrite symbols_preserved. eauto.
- eapply function_ptr_translated; eauto.
+ econstructor.
+ eapply Genv.init_mem_transf; eauto.
+ simpl. rewrite symbols_preserved. eauto.
+ eapply function_ptr_translated; eauto.
rewrite <- H3; apply sig_preserved.
constructor. constructor.
Qed.
@@ -260,14 +260,14 @@ Proof.
eexact public_preserved.
eexact transf_initial_states.
eexact transf_final_states.
- exact step_simulation.
+ exact step_simulation.
Qed.
End PRESERVATION.
-
-
+
+
diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v
index d4bd4f5c..ffe607e4 100644
--- a/backend/SelectDivproof.v
+++ b/backend/SelectDivproof.v
@@ -50,17 +50,17 @@ Lemma Zdiv_mul_pos:
Zdiv n d = Zdiv (m * n) (two_p (N + l)).
Proof.
intros m l l_pos [LO HI] n RANGE.
- exploit (Z_div_mod_eq n d). auto.
+ exploit (Z_div_mod_eq n d). auto.
set (q := n / d).
set (r := n mod d).
intro EUCL.
assert (0 <= r <= d - 1).
unfold r. generalize (Z_mod_lt n d d_pos). omega.
- assert (0 <= m).
+ assert (0 <= m).
apply Zmult_le_0_reg_r with d. auto.
- exploit (two_p_gt_ZERO (N + l)). omega. omega.
+ exploit (two_p_gt_ZERO (N + l)). omega. omega.
set (k := m * d - two_p (N + l)).
- assert (0 <= k <= two_p l).
+ assert (0 <= k <= two_p l).
unfold k; omega.
assert ((m * n - two_p (N + l) * q) * d = k * n + two_p (N + l) * r).
unfold k. rewrite EUCL. ring.
@@ -70,14 +70,14 @@ Proof.
apply Zle_trans with (two_p l * n).
apply Zmult_le_compat_r. omega. omega.
replace (N + l) with (l + N) by omega.
- rewrite two_p_is_exp.
+ rewrite two_p_is_exp.
replace (two_p l * two_p N - two_p l)
with (two_p l * (two_p N - 1))
by ring.
apply Zmult_le_compat_l. omega. exploit (two_p_gt_ZERO l). omega. omega.
omega. omega.
assert (0 <= two_p (N + l) * r).
- apply Zmult_le_0_compat.
+ apply Zmult_le_0_compat.
exploit (two_p_gt_ZERO (N + l)). omega. omega.
omega.
assert (two_p (N + l) * r <= two_p (N + l) * d - two_p (N + l)).
@@ -87,23 +87,23 @@ Proof.
omega.
exploit (two_p_gt_ZERO (N + l)). omega. omega.
assert (0 <= m * n - two_p (N + l) * q).
- apply Zmult_le_reg_r with d. auto.
- replace (0 * d) with 0 by ring. rewrite H2. omega.
+ apply Zmult_le_reg_r with d. auto.
+ replace (0 * d) with 0 by ring. rewrite H2. omega.
assert (m * n - two_p (N + l) * q < two_p (N + l)).
apply Zmult_lt_reg_r with d. omega.
- rewrite H2.
+ rewrite H2.
apply Zle_lt_trans with (two_p (N + l) * d - two_p l).
- omega.
+ omega.
exploit (two_p_gt_ZERO l). omega. omega.
- symmetry. apply Zdiv_unique with (m * n - two_p (N + l) * q).
+ symmetry. apply Zdiv_unique with (m * n - two_p (N + l) * q).
ring. omega.
Qed.
Lemma Zdiv_unique_2:
forall x y q, y > 0 -> 0 < y * q - x <= y -> Zdiv x y = q - 1.
Proof.
- intros. apply Zdiv_unique with (x - (q - 1) * y). ring.
- replace ((q - 1) * y) with (y * q - y) by ring. omega.
+ intros. apply Zdiv_unique with (x - (q - 1) * y). ring.
+ replace ((q - 1) * y) with (y * q - y) by ring. omega.
Qed.
Lemma Zdiv_mul_opp:
@@ -116,29 +116,29 @@ Lemma Zdiv_mul_opp:
Proof.
intros m l l_pos [LO HI] n RANGE.
replace (m * (-n)) with (- (m * n)) by ring.
- exploit (Z_div_mod_eq n d). auto.
+ exploit (Z_div_mod_eq n d). auto.
set (q := n / d).
set (r := n mod d).
intro EUCL.
assert (0 <= r <= d - 1).
unfold r. generalize (Z_mod_lt n d d_pos). omega.
- assert (0 <= m).
+ assert (0 <= m).
apply Zmult_le_0_reg_r with d. auto.
exploit (two_p_gt_ZERO (N + l)). omega. omega.
cut (Zdiv (- (m * n)) (two_p (N + l)) = -q - 1).
omega.
- apply Zdiv_unique_2.
+ apply Zdiv_unique_2.
apply two_p_gt_ZERO. omega.
replace (two_p (N + l) * - q - - (m * n))
with (m * n - two_p (N + l) * q)
by ring.
set (k := m * d - two_p (N + l)).
- assert (0 < k <= two_p l).
+ assert (0 < k <= two_p l).
unfold k; omega.
assert ((m * n - two_p (N + l) * q) * d = k * n + two_p (N + l) * r).
unfold k. rewrite EUCL. ring.
split.
- apply Zmult_lt_reg_r with d. omega.
+ apply Zmult_lt_reg_r with d. omega.
replace (0 * d) with 0 by omega.
rewrite H2.
assert (0 < k * n). apply Zmult_lt_0_compat; omega.
@@ -146,10 +146,10 @@ Proof.
apply Zmult_le_0_compat. exploit (two_p_gt_ZERO (N + l)); omega. omega.
omega.
apply Zmult_le_reg_r with d. omega.
- rewrite H2.
+ rewrite H2.
assert (k * n <= two_p (N + l)).
- rewrite Zplus_comm. rewrite two_p_is_exp; try omega.
- apply Zle_trans with (two_p l * n). apply Zmult_le_compat_r. omega. omega.
+ rewrite Zplus_comm. rewrite two_p_is_exp; try omega.
+ apply Zle_trans with (two_p l * n). apply Zmult_le_compat_r. omega. omega.
apply Zmult_le_compat_l. omega. exploit (two_p_gt_ZERO l). omega. omega.
assert (two_p (N + l) * r <= two_p (N + l) * d - two_p (N + l)).
replace (two_p (N + l) * d - two_p (N + l))
@@ -170,12 +170,12 @@ Lemma Zquot_mul:
Z.quot n d = Zdiv (m * n) (two_p (N + l)) + (if zlt n 0 then 1 else 0).
Proof.
intros. destruct (zlt n 0).
- exploit (Zdiv_mul_opp m l H H0 (-n)). omega.
+ exploit (Zdiv_mul_opp m l H H0 (-n)). omega.
replace (- - n) with n by ring.
replace (Z.quot n d) with (- Z.quot (-n) d).
rewrite Zquot_Zdiv_pos by omega. omega.
rewrite Z.quot_opp_l by omega. ring.
- rewrite Zplus_0_r. rewrite Zquot_Zdiv_pos by omega.
+ rewrite Zplus_0_r. rewrite Zquot_Zdiv_pos by omega.
apply Zdiv_mul_pos; omega.
Qed.
@@ -202,13 +202,13 @@ Proof with (try discriminate).
destruct (zlt m Int.modulus)...
destruct (zle 0 p)...
destruct (zlt p 32)...
- simpl in EQ. inv EQ.
- split. auto. split. auto. intros.
- replace (32 + p') with (31 + (p' + 1)) by omega.
+ simpl in EQ. inv EQ.
+ split. auto. split. auto. intros.
+ replace (32 + p') with (31 + (p' + 1)) by omega.
apply Zquot_mul; try omega.
replace (31 + (p' + 1)) with (32 + p') by omega. omega.
- change (Int.min_signed <= n < Int.half_modulus).
- unfold Int.max_signed in H. omega.
+ change (Int.min_signed <= n < Int.half_modulus).
+ unfold Int.max_signed in H. omega.
Qed.
Lemma divu_mul_params_sound:
@@ -230,7 +230,7 @@ Proof with (try discriminate).
destruct (zlt m Int.modulus)...
destruct (zle 0 p)...
destruct (zlt p 32)...
- simpl in EQ. inv EQ.
+ simpl in EQ. inv EQ.
split. auto. split. auto. intros.
apply Zdiv_mul_pos; try omega. assumption.
Qed.
@@ -245,23 +245,23 @@ Proof.
intros. set (n := Int.signed x). set (d := Int.signed y) in *.
exploit divs_mul_params_sound; eauto. intros (A & B & C).
split. auto. split. auto.
- unfold Int.divs. fold n; fold d. rewrite C by (apply Int.signed_range).
- rewrite two_p_is_exp by omega. rewrite <- Zdiv_Zdiv.
+ unfold Int.divs. fold n; fold d. rewrite C by (apply Int.signed_range).
+ rewrite two_p_is_exp by omega. rewrite <- Zdiv_Zdiv.
rewrite Int.shru_lt_zero. unfold Int.add. apply Int.eqm_samerepr. apply Int.eqm_add.
- rewrite Int.shr_div_two_p. apply Int.eqm_unsigned_repr_r. apply Int.eqm_refl2.
+ rewrite Int.shr_div_two_p. apply Int.eqm_unsigned_repr_r. apply Int.eqm_refl2.
rewrite Int.unsigned_repr. f_equal.
rewrite Int.signed_repr. rewrite Int.modulus_power. f_equal. ring.
- cut (Int.min_signed <= n * m / Int.modulus < Int.half_modulus).
- unfold Int.max_signed; omega.
- apply Zdiv_interval_1. generalize Int.min_signed_neg; omega. apply Int.half_modulus_pos.
+ cut (Int.min_signed <= n * m / Int.modulus < Int.half_modulus).
+ unfold Int.max_signed; omega.
+ apply Zdiv_interval_1. generalize Int.min_signed_neg; omega. apply Int.half_modulus_pos.
apply Int.modulus_pos.
split. apply Zle_trans with (Int.min_signed * m). apply Zmult_le_compat_l_neg. omega. generalize Int.min_signed_neg; omega.
apply Zmult_le_compat_r. unfold n; generalize (Int.signed_range x); tauto. tauto.
- apply Zle_lt_trans with (Int.half_modulus * m).
+ apply Zle_lt_trans with (Int.half_modulus * m).
apply Zmult_le_compat_r. generalize (Int.signed_range x); unfold n, Int.max_signed; omega. tauto.
- apply Zmult_lt_compat_l. generalize Int.half_modulus_pos; omega. tauto.
+ apply Zmult_lt_compat_l. generalize Int.half_modulus_pos; omega. tauto.
assert (32 < Int.max_unsigned) by (compute; auto). omega.
- unfold Int.lt; fold n. rewrite Int.signed_zero. destruct (zlt n 0); apply Int.eqm_unsigned_repr.
+ unfold Int.lt; fold n. rewrite Int.signed_zero. destruct (zlt n 0); apply Int.eqm_unsigned_repr.
apply two_p_gt_ZERO. omega.
apply two_p_gt_ZERO. omega.
Qed.
@@ -274,8 +274,8 @@ Theorem divs_mul_shift_1:
Int.divs x y = Int.add (Int.shr (Int.mulhs x (Int.repr m)) (Int.repr p))
(Int.shru x (Int.repr 31)).
Proof.
- intros. exploit divs_mul_shift_gen; eauto. instantiate (1 := x).
- intros (A & B & C). split. auto. rewrite C.
+ intros. exploit divs_mul_shift_gen; eauto. instantiate (1 := x).
+ intros (A & B & C). split. auto. rewrite C.
unfold Int.mulhs. rewrite Int.signed_repr. auto.
generalize Int.min_signed_neg; unfold Int.max_signed; omega.
Qed.
@@ -288,17 +288,17 @@ Theorem divs_mul_shift_2:
Int.divs x y = Int.add (Int.shr (Int.add (Int.mulhs x (Int.repr m)) x) (Int.repr p))
(Int.shru x (Int.repr 31)).
Proof.
- intros. exploit divs_mul_shift_gen; eauto. instantiate (1 := x).
+ intros. exploit divs_mul_shift_gen; eauto. instantiate (1 := x).
intros (A & B & C). split. auto. rewrite C. f_equal. f_equal.
rewrite Int.add_signed. unfold Int.mulhs. set (n := Int.signed x).
transitivity (Int.repr (n * (m - Int.modulus) / Int.modulus + n)).
- f_equal.
+ f_equal.
replace (n * (m - Int.modulus)) with (n * m + (-n) * Int.modulus) by ring.
- rewrite Z_div_plus. ring. apply Int.modulus_pos.
- apply Int.eqm_samerepr. apply Int.eqm_add; auto with ints.
- apply Int.eqm_sym. eapply Int.eqm_trans. apply Int.eqm_signed_unsigned.
- apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl2. f_equal. f_equal.
- rewrite Int.signed_repr_eq. rewrite Zmod_small by assumption.
+ rewrite Z_div_plus. ring. apply Int.modulus_pos.
+ apply Int.eqm_samerepr. apply Int.eqm_add; auto with ints.
+ apply Int.eqm_sym. eapply Int.eqm_trans. apply Int.eqm_signed_unsigned.
+ apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl2. f_equal. f_equal.
+ rewrite Int.signed_repr_eq. rewrite Zmod_small by assumption.
apply zlt_false. omega.
Qed.
@@ -309,19 +309,19 @@ Theorem divu_mul_shift:
Int.divu x y = Int.shru (Int.mulhu x (Int.repr m)) (Int.repr p).
Proof.
intros. exploit divu_mul_params_sound; eauto. intros (A & B & C).
- split. auto.
- rewrite Int.shru_div_two_p. rewrite Int.unsigned_repr.
+ split. auto.
+ rewrite Int.shru_div_two_p. rewrite Int.unsigned_repr.
unfold Int.divu, Int.mulhu. f_equal. rewrite C by apply Int.unsigned_range.
rewrite two_p_is_exp by omega. rewrite <- Zdiv_Zdiv by (apply two_p_gt_ZERO; omega).
- f_equal. rewrite (Int.unsigned_repr m).
+ f_equal. rewrite (Int.unsigned_repr m).
rewrite Int.unsigned_repr. f_equal. ring.
cut (0 <= Int.unsigned x * m / Int.modulus < Int.modulus).
unfold Int.max_unsigned; omega.
apply Zdiv_interval_1. omega. compute; auto. compute; auto.
- split. simpl. apply Z.mul_nonneg_nonneg. generalize (Int.unsigned_range x); omega. omega.
+ split. simpl. apply Z.mul_nonneg_nonneg. generalize (Int.unsigned_range x); omega. omega.
apply Zle_lt_trans with (Int.modulus * m).
- apply Zmult_le_compat_r. generalize (Int.unsigned_range x); omega. omega.
- apply Zmult_lt_compat_l. compute; auto. omega.
+ apply Zmult_le_compat_r. generalize (Int.unsigned_range x); omega. omega.
+ apply Zmult_lt_compat_l. compute; auto. omega.
unfold Int.max_unsigned; omega.
assert (32 < Int.max_unsigned) by (compute; auto). omega.
Qed.
@@ -347,10 +347,10 @@ Proof.
(Vint (Int.mulhu x (Int.repr M)))).
{ EvalOp. econstructor. econstructor; eauto. econstructor. EvalOp. simpl; reflexivity. constructor.
auto. }
- exploit eval_shruimm. eexact H1. instantiate (1 := Int.repr p).
- intros [v [P Q]]. simpl in Q.
- replace (Int.ltu (Int.repr p) Int.iwordsize) with true in Q.
- inv Q. rewrite B. auto.
+ exploit eval_shruimm. eexact H1. instantiate (1 := Int.repr p).
+ intros [v [P Q]]. simpl in Q.
+ replace (Int.ltu (Int.repr p) Int.iwordsize) with true in Q.
+ inv Q. rewrite B. auto.
unfold Int.ltu. rewrite Int.unsigned_repr. rewrite zlt_true; auto. tauto.
assert (32 < Int.max_unsigned) by (compute; auto). omega.
Qed.
@@ -363,17 +363,17 @@ Theorem eval_divuimm:
Proof.
unfold divuimm; intros. generalize H0; intros DIV.
destruct x; simpl in DIV; try discriminate.
- destruct (Int.eq n2 Int.zero) eqn:Z2; inv DIV.
+ destruct (Int.eq n2 Int.zero) eqn:Z2; inv DIV.
destruct (Int.is_power2 n2) as [l | ] eqn:P2.
-- erewrite Int.divu_pow2 by eauto.
- replace (Vint (Int.shru i l)) with (Val.shru (Vint i) (Vint l)).
+- erewrite Int.divu_pow2 by eauto.
+ replace (Vint (Int.shru i l)) with (Val.shru (Vint i) (Vint l)).
apply eval_shruimm; auto.
- simpl. erewrite Int.is_power2_range; eauto.
+ simpl. erewrite Int.is_power2_range; eauto.
- destruct (Compopts.optim_for_size tt).
+ eapply eval_divu_base; eauto. EvalOp.
+ destruct (divu_mul_params (Int.unsigned n2)) as [[p M] | ] eqn:PARAMS.
* exists (Vint (Int.divu i n2)); split; auto.
- econstructor; eauto. eapply eval_divu_mul; eauto.
+ econstructor; eauto. eapply eval_divu_mul; eauto.
* eapply eval_divu_base; eauto. EvalOp.
Qed.
@@ -386,7 +386,7 @@ Theorem eval_divu:
Proof.
unfold divu; intros until b. destruct (divu_match b); intros.
- inv H0. inv H5. simpl in H7. inv H7. eapply eval_divuimm; eauto.
-- eapply eval_divu_base; eauto.
+- eapply eval_divu_base; eauto.
Qed.
Lemma eval_mod_from_div:
@@ -395,8 +395,8 @@ Lemma eval_mod_from_div:
nth_error le O = Some (Vint x) ->
eval_expr ge sp e m le (mod_from_div a n) (Vint (Int.sub x (Int.mul y n))).
Proof.
- unfold mod_from_div; intros.
- exploit eval_mulimm; eauto. instantiate (1 := n). intros [v [A B]].
+ unfold mod_from_div; intros.
+ exploit eval_mulimm; eauto. instantiate (1 := n). intros [v [A B]].
simpl in B. inv B. EvalOp.
Qed.
@@ -408,7 +408,7 @@ Theorem eval_moduimm:
Proof.
unfold moduimm; intros. generalize H0; intros MOD.
destruct x; simpl in MOD; try discriminate.
- destruct (Int.eq n2 Int.zero) eqn:Z2; inv MOD.
+ destruct (Int.eq n2 Int.zero) eqn:Z2; inv MOD.
destruct (Int.is_power2 n2) as [l | ] eqn:P2.
- erewrite Int.modu_and by eauto.
change (Vint (Int.and i (Int.sub n2 Int.one)))
@@ -417,10 +417,10 @@ Proof.
- destruct (Compopts.optim_for_size tt).
+ eapply eval_modu_base; eauto. EvalOp.
+ destruct (divu_mul_params (Int.unsigned n2)) as [[p M] | ] eqn:PARAMS.
- * econstructor; split.
- econstructor; eauto. eapply eval_mod_from_div.
- eapply eval_divu_mul; eauto. simpl; eauto. simpl; eauto.
- rewrite Int.modu_divu. auto.
+ * econstructor; split.
+ econstructor; eauto. eapply eval_mod_from_div.
+ eapply eval_divu_mul; eauto. simpl; eauto. simpl; eauto.
+ rewrite Int.modu_divu. auto.
red; intros; subst n2; discriminate.
* eapply eval_modu_base; eauto. EvalOp.
Qed.
@@ -434,7 +434,7 @@ Theorem eval_modu:
Proof.
unfold modu; intros until b. destruct (modu_match b); intros.
- inv H0. inv H5. simpl in H7. inv H7. eapply eval_moduimm; eauto.
-- eapply eval_modu_base; eauto.
+- eapply eval_modu_base; eauto.
Qed.
Lemma eval_divs_mul:
@@ -451,10 +451,10 @@ Proof.
(Vint (Int.mulhs x (Int.repr M)))).
{ EvalOp. econstructor. eauto. econstructor. EvalOp. simpl; reflexivity. constructor.
auto. }
- exploit eval_shruimm. eexact V. instantiate (1 := Int.repr (Int.zwordsize - 1)).
- intros [v1 [Y LD]]. simpl in LD.
- change (Int.ltu (Int.repr 31) Int.iwordsize) with true in LD.
- simpl in LD. inv LD.
+ exploit eval_shruimm. eexact V. instantiate (1 := Int.repr (Int.zwordsize - 1)).
+ intros [v1 [Y LD]]. simpl in LD.
+ change (Int.ltu (Int.repr 31) Int.iwordsize) with true in LD.
+ simpl in LD. inv LD.
assert (RANGE: 0 <= p < 32 -> Int.ltu (Int.repr p) Int.iwordsize = true).
{ intros. unfold Int.ltu. rewrite Int.unsigned_repr. rewrite zlt_true by tauto. auto.
assert (32 < Int.max_unsigned) by (compute; auto). omega. }
@@ -463,15 +463,15 @@ Proof.
exploit eval_shrimm. eexact X. instantiate (1 := Int.repr p). intros [v1 [Z LD]].
simpl in LD. rewrite RANGE in LD by auto. inv LD.
exploit eval_add. eexact Z. eexact Y. intros [v1 [W LD]].
- simpl in LD. inv LD.
+ simpl in LD. inv LD.
rewrite B. exact W.
- exploit (divs_mul_shift_2 x); eauto. intros [A B].
- exploit eval_add. eexact X. eexact V. intros [v1 [Z LD]].
- simpl in LD. inv LD.
+ exploit eval_add. eexact X. eexact V. intros [v1 [Z LD]].
+ simpl in LD. inv LD.
exploit eval_shrimm. eexact Z. instantiate (1 := Int.repr p). intros [v1 [U LD]].
simpl in LD. rewrite RANGE in LD by auto. inv LD.
exploit eval_add. eexact U. eexact Y. intros [v1 [W LD]].
- simpl in LD. inv LD.
+ simpl in LD. inv LD.
rewrite B. exact W.
Qed.
@@ -484,7 +484,7 @@ Proof.
unfold divsimm; intros. generalize H0; intros DIV.
destruct x; simpl in DIV; try discriminate.
destruct (Int.eq n2 Int.zero
- || Int.eq i (Int.repr Int.min_signed) && Int.eq n2 Int.mone) eqn:Z2; inv DIV.
+ || Int.eq i (Int.repr Int.min_signed) && Int.eq n2 Int.mone) eqn:Z2; inv DIV.
destruct (Int.is_power2 n2) as [l | ] eqn:P2.
- destruct (Int.ltu l (Int.repr 31)) eqn:LT31.
+ eapply eval_shrximm; eauto. eapply Val.divs_pow2; eauto.
@@ -493,7 +493,7 @@ Proof.
+ eapply eval_divs_base; eauto. EvalOp.
+ destruct (divs_mul_params (Int.signed n2)) as [[p M] | ] eqn:PARAMS.
* exists (Vint (Int.divs i n2)); split; auto.
- econstructor; eauto. eapply eval_divs_mul; eauto.
+ econstructor; eauto. eapply eval_divs_mul; eauto.
* eapply eval_divs_base; eauto. EvalOp.
Qed.
@@ -506,7 +506,7 @@ Theorem eval_divs:
Proof.
unfold divs; intros until b. destruct (divs_match b); intros.
- inv H0. inv H5. simpl in H7. inv H7. eapply eval_divsimm; eauto.
-- eapply eval_divs_base; eauto.
+- eapply eval_divs_base; eauto.
Qed.
Theorem eval_modsimm:
@@ -515,25 +515,25 @@ Theorem eval_modsimm:
Val.mods x (Vint n2) = Some z ->
exists v, eval_expr ge sp e m le (modsimm e1 n2) v /\ Val.lessdef z v.
Proof.
- unfold modsimm; intros.
+ unfold modsimm; intros.
exploit Val.mods_divs; eauto. intros [y [A B]].
generalize A; intros DIV.
destruct x; simpl in DIV; try discriminate.
destruct (Int.eq n2 Int.zero
- || Int.eq i (Int.repr Int.min_signed) && Int.eq n2 Int.mone) eqn:Z2; inv DIV.
+ || Int.eq i (Int.repr Int.min_signed) && Int.eq n2 Int.mone) eqn:Z2; inv DIV.
destruct (Int.is_power2 n2) as [l | ] eqn:P2.
- destruct (Int.ltu l (Int.repr 31)) eqn:LT31.
+ exploit (eval_shrximm ge sp e m (Vint i :: le) (Eletvar O)).
- constructor. simpl; eauto. eapply Val.divs_pow2; eauto.
- intros [v1 [X LD]]. inv LD.
- econstructor; split. econstructor. eauto.
- apply eval_mod_from_div. eexact X. simpl; eauto.
+ constructor. simpl; eauto. eapply Val.divs_pow2; eauto.
+ intros [v1 [X LD]]. inv LD.
+ econstructor; split. econstructor. eauto.
+ apply eval_mod_from_div. eexact X. simpl; eauto.
simpl. auto.
+ eapply eval_mods_base; eauto. EvalOp.
- destruct (Compopts.optim_for_size tt).
+ eapply eval_mods_base; eauto. EvalOp.
+ destruct (divs_mul_params (Int.signed n2)) as [[p M] | ] eqn:PARAMS.
- * econstructor; split.
+ * econstructor; split.
econstructor. eauto. apply eval_mod_from_div with (x := i); auto.
eapply eval_divs_mul with (x := i); eauto.
simpl. auto.
@@ -549,7 +549,7 @@ Theorem eval_mods:
Proof.
unfold mods; intros until b. destruct (mods_match b); intros.
- inv H0. inv H5. simpl in H7. inv H7. eapply eval_modsimm; eauto.
-- eapply eval_mods_base; eauto.
+- eapply eval_mods_base; eauto.
Qed.
(** * Floating-point division *)
@@ -563,10 +563,10 @@ Proof.
intros until y. unfold divf. destruct (divf_match b); intros.
- unfold divfimm. destruct (Float.exact_inverse n2) as [n2' | ] eqn:EINV.
+ inv H0. inv H4. simpl in H6. inv H6. econstructor; split.
- EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
- simpl; eauto.
- destruct x; simpl; auto. erewrite Float.div_mul_inverse; eauto.
- + TrivialExists.
+ EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
+ simpl; eauto.
+ destruct x; simpl; auto. erewrite Float.div_mul_inverse; eauto.
+ + TrivialExists.
- TrivialExists.
Qed.
@@ -579,10 +579,10 @@ Proof.
intros until y. unfold divfs. destruct (divfs_match b); intros.
- unfold divfsimm. destruct (Float32.exact_inverse n2) as [n2' | ] eqn:EINV.
+ inv H0. inv H4. simpl in H6. inv H6. econstructor; split.
- EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
- simpl; eauto.
- destruct x; simpl; auto. erewrite Float32.div_mul_inverse; eauto.
- + TrivialExists.
+ EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
+ simpl; eauto.
+ destruct x; simpl; auto. erewrite Float32.div_mul_inverse; eauto.
+ + TrivialExists.
- TrivialExists.
Qed.
diff --git a/backend/SelectLongproof.v b/backend/SelectLongproof.v
index cdfb1107..35d53215 100644
--- a/backend/SelectLongproof.v
+++ b/backend/SelectLongproof.v
@@ -131,7 +131,7 @@ Remark eval_builtin_1:
builtin_implements 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.
+ intros. econstructor. econstructor. eauto. constructor. apply H0.
Qed.
Remark eval_builtin_2:
@@ -181,10 +181,10 @@ 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]].
+ 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).
+ 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.
@@ -202,9 +202,9 @@ Lemma eval_splitlong_strict:
eval_expr ge sp e m le (f a1 a2) v) ->
eval_expr ge sp e m le (splitlong a f) v.
Proof.
- intros until v.
+ intros until v.
unfold splitlong. case (splitlong_match a); intros.
-- InvEval. destruct v1; simpl in H; try discriminate. destruct v0; inv H.
+- InvEval. destruct v1; simpl in H; try discriminate. destruct v0; inv H.
apply H0. rewrite Int64.hi_ofwords; auto. rewrite Int64.lo_ofwords; auto.
- EvalOp. apply H0; EvalOp.
Qed.
@@ -236,10 +236,10 @@ Proof.
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)
+- 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.
+ EvalOp. EvalOp. EvalOp. EvalOp.
intros [v [A B]].
exists v; split.
econstructor; eauto.
@@ -247,7 +247,7 @@ Proof.
destruct v0; try (rewrite UNDEF; auto).
destruct vb; try (rewrite UNDEF; auto).
erewrite B; simpl; eauto. rewrite Int64.ofwords_recompose. auto.
-- InvEval. subst vb.
+- InvEval. subst vb.
exploit (EXEC (va :: le)
(Eop Ohighlong (Eletvar 0 ::: Enil)) (Eop Olowlong (Eletvar 0 ::: Enil))
(lift h2) (lift l2)).
@@ -256,15 +256,15 @@ Proof.
exists v; split.
econstructor; eauto.
destruct va; try (rewrite UNDEF; auto).
- destruct v1; simpl in *; 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.
+ 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.
+ 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.
@@ -287,13 +287,13 @@ Proof.
intros. destruct v1; simpl in H; try discriminate. destruct v2; inv H.
rewrite Int64.hi_ofwords; rewrite Int64.lo_ofwords; auto.
}
- intros until v.
+ intros until v.
unfold splitlong2. case (splitlong2_match a b); intros.
-- InvEval. exploit INV. eexact H. intros [EQ1 EQ2]. exploit INV. eexact H0. intros [EQ3 EQ4].
+- InvEval. exploit INV. eexact H. intros [EQ1 EQ2]. exploit INV. eexact H0. intros [EQ3 EQ4].
subst. auto.
-- InvEval. exploit INV; eauto. intros [EQ1 EQ2]. subst.
- econstructor. eauto. apply H1; EvalOp.
-- InvEval. exploit INV; eauto. intros [EQ1 EQ2]. subst.
+- InvEval. exploit INV; eauto. intros [EQ1 EQ2]. subst.
+ econstructor. eauto. apply H1; EvalOp.
+- InvEval. exploit INV; eauto. intros [EQ1 EQ2]. subst.
econstructor. eauto. apply H1; EvalOp.
- EvalOp. apply H1; EvalOp.
Qed.
@@ -304,9 +304,9 @@ Lemma is_longconst_sound:
eval_expr ge sp e m le a x ->
x = Vlong n.
Proof.
- unfold is_longconst; intros until n; intros LC.
+ unfold is_longconst; intros until n; intros LC.
destruct (is_longconst_match a); intros.
- inv LC. InvEval. simpl in H5. inv H5. auto.
+ inv LC. InvEval. simpl in H5. inv H5. auto.
discriminate.
Qed.
@@ -316,35 +316,35 @@ Lemma is_longconst_zero_sound:
eval_expr ge sp e m le a x ->
x = Vlong Int64.zero.
Proof.
- unfold is_longconst_zero; intros.
+ 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.
+ 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.
+ 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.
+ 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:
+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.
+ intros. EvalOp. rewrite Int64.ofwords_recompose; auto.
Qed.
Theorem eval_intoflong: unary_constructor_sound intoflong Val.loword.
@@ -352,10 +352,10 @@ 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.
+ 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.
+ apply Int64.same_bits_eq; intros.
rewrite Int64.testbit_repr by auto.
rewrite Int64.bits_ofwords by auto.
fold (Int.testbit i i0).
@@ -370,17 +370,17 @@ Proof.
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.
+ 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.
+ 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.
+ 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.
@@ -389,19 +389,19 @@ Qed.
Theorem eval_negl: unary_constructor_sound negl Val.negl.
Proof.
unfold negl; red; intros. destruct (is_longconst a) eqn:E.
- econstructor; split. apply eval_longconst.
+ 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.
+ 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.
+ intros; subst. simpl in *. inv B; inv D.
simpl. unfold Int.not. rewrite <- Int64.decompose_xor. auto.
destruct x; auto.
Qed.
@@ -412,7 +412,7 @@ Theorem eval_longoffloat:
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.
+ intros; unfold longoffloat. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
Qed.
@@ -422,7 +422,7 @@ Theorem eval_longuoffloat:
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.
+ intros; unfold longuoffloat. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
Qed.
@@ -432,7 +432,7 @@ Theorem eval_floatoflong:
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.
+ intros; unfold floatoflong. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
Qed.
@@ -442,7 +442,7 @@ Theorem eval_floatoflongu:
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.
+ intros; unfold floatoflongu. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
Qed.
@@ -455,7 +455,7 @@ Proof.
intros; unfold longofsingle.
destruct x; simpl in H0; inv H0. destruct (Float32.to_long f) as [n|] eqn:EQ; simpl in H2; inv H2.
exploit eval_floatofsingle; eauto. intros (v & A & B). simpl in B. inv B.
- apply Float32.to_long_double in EQ.
+ apply Float32.to_long_double in EQ.
eapply eval_longoffloat; eauto. simpl.
change (Float.of_single f) with (Float32.to_double f); rewrite EQ; auto.
Qed.
@@ -469,7 +469,7 @@ Proof.
intros; unfold longuofsingle.
destruct x; simpl in H0; inv H0. destruct (Float32.to_longu f) as [n|] eqn:EQ; simpl in H2; inv H2.
exploit eval_floatofsingle; eauto. intros (v & A & B). simpl in B. inv B.
- apply Float32.to_longu_double in EQ.
+ apply Float32.to_longu_double in EQ.
eapply eval_longuoffloat; eauto. simpl.
change (Float.of_single f) with (Float32.to_double f); rewrite EQ; auto.
Qed.
@@ -480,7 +480,7 @@ Theorem eval_singleoflong:
Val.singleoflong x = Some y ->
exists v, eval_expr ge sp e m le (singleoflong hf a) v /\ Val.lessdef y v.
Proof.
- intros; unfold singleoflong. econstructor; split.
+ intros; unfold singleoflong. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
Qed.
@@ -490,17 +490,17 @@ Theorem eval_singleoflongu:
Val.singleoflongu x = Some y ->
exists v, eval_expr ge sp e m le (singleoflongu hf a) v /\ Val.lessdef y v.
Proof.
- intros; unfold singleoflongu. econstructor; split.
+ intros; unfold singleoflongu. econstructor; split.
eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
Qed.
Theorem eval_andl: binary_constructor_sound andl Val.andl.
Proof.
red; intros. unfold andl. apply eval_splitlong2; auto.
- intros.
+ 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.
+ 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.
@@ -509,10 +509,10 @@ Qed.
Theorem eval_orl: binary_constructor_sound orl Val.orl.
Proof.
red; intros. unfold orl. apply eval_splitlong2; auto.
- intros.
+ 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.
+ 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.
@@ -521,10 +521,10 @@ Qed.
Theorem eval_xorl: binary_constructor_sound xorl Val.xorl.
Proof.
red; intros. unfold xorl. apply eval_splitlong2; auto.
- intros.
+ 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.
+ 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.
@@ -536,7 +536,7 @@ Lemma is_intconst_sound:
eval_expr ge sp e m le a x ->
x = Vint n.
Proof.
- unfold is_intconst; intros until n; intros LC.
+ 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.
@@ -561,28 +561,28 @@ 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).
+ 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.
+ 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.
+ 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.
+ 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.
@@ -593,8 +593,8 @@ 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).
+ 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.
@@ -603,8 +603,8 @@ Proof.
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.
+ 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.
@@ -613,9 +613,9 @@ Proof.
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).
+ 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. DeclHelper. UseHelper. auto.
@@ -638,8 +638,8 @@ Lemma eval_shrluimm:
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).
+ 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.
@@ -648,8 +648,8 @@ Proof.
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.
+ 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.
@@ -658,9 +658,9 @@ Proof.
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).
+ 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. DeclHelper. UseHelper. auto.
@@ -683,8 +683,8 @@ Lemma eval_shrlimm:
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).
+ 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.
@@ -693,8 +693,8 @@ Proof.
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.
+ 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.
@@ -705,11 +705,11 @@ Proof.
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.
+ 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).
+ 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. DeclHelper. UseHelper. auto.
@@ -733,18 +733,18 @@ Proof.
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.
+ 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 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.
+ 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.
+ 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.
@@ -756,19 +756,19 @@ Proof.
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.
+ 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 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.
+ 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.
+ 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.
@@ -776,22 +776,22 @@ Qed.
Lemma eval_mull_base: binary_constructor_sound mull_base Val.mull.
Proof.
unfold mull_base; red; intros. apply eval_splitlong2; auto.
-- intros.
+- 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.
+ 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.
+ 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.
+ 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.
+ 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:
@@ -799,30 +799,30 @@ Lemma eval_mullimm:
Proof.
unfold mullimm; red; intros.
predSpec Int64.eq Int64.eq_spec n Int64.zero.
- subst n. econstructor; split. apply eval_longconst.
+ 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.
+ subst n. exists x; split; auto.
destruct x; simpl; auto. rewrite Int64.mul_one. auto.
- destruct (Int64.is_power2 n) as [l|] eqn:P2.
+ 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.
+ 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).
+ assert (Int64.zwordsize < Int.max_unsigned) by (compute; auto).
omega.
}
- simpl in B.
+ 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.
+ 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.
+ apply eval_mull_base; auto. apply eval_longconst.
Qed.
Theorem eval_mull: binary_constructor_sound (mull hf) Val.mull.
@@ -830,13 +830,13 @@ 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 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.
+ 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.
@@ -851,12 +851,12 @@ Lemma eval_binop_long:
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.
+ 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. 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.
@@ -868,12 +868,12 @@ Theorem eval_divl:
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. 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.
+ || Int64.eq p (Int64.repr Int64.min_signed) && Int64.eq q Int64.mone); inv H1.
auto.
- DeclHelper. UseHelper.
+ DeclHelper. UseHelper.
Qed.
Theorem eval_modl:
@@ -883,10 +883,10 @@ Theorem eval_modl:
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. 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.
+ || Int64.eq p (Int64.repr Int64.min_signed) && Int64.eq q Int64.mone); inv H1.
auto.
DeclHelper. UseHelper.
Qed.
@@ -898,38 +898,38 @@ Theorem eval_divlu:
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.
+ intros. unfold divlu.
set (default := Eexternal hf.(i64_udiv) 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. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. 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 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.
+ 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.
+ 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).
+ 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.
+ erewrite Int64.divu_pow2 by eauto.
+ unfold Int64.shru', Int64.shru. rewrite EQ. auto.
unfold Int64.ltu, Int.ltu. rewrite EQ. auto.
- auto.
Qed.
@@ -941,27 +941,27 @@ Theorem eval_modlu:
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.
+ intros. unfold modlu.
set (default := Eexternal hf.(i64_umod) 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. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. 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 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.
+ 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.
+ 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.
@@ -970,26 +970,26 @@ 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.
+ 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.
+ 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.
+ 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 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.
+ 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).
+ 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.
@@ -998,14 +998,14 @@ Lemma eval_cmpl_eq_zero:
eval_expr ge sp e m le a (Vlong x) ->
eval_expr ge sp e m le (cmpl_eq_zero a) (Val.of_bool (Int64.eq x Int64.zero)).
Proof.
- intros. unfold cmpl_eq_zero.
+ intros. unfold cmpl_eq_zero.
eapply eval_splitlong_strict; eauto. intros.
exploit eval_or. eexact H0. eexact H1. intros [v1 [A1 B1]]. simpl in B1; inv B1.
exploit eval_comp. eexact A1. instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
- instantiate (1 := Ceq). intros [v2 [A2 B2]].
- unfold Val.cmp in B2; simpl in B2.
- rewrite <- decompose_cmpl_eq_zero in B2.
- rewrite Int64.ofwords_recompose in B2.
+ instantiate (1 := Ceq). intros [v2 [A2 B2]].
+ unfold Val.cmp in B2; simpl in B2.
+ rewrite <- decompose_cmpl_eq_zero in B2.
+ rewrite Int64.ofwords_recompose in B2.
destruct (Int64.eq x Int64.zero); inv B2; auto.
Qed.
@@ -1014,14 +1014,14 @@ Lemma eval_cmpl_ne_zero:
eval_expr ge sp e m le a (Vlong x) ->
eval_expr ge sp e m le (cmpl_ne_zero a) (Val.of_bool (negb (Int64.eq x Int64.zero))).
Proof.
- intros. unfold cmpl_ne_zero.
+ intros. unfold cmpl_ne_zero.
eapply eval_splitlong_strict; eauto. intros.
exploit eval_or. eexact H0. eexact H1. intros [v1 [A1 B1]]. simpl in B1; inv B1.
exploit eval_comp. eexact A1. instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
- instantiate (1 := Cne). intros [v2 [A2 B2]].
- unfold Val.cmp in B2; simpl in B2.
- rewrite <- decompose_cmpl_eq_zero in B2.
- rewrite Int64.ofwords_recompose in B2.
+ instantiate (1 := Cne). intros [v2 [A2 B2]].
+ unfold Val.cmp in B2; simpl in B2.
+ rewrite <- decompose_cmpl_eq_zero in B2.
+ rewrite Int64.ofwords_recompose in B2.
destruct (negb (Int64.eq x Int64.zero)); inv B2; auto.
Qed.
@@ -1035,7 +1035,7 @@ Lemma eval_cmplu_gen:
else Int.cmpu ch (Int64.hiword x) (Int64.hiword y))).
Proof.
intros. unfold cmplu_gen. eapply eval_splitlong2_strict; eauto. intros.
- econstructor. econstructor. EvalOp. simpl. eauto.
+ econstructor. econstructor. EvalOp. simpl. eauto.
destruct (Int.eq (Int64.hiword x) (Int64.hiword y)); EvalOp.
Qed.
@@ -1051,34 +1051,34 @@ Proof.
Qed.
Theorem eval_cmplu:
- forall c le a x b y v,
+ forall c le a x b y v,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.cmplu c x y = Some v ->
eval_expr ge sp e m le (cmplu c a b) v.
Proof.
intros. unfold Val.cmplu in H1.
- destruct x; simpl in H1; try discriminate. destruct y; inv H1.
+ destruct x; simpl in H1; try discriminate. destruct y; inv H1.
rename i into x. rename i0 into y.
destruct c; simpl.
- (* Ceq *)
exploit eval_xorl. eexact H. eexact H0. intros [v1 [A B]]. simpl in B. inv B.
- rewrite int64_eq_xor. apply eval_cmpl_eq_zero; auto.
+ rewrite int64_eq_xor. apply eval_cmpl_eq_zero; auto.
- (* Cne *)
exploit eval_xorl. eexact H. eexact H0. intros [v1 [A B]]. simpl in B. inv B.
- rewrite int64_eq_xor. apply eval_cmpl_ne_zero; auto.
+ rewrite int64_eq_xor. apply eval_cmpl_ne_zero; auto.
- (* Clt *)
exploit (eval_cmplu_gen Clt Clt). eexact H. eexact H0. simpl.
rewrite <- Int64.decompose_ltu. rewrite ! Int64.ofwords_recompose. auto.
- (* Cle *)
- exploit (eval_cmplu_gen Clt Cle). eexact H. eexact H0. intros.
+ exploit (eval_cmplu_gen Clt Cle). eexact H. eexact H0. intros.
rewrite <- (Int64.ofwords_recompose x). rewrite <- (Int64.ofwords_recompose y).
- rewrite Int64.decompose_leu. auto.
+ rewrite Int64.decompose_leu. auto.
- (* Cgt *)
exploit (eval_cmplu_gen Cgt Cgt). eexact H. eexact H0. simpl.
rewrite Int.eq_sym. rewrite <- Int64.decompose_ltu. rewrite ! Int64.ofwords_recompose. auto.
- (* Cge *)
- exploit (eval_cmplu_gen Cgt Cge). eexact H. eexact H0. intros.
+ exploit (eval_cmplu_gen Cgt Cge). eexact H. eexact H0. intros.
rewrite <- (Int64.ofwords_recompose x). rewrite <- (Int64.ofwords_recompose y).
rewrite Int64.decompose_leu. rewrite Int.eq_sym. auto.
Qed.
@@ -1093,7 +1093,7 @@ Lemma eval_cmpl_gen:
else Int.cmp ch (Int64.hiword x) (Int64.hiword y))).
Proof.
intros. unfold cmpl_gen. eapply eval_splitlong2_strict; eauto. intros.
- econstructor. econstructor. EvalOp. simpl. eauto.
+ econstructor. econstructor. EvalOp. simpl. eauto.
destruct (Int.eq (Int64.hiword x) (Int64.hiword y)); EvalOp.
Qed.
@@ -1101,29 +1101,29 @@ Remark decompose_cmpl_lt_zero:
forall h l,
Int64.lt (Int64.ofwords h l) Int64.zero = Int.lt h Int.zero.
Proof.
- intros.
+ 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.
+ rewrite Int64.decompose_shru_2.
change (Int.sub (Int.repr 63) Int.iwordsize)
- with (Int.repr (Int.zwordsize - 1)).
- rewrite Int.shru_lt_zero.
+ 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.
elim Int64.one_not_zero. auto.
vm_compute. intuition congruence.
Qed.
Theorem eval_cmpl:
- forall c le a x b y v,
+ forall c le a x b y v,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.cmpl c x y = Some v ->
eval_expr ge sp e m le (cmpl c a b) v.
Proof.
intros. unfold Val.cmpl in H1.
- destruct x; simpl in H1; try discriminate. destruct y; inv H1.
+ destruct x; simpl in H1; try discriminate. destruct y; inv H1.
rename i into x. rename i0 into y.
destruct c; simpl.
- (* Ceq *)
@@ -1135,8 +1135,8 @@ Proof.
- (* Clt *)
destruct (is_longconst_zero b) eqn:LC.
+ exploit is_longconst_zero_sound; eauto. intros EQ; inv EQ; clear H0.
- exploit eval_highlong. eexact H. intros [v1 [A1 B1]]. simpl in B1. inv B1.
- exploit eval_comp. eexact A1.
+ exploit eval_highlong. eexact H. intros [v1 [A1 B1]]. simpl in B1. inv B1.
+ exploit eval_comp. eexact A1.
instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
instantiate (1 := Clt). intros [v2 [A2 B2]].
unfold Val.cmp in B2. simpl in B2.
@@ -1145,9 +1145,9 @@ Proof.
+ exploit (eval_cmpl_gen Clt Clt). eexact H. eexact H0. simpl.
rewrite <- Int64.decompose_lt. rewrite ! Int64.ofwords_recompose. auto.
- (* Cle *)
- exploit (eval_cmpl_gen Clt Cle). eexact H. eexact H0. intros.
+ exploit (eval_cmpl_gen Clt Cle). eexact H. eexact H0. intros.
rewrite <- (Int64.ofwords_recompose x). rewrite <- (Int64.ofwords_recompose y).
- rewrite Int64.decompose_le. auto.
+ rewrite Int64.decompose_le. auto.
- (* Cgt *)
exploit (eval_cmpl_gen Cgt Cgt). eexact H. eexact H0. simpl.
rewrite Int.eq_sym. rewrite <- Int64.decompose_lt. rewrite ! Int64.ofwords_recompose. auto.
@@ -1155,13 +1155,13 @@ Proof.
destruct (is_longconst_zero b) eqn:LC.
+ exploit is_longconst_zero_sound; eauto. intros EQ; inv EQ; clear H0.
exploit eval_highlong. eexact H. intros [v1 [A1 B1]]. simpl in B1; inv B1.
- exploit eval_comp. eexact A1.
+ exploit eval_comp. eexact A1.
instantiate (2 := Eop (Ointconst Int.zero) Enil). EvalOp.
instantiate (1 := Cge). intros [v2 [A2 B2]].
- unfold Val.cmp in B2; simpl in B2.
+ unfold Val.cmp in B2; simpl in B2.
rewrite <- (Int64.ofwords_recompose x). rewrite decompose_cmpl_lt_zero.
destruct (negb (Int.lt (Int64.hiword x) Int.zero)); inv B2; auto.
-+ exploit (eval_cmpl_gen Cgt Cge). eexact H. eexact H0. intros.
++ exploit (eval_cmpl_gen Cgt Cge). eexact H. eexact H0. intros.
rewrite <- (Int64.ofwords_recompose x). rewrite <- (Int64.ofwords_recompose y).
rewrite Int64.decompose_le. rewrite Int.eq_sym. auto.
Qed.
diff --git a/backend/Selection.v b/backend/Selection.v
index dea8a008..dcefdd1c 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -83,10 +83,10 @@ Definition sel_constant (cst: Cminor.constant) : expr :=
Definition sel_unop (op: Cminor.unary_operation) (arg: expr) : expr :=
match op with
- | Cminor.Ocast8unsigned => cast8unsigned arg
- | Cminor.Ocast8signed => cast8signed arg
- | Cminor.Ocast16unsigned => cast16unsigned arg
- | Cminor.Ocast16signed => cast16signed arg
+ | Cminor.Ocast8unsigned => cast8unsigned arg
+ | Cminor.Ocast8signed => cast8signed arg
+ | Cminor.Ocast16unsigned => cast16unsigned arg
+ | Cminor.Ocast16signed => cast16signed arg
| Cminor.Onegint => negint arg
| Cminor.Onotint => notint arg
| Cminor.Onegf => negf arg
@@ -295,7 +295,7 @@ Fixpoint sel_stmt (ge: Cminor.genv) (s: Cminor.stmt) : res stmt :=
| Cminor.Sbuiltin optid ef args =>
OK (Sbuiltin (sel_builtin_res optid) ef
(sel_builtin_args args (Machregs.builtin_constraints ef)))
- | Cminor.Stailcall sg fn args =>
+ | Cminor.Stailcall sg fn args =>
OK (match classify_call ge fn with
| Call_imm id => Stailcall sg (inr _ id) (sel_exprlist args)
| _ => Stailcall sg (inl _ (sel_expr fn)) (sel_exprlist args)
@@ -362,7 +362,7 @@ Definition globdef_of_interest (gd: globdef) : bool :=
Definition record_globdef (globs: PTree.t globdef) (id_gd: ident * globdef) :=
let (id, gd) := id_gd in
- if globdef_of_interest gd
+ if globdef_of_interest gd
then PTree.set id gd globs
else PTree.remove id globs.
diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v
index 8ea4c56e..8051df59 100644
--- a/backend/Selectionproof.v
+++ b/backend/Selectionproof.v
@@ -53,20 +53,20 @@ Hypothesis TRANSFPROG: transform_partial_program (sel_fundef hf ge) prog = OK tp
Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
- intros. eapply Genv.find_symbol_transf_partial; eauto.
+ intros. eapply Genv.find_symbol_transf_partial; eauto.
Qed.
Lemma public_preserved:
forall (s: ident), Genv.public_symbol tge s = Genv.public_symbol ge s.
Proof.
- intros. eapply Genv.public_symbol_transf_partial; eauto.
+ intros. eapply Genv.public_symbol_transf_partial; eauto.
Qed.
Lemma function_ptr_translated:
forall (b: block) (f: Cminor.fundef),
Genv.find_funct_ptr ge b = Some f ->
exists tf, Genv.find_funct_ptr tge b = Some tf /\ sel_fundef hf ge f = OK tf.
-Proof.
+Proof.
intros. eapply Genv.find_funct_ptr_transf_partial; eauto.
Qed.
@@ -75,7 +75,7 @@ Lemma functions_translated:
Genv.find_funct ge v = Some f ->
Val.lessdef v v' ->
exists tf, Genv.find_funct tge v' = Some tf /\ sel_fundef hf ge f = OK tf.
-Proof.
+Proof.
intros. inv H0.
eapply Genv.find_funct_transf_partial; eauto.
simpl in H. discriminate.
@@ -84,13 +84,13 @@ Qed.
Lemma sig_function_translated:
forall f tf, sel_fundef hf ge f = OK tf -> funsig tf = Cminor.funsig f.
Proof.
- intros. destruct f; monadInv H; auto. monadInv EQ. auto.
+ intros. destruct f; monadInv H; auto. monadInv EQ. auto.
Qed.
Lemma stackspace_function_translated:
forall f tf, sel_function hf ge f = OK tf -> fn_stackspace tf = Cminor.fn_stackspace f.
Proof.
- intros. monadInv H. auto.
+ intros. monadInv H. auto.
Qed.
Lemma varinfo_preserved:
@@ -103,14 +103,14 @@ Lemma helper_declared_preserved:
forall id name sg, helper_declared ge id name sg -> helper_declared tge id name sg.
Proof.
intros id name sg (b & A & B).
- exploit function_ptr_translated; eauto. simpl. intros (tf & P & Q). inv Q.
+ exploit function_ptr_translated; eauto. simpl. intros (tf & P & Q). inv Q.
exists b; split; auto. rewrite symbols_preserved. auto.
Qed.
Let HELPERS': helper_functions_declared tge hf.
Proof.
red in HELPERS; decompose [Logic.and] HELPERS.
- red. auto 20 using helper_declared_preserved.
+ red. auto 20 using helper_declared_preserved.
Qed.
Section CMCONSTR.
@@ -127,15 +127,15 @@ Lemma eval_condexpr_of_expr:
Proof.
intros until a. functional induction (condexpr_of_expr a); intros.
(* compare *)
- inv H. econstructor; eauto.
+ inv H. econstructor; eauto.
simpl in H6. inv H6. apply Val.bool_of_val_of_optbool. auto.
(* condition *)
inv H. econstructor; eauto. destruct va; eauto.
(* let *)
inv H. econstructor; eauto.
(* default *)
- econstructor. constructor. eauto. constructor.
- simpl. inv H0. auto.
+ econstructor. constructor. eauto. constructor.
+ simpl. inv H0. auto.
Qed.
Lemma eval_load:
@@ -145,10 +145,10 @@ Lemma eval_load:
eval_expr tge sp e m le (load chunk a) v'.
Proof.
intros. generalize H0; destruct v; simpl; intro; try discriminate.
- unfold load.
+ unfold load.
generalize (eval_addressing _ _ _ _ _ chunk _ _ _ _ H (refl_equal _)).
- destruct (addressing chunk a). intros [vl [EV EQ]].
- eapply eval_Eload; eauto.
+ destruct (addressing chunk a). intros [vl [EV EQ]].
+ eapply eval_Eload; eauto.
Qed.
Lemma eval_store:
@@ -162,8 +162,8 @@ Proof.
intros. generalize H1; destruct v1; simpl; intro; try discriminate.
unfold store.
generalize (eval_addressing _ _ _ _ _ chunk _ _ _ _ H (refl_equal _)).
- destruct (addressing chunk a1). intros [vl [EV EQ]].
- eapply step_store; eauto.
+ destruct (addressing chunk a1). intros [vl [EV EQ]].
+ eapply step_store; eauto.
Qed.
(** Correctness of instruction selection for operators *)
@@ -269,7 +269,7 @@ Lemma expr_is_addrof_ident_correct:
expr_is_addrof_ident e = Some id ->
e = Cminor.Econst (Cminor.Oaddrsymbol id Int.zero).
Proof.
- intros e id. unfold expr_is_addrof_ident.
+ intros e id. unfold expr_is_addrof_ident.
destruct e; try congruence.
destruct c; try congruence.
predSpec Int.eq Int.eq_spec i0 Int.zero; congruence.
@@ -285,14 +285,14 @@ Lemma classify_call_correct:
| Call_builtin ef => fd = External ef
end.
Proof.
- unfold classify_call; intros.
- destruct (expr_is_addrof_ident a) as [id|] eqn:?.
+ unfold classify_call; intros.
+ destruct (expr_is_addrof_ident a) as [id|] eqn:?.
exploit expr_is_addrof_ident_correct; eauto. intros EQ; subst a.
- inv H. inv H2.
- destruct (Genv.find_symbol ge id) as [b|] eqn:?.
- rewrite Genv.find_funct_find_funct_ptr in H0.
- rewrite H0.
- destruct fd. exists b; auto.
+ inv H. inv H2.
+ destruct (Genv.find_symbol ge id) as [b|] eqn:?.
+ rewrite Genv.find_funct_find_funct_ptr in H0.
+ rewrite H0.
+ destruct fd. exists b; auto.
destruct (ef_inline e0). auto. exists b; auto.
simpl in H0. discriminate.
auto.
@@ -344,18 +344,18 @@ Proof.
inv WF.
assert (eval_expr tge sp e m le (make_cmp_eq (Eletvar arg) key) (Val.of_bool (zeq i key))).
{ eapply eval_make_cmp_eq; eauto. constructor; auto. }
- eapply eval_XEcondition with (va := zeq i key).
- eapply eval_condexpr_of_expr; eauto. destruct (zeq i key); constructor; auto.
- destruct (zeq i key); simpl.
+ eapply eval_XEcondition with (va := zeq i key).
+ eapply eval_condexpr_of_expr; eauto. destruct (zeq i key); constructor; auto.
+ destruct (zeq i key); simpl.
+ inv MATCH. constructor.
+ eapply IHt; eauto.
- (* lt test *)
inv WF.
assert (eval_expr tge sp e m le (make_cmp_ltu (Eletvar arg) key) (Val.of_bool (zlt i key))).
{ eapply eval_make_cmp_ltu; eauto. constructor; auto. }
- eapply eval_XEcondition with (va := zlt i key).
- eapply eval_condexpr_of_expr; eauto. destruct (zlt i key); constructor; auto.
- destruct (zlt i key); simpl.
+ eapply eval_XEcondition with (va := zlt i key).
+ eapply eval_condexpr_of_expr; eauto. destruct (zlt i key); constructor; auto.
+ destruct (zlt i key); simpl.
+ eapply IHt1; eauto.
+ eapply IHt2; eauto.
- (* jump table *)
@@ -366,13 +366,13 @@ Proof.
set (i' := (i - ofs) mod modulus) in *.
assert (eval_expr tge sp e m (v' :: le) (make_cmp_ltu (Eletvar O) sz) (Val.of_bool (zlt i' sz))).
{ eapply eval_make_cmp_ltu; eauto. constructor; auto. }
- econstructor. eauto.
+ econstructor. eauto.
eapply eval_XEcondition with (va := zlt i' sz).
eapply eval_condexpr_of_expr; eauto. destruct (zlt i' sz); constructor; auto.
destruct (zlt i' sz); simpl.
+ exploit (eval_make_to_int sp e m (v' :: le) (Eletvar O)).
- constructor. simpl; eauto. eauto. intros (v'' & C & D). inv D.
- econstructor; eauto. congruence.
+ constructor. simpl; eauto. eauto. intros (v'' & C & D). inv D.
+ econstructor; eauto. congruence.
+ eapply IHt; eauto.
Qed.
@@ -398,37 +398,37 @@ Lemma sel_switch_int_correct:
eval_expr tge sp e m le arg (Vint i) ->
eval_exitexpr tge sp e m le (XElet arg (sel_switch_int O t)) (switch_target (Int.unsigned i) dfl cases).
Proof.
- assert (INTCONST: forall n sp e m le,
+ assert (INTCONST: forall n sp e m le,
eval_expr tge sp e m le (Eop (Ointconst n) Enil) (Vint n)).
- { intros. econstructor. constructor. auto. }
+ { intros. econstructor. constructor. auto. }
intros. eapply sel_switch_correct with (R := Rint); eauto.
- intros until n; intros EVAL R RANGE.
- exploit eval_comp. eexact EVAL. apply (INTCONST (Int.repr n)).
- instantiate (1 := Ceq). intros (vb & A & B).
- inv R. unfold Val.cmp in B. simpl in B. revert B.
- predSpec Int.eq Int.eq_spec n0 (Int.repr n); intros B; inv B.
- rewrite Int.unsigned_repr. unfold proj_sumbool; rewrite zeq_true; auto.
+ exploit eval_comp. eexact EVAL. apply (INTCONST (Int.repr n)).
+ instantiate (1 := Ceq). intros (vb & A & B).
+ inv R. unfold Val.cmp in B. simpl in B. revert B.
+ predSpec Int.eq Int.eq_spec n0 (Int.repr n); intros B; inv B.
+ rewrite Int.unsigned_repr. unfold proj_sumbool; rewrite zeq_true; auto.
unfold Int.max_unsigned; omega.
unfold proj_sumbool; rewrite zeq_false; auto.
red; intros; elim H1. rewrite <- (Int.repr_unsigned n0). congruence.
- intros until n; intros EVAL R RANGE.
- exploit eval_compu. eexact EVAL. apply (INTCONST (Int.repr n)).
- instantiate (1 := Clt). intros (vb & A & B).
+ exploit eval_compu. eexact EVAL. apply (INTCONST (Int.repr n)).
+ instantiate (1 := Clt). intros (vb & A & B).
inv R. unfold Val.cmpu in B. simpl in B.
- unfold Int.ltu in B. rewrite Int.unsigned_repr in B.
- destruct (zlt (Int.unsigned n0) n); inv B; auto.
+ unfold Int.ltu in B. rewrite Int.unsigned_repr in B.
+ destruct (zlt (Int.unsigned n0) n); inv B; auto.
unfold Int.max_unsigned; omega.
- intros until n; intros EVAL R RANGE.
exploit eval_sub. eexact EVAL. apply (INTCONST (Int.repr n)). intros (vb & A & B).
- inv R. simpl in B. inv B. econstructor; split; eauto.
+ inv R. simpl in B. inv B. econstructor; split; eauto.
replace ((Int.unsigned n0 - n) mod Int.modulus)
with (Int.unsigned (Int.sub n0 (Int.repr n))).
constructor.
- unfold Int.sub. rewrite Int.unsigned_repr_eq. f_equal. f_equal.
+ unfold Int.sub. rewrite Int.unsigned_repr_eq. f_equal. f_equal.
apply Int.unsigned_repr. unfold Int.max_unsigned; omega.
-- intros until i0; intros EVAL R. exists v; split; auto.
+- intros until i0; intros EVAL R. exists v; split; auto.
inv R. rewrite Zmod_small by (apply Int.unsigned_range). constructor.
-- constructor.
+- constructor.
- apply Int.unsigned_range.
Qed.
@@ -441,30 +441,30 @@ Proof.
intros. eapply sel_switch_correct with (R := Rlong); eauto.
- intros until n; intros EVAL R RANGE.
eapply eval_cmpl. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
- inv R. unfold Val.cmpl. simpl. f_equal; f_equal. unfold Int64.eq.
- rewrite Int64.unsigned_repr. destruct (zeq (Int64.unsigned n0) n); auto.
+ inv R. unfold Val.cmpl. simpl. f_equal; f_equal. unfold Int64.eq.
+ rewrite Int64.unsigned_repr. destruct (zeq (Int64.unsigned n0) n); auto.
unfold Int64.max_unsigned; omega.
-- intros until n; intros EVAL R RANGE.
+- intros until n; intros EVAL R RANGE.
eapply eval_cmplu. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
- inv R. unfold Val.cmplu. simpl. f_equal; f_equal. unfold Int64.ltu.
- rewrite Int64.unsigned_repr. destruct (zlt (Int64.unsigned n0) n); auto.
+ inv R. unfold Val.cmplu. simpl. f_equal; f_equal. unfold Int64.ltu.
+ rewrite Int64.unsigned_repr. destruct (zlt (Int64.unsigned n0) n); auto.
unfold Int64.max_unsigned; omega.
- intros until n; intros EVAL R RANGE.
exploit eval_subl. eexact EVAL. apply eval_longconst with (n := Int64.repr n).
intros (vb & A & B).
- inv R. simpl in B. inv B. econstructor; split; eauto.
+ inv R. simpl in B. inv B. econstructor; split; eauto.
replace ((Int64.unsigned n0 - n) mod Int64.modulus)
with (Int64.unsigned (Int64.sub n0 (Int64.repr n))).
constructor.
- unfold Int64.sub. rewrite Int64.unsigned_repr_eq. f_equal. f_equal.
+ unfold Int64.sub. rewrite Int64.unsigned_repr_eq. f_equal. f_equal.
apply Int64.unsigned_repr. unfold Int64.max_unsigned; omega.
-- intros until i0; intros EVAL R.
- exploit eval_lowlong. eexact EVAL. intros (vb & A & B).
+- intros until i0; intros EVAL R.
+ exploit eval_lowlong. eexact EVAL. intros (vb & A & B).
inv R. simpl in B. inv B. econstructor; split; eauto.
replace (Int64.unsigned n mod Int.modulus) with (Int.unsigned (Int64.loword n)).
constructor.
- unfold Int64.loword. apply Int.unsigned_repr_eq.
-- constructor.
+ unfold Int64.loword. apply Int.unsigned_repr_eq.
+- constructor.
- apply Int64.unsigned_range.
Qed.
@@ -481,24 +481,24 @@ Lemma eval_unop_lessdef:
eval_unop op v1 = Some v -> Val.lessdef v1 v1' ->
exists v', eval_unop op v1' = Some v' /\ Val.lessdef v v'.
Proof.
- intros until v; intros EV LD. inv LD.
+ intros until v; intros EV LD. inv LD.
exists v; auto.
destruct op; simpl in *; inv EV; TrivialExists.
Qed.
Lemma eval_binop_lessdef:
forall op v1 v1' v2 v2' v m m',
- eval_binop op v1 v2 m = Some v ->
+ eval_binop op v1 v2 m = Some v ->
Val.lessdef v1 v1' -> Val.lessdef v2 v2' -> Mem.extends m m' ->
exists v', eval_binop op v1' v2' m' = Some v' /\ Val.lessdef v v'.
Proof.
intros until m'; intros EV LD1 LD2 ME.
assert (exists v', eval_binop op v1' v2' m = Some v' /\ Val.lessdef v v').
- inv LD1. inv LD2. exists v; auto.
+ inv LD1. inv LD2. exists v; auto.
destruct op; destruct v1'; simpl in *; inv EV; TrivialExists.
destruct op; simpl in *; inv EV; TrivialExists.
- destruct op; try (exact H).
- simpl in *. TrivialExists. inv EV. apply Val.of_optbool_lessdef.
+ destruct op; try (exact H).
+ simpl in *. TrivialExists. inv EV. apply Val.of_optbool_lessdef.
intros. apply Val.cmpu_bool_lessdef with (Mem.valid_pointer m) v1 v2; auto.
intros; eapply Mem.valid_pointer_extends; eauto.
Qed.
@@ -529,7 +529,7 @@ Proof.
Qed.
Lemma set_params_lessdef:
- forall il vl1 vl2,
+ forall il vl1 vl2,
Val.lessdef_list vl1 vl2 ->
env_lessdef (set_params vl1 il) (set_params vl2 il).
Proof.
@@ -558,10 +558,10 @@ Proof.
(* Evar *)
exploit H0; eauto. intros [v' [A B]]. exists v'; split; auto. constructor; auto.
(* Econst *)
- destruct cst; simpl in *; inv H.
- exists (Vint i); split; auto. econstructor. constructor. auto.
+ destruct cst; simpl in *; inv H.
+ exists (Vint i); split; auto. econstructor. constructor. auto.
exists (Vfloat f); split; auto. econstructor. constructor. auto.
- exists (Vsingle f); split; auto. econstructor. constructor. auto.
+ exists (Vsingle 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.
@@ -571,13 +571,13 @@ Proof.
exploit IHeval_expr; eauto. intros [v1' [A B]].
exploit eval_unop_lessdef; eauto. intros [v' [C D]].
exploit eval_sel_unop; eauto. intros [v'' [E F]].
- exists v''; split; eauto. eapply Val.lessdef_trans; eauto.
+ exists v''; split; eauto. eapply Val.lessdef_trans; eauto.
(* Ebinop *)
exploit IHeval_expr1; eauto. intros [v1' [A B]].
exploit IHeval_expr2; eauto. intros [v2' [C D]].
exploit eval_binop_lessdef; eauto. intros [v' [E F]].
exploit eval_sel_binop. eexact A. eexact C. eauto. intros [v'' [P Q]].
- exists v''; split; eauto. eapply Val.lessdef_trans; eauto.
+ exists v''; split; eauto. eapply Val.lessdef_trans; eauto.
(* Eload *)
exploit IHeval_expr; eauto. intros [vaddr' [A B]].
exploit Mem.loadv_extends; eauto. intros [v' [C D]].
@@ -591,7 +591,7 @@ Lemma sel_exprlist_correct:
env_lessdef e e' -> Mem.extends m m' ->
exists v', eval_exprlist tge sp e' m' le (sel_exprlist hf a) v' /\ Val.lessdef_list v v'.
Proof.
- induction 1; intros; simpl.
+ induction 1; intros; simpl.
exists (@nil val); split; auto. constructor.
exploit sel_expr_correct; eauto. intros [v1' [A B]].
exploit IHeval_exprlist; eauto. intros [vl' [C D]].
@@ -606,7 +606,7 @@ Lemma sel_builtin_arg_correct:
CminorSel.eval_builtin_arg tge sp e' m' (sel_builtin_arg hf a c) v'
/\ Val.lessdef v v'.
Proof.
- intros. unfold sel_builtin_arg.
+ intros. unfold sel_builtin_arg.
exploit sel_expr_correct; eauto. intros (v1 & A & B).
exists v1; split; auto.
destruct (builtin_arg_ok (builtin_arg (sel_expr hf a)) c).
@@ -627,7 +627,7 @@ Lemma sel_builtin_args_correct:
/\ Val.lessdef_list vl vl'.
Proof.
induction 3; intros; simpl.
-- exists (@nil val); split; constructor.
+- exists (@nil val); split; constructor.
- exploit sel_builtin_arg_correct; eauto. intros (v1' & A & B).
edestruct IHeval_exprlist as (vl' & C & D).
exists (v1' :: vl'); split; auto. constructor; eauto.
@@ -638,7 +638,7 @@ Lemma sel_builtin_res_correct:
env_lessdef e e' -> Val.lessdef v v' ->
env_lessdef (set_optvar oid v e) (set_builtin_res (sel_builtin_res oid) v' e').
Proof.
- intros. destruct oid; simpl; auto. apply set_var_lessdef; auto.
+ intros. destruct oid; simpl; auto. apply set_var_lessdef; auto.
Qed.
(** Semantic preservation for functions and statements. *)
@@ -727,12 +727,12 @@ Proof.
(* tailcall *)
destruct (classify_call ge e); simpl; auto.
(* seq *)
- exploit (IHs1 (Cminor.Kseq s2 k)). constructor; eauto. eauto.
+ exploit (IHs1 (Cminor.Kseq s2 k)). constructor; eauto. eauto.
destruct (Cminor.find_label lbl s1 (Cminor.Kseq s2 k)) as [[sx kx] | ];
destruct (find_label lbl x (Kseq x0 k')) as [[sy ky] | ];
intuition. apply IHs2; auto.
(* ifthenelse *)
- exploit (IHs1 k); eauto.
+ exploit (IHs1 k); eauto.
destruct (Cminor.find_label lbl s1 k) as [[sx kx] | ];
destruct (find_label lbl x k') as [[sy ky] | ];
intuition. apply IHs2; auto.
@@ -741,7 +741,7 @@ Proof.
(* block *)
apply IHs. constructor; auto. auto.
(* switch *)
- destruct b.
+ destruct b.
destruct (validate_switch Int64.modulus n l (compile_switch Int64.modulus n l)); inv SE.
simpl; auto.
destruct (validate_switch Int.modulus n l (compile_switch Int.modulus n l)); inv SE.
@@ -772,10 +772,10 @@ Proof.
inv MC. left; econstructor; split. econstructor. constructor; auto.
- (* skip call *)
exploit Mem.free_parallel_extends; eauto. intros [m2' [A B]].
- left; econstructor; split.
- econstructor. inv MC; simpl in H; simpl; auto.
- eauto.
- erewrite stackspace_function_translated; eauto.
+ left; econstructor; split.
+ econstructor. inv MC; simpl in H; simpl; auto.
+ eauto.
+ erewrite stackspace_function_translated; eauto.
constructor; auto.
- (* assign *)
exploit sel_expr_correct; eauto. intros [v' [A B]].
@@ -790,18 +790,18 @@ Proof.
eapply eval_store; eauto.
constructor; auto.
- (* Scall *)
- exploit classify_call_correct; eauto.
+ exploit classify_call_correct; eauto.
destruct (classify_call ge a) as [ | id | ef].
+ (* indirect *)
exploit sel_expr_correct; eauto. intros [vf' [A B]].
exploit sel_exprlist_correct; eauto. intros [vargs' [C D]].
exploit functions_translated; eauto. intros (fd' & U & V).
left; econstructor; split.
- econstructor; eauto. econstructor; eauto.
+ econstructor; eauto. econstructor; eauto.
apply sig_function_translated; auto.
constructor; auto. constructor; auto.
+ (* direct *)
- intros [b [U V]].
+ intros [b [U V]].
exploit sel_exprlist_correct; eauto. intros [vargs' [C D]].
exploit functions_translated; eauto. intros (fd' & X & Y).
left; econstructor; split.
@@ -812,7 +812,7 @@ Proof.
+ (* turned into Sbuiltin *)
intros EQ. subst fd.
exploit sel_builtin_args_correct; eauto. intros [vargs' [C D]].
- right; split. simpl. omega. split. auto.
+ right; split. simpl. omega. split. auto.
econstructor; eauto.
- (* Stailcall *)
exploit Mem.free_parallel_extends; eauto. intros [m2' [P Q]].
@@ -821,21 +821,21 @@ Proof.
exploit sel_exprlist_correct; eauto. intros [vargs' [C D]].
exploit functions_translated; eauto. intros [fd' [E F]].
left; econstructor; split.
- exploit classify_call_correct; eauto.
- destruct (classify_call ge a) as [ | id | ef]; intros.
+ exploit classify_call_correct; eauto.
+ destruct (classify_call ge a) as [ | id | ef]; intros.
econstructor; eauto. econstructor; eauto. apply sig_function_translated; auto.
- destruct H2 as [b [U V]]. subst vf. inv B.
+ destruct H2 as [b [U V]]. subst vf. inv B.
econstructor; eauto. econstructor; eauto. rewrite symbols_preserved; eauto. apply sig_function_translated; auto.
econstructor; eauto. econstructor; eauto. apply sig_function_translated; auto.
constructor; auto. apply call_cont_commut; auto.
- (* Sbuiltin *)
exploit sel_builtin_args_correct; eauto. intros [vargs' [P Q]].
- exploit external_call_mem_extends; eauto.
+ exploit external_call_mem_extends; eauto.
intros [vres' [m2 [A [B [C D]]]]].
left; econstructor; split.
econstructor. eauto. eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- constructor; auto. apply sel_builtin_res_correct; auto.
+ constructor; auto. apply sel_builtin_res_correct; auto.
- (* Seq *)
left; econstructor; split.
constructor. constructor; auto. constructor; auto.
@@ -843,7 +843,7 @@ Proof.
exploit sel_expr_correct; eauto. intros [v' [A B]].
assert (Val.bool_of_val v' b). inv B. auto. inv H0.
left; exists (State f' (if b then x else x0) k' sp e' m'); split.
- econstructor; eauto. eapply eval_condexpr_of_expr; eauto.
+ econstructor; eauto. eapply eval_condexpr_of_expr; eauto.
constructor; auto. destruct b; auto.
- (* Sloop *)
left; econstructor; split. constructor. constructor; auto.
@@ -861,26 +861,26 @@ Proof.
+ set (ct := compile_switch Int.modulus default cases) in *.
destruct (validate_switch Int.modulus default cases ct) eqn:VALID; inv TS.
exploit sel_expr_correct; eauto. intros [v' [A B]]. inv B.
- left; econstructor; split.
- econstructor. eapply sel_switch_int_correct; eauto.
+ left; econstructor; split.
+ econstructor. eapply sel_switch_int_correct; eauto.
constructor; auto.
+ set (ct := compile_switch Int64.modulus default cases) in *.
destruct (validate_switch Int64.modulus default cases ct) eqn:VALID; inv TS.
exploit sel_expr_correct; eauto. intros [v' [A B]]. inv B.
left; econstructor; split.
- econstructor. eapply sel_switch_long_correct; eauto.
+ econstructor. eapply sel_switch_long_correct; eauto.
constructor; auto.
- (* Sreturn None *)
exploit Mem.free_parallel_extends; eauto. intros [m2' [P Q]].
erewrite <- stackspace_function_translated in P by eauto.
- left; econstructor; split.
- econstructor. simpl; eauto.
+ left; econstructor; split.
+ econstructor. simpl; eauto.
constructor; auto. apply call_cont_commut; auto.
- (* Sreturn Some *)
exploit Mem.free_parallel_extends; eauto. intros [m2' [P Q]].
erewrite <- stackspace_function_translated in P by eauto.
exploit sel_expr_correct; eauto. intros [v' [A B]].
- left; econstructor; split.
+ left; econstructor; split.
econstructor; eauto.
constructor; auto. apply call_cont_commut; auto.
- (* Slabel *)
@@ -890,39 +890,39 @@ Proof.
{ monadInv TF; simpl; auto. }
exploit (find_label_commut lbl (Cminor.fn_body f) (Cminor.call_cont k)).
apply call_cont_commut; eauto. eauto.
- rewrite H.
+ rewrite H.
destruct (find_label lbl (fn_body f') (call_cont k'0))
as [[s'' k'']|] eqn:?; intros; try contradiction.
- destruct H1.
+ destruct H1.
left; econstructor; split.
- econstructor; eauto.
+ econstructor; eauto.
constructor; auto.
- (* internal function *)
monadInv TF. generalize EQ; intros TF; monadInv TF.
- exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl.
+ exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl.
intros [m2' [A B]].
left; econstructor; split.
econstructor; simpl; eauto.
constructor; simpl; auto. apply set_locals_lessdef. apply set_params_lessdef; auto.
- (* external call *)
monadInv TF.
- exploit external_call_mem_extends; eauto.
+ exploit external_call_mem_extends; eauto.
intros [vres' [m2 [A [B [C D]]]]].
left; econstructor; split.
econstructor. eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
constructor; auto.
- (* external call turned into a Sbuiltin *)
- exploit external_call_mem_extends; eauto.
+ exploit external_call_mem_extends; eauto.
intros [vres' [m2 [A [B [C D]]]]].
left; econstructor; split.
econstructor. eauto. eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
constructor; auto.
- (* return *)
- inv MC.
- left; econstructor; split.
- econstructor.
+ inv MC.
+ left; econstructor; split.
+ econstructor.
constructor; auto. destruct optid; simpl; auto. apply set_var_lessdef; auto.
- (* return of an external call turned into a Sbuiltin *)
right; split. simpl; omega. split. auto. constructor; auto.
@@ -975,20 +975,20 @@ Proof.
- auto.
- apply IHgl. red. destruct a as [id1 gd1]; simpl; intros.
unfold Genv.find_symbol; simpl. rewrite PTree.gsspec. destruct (peq id id1).
- + subst id1. exists (Genv.genv_next ge); split; auto.
+ + subst id1. exists (Genv.genv_next ge); split; auto.
replace gd1 with (@Gfun Cminor.fundef unit fd).
- unfold Genv.find_funct_ptr; simpl. apply PTree.gss.
+ unfold Genv.find_funct_ptr; simpl. apply PTree.gss.
destruct (globdef_of_interest gd1).
rewrite PTree.gss in H0; congruence.
rewrite PTree.grs in H0; congruence.
+ assert (m!id = Some (Gfun fd)).
- { destruct (globdef_of_interest gd1).
+ { destruct (globdef_of_interest gd1).
rewrite PTree.gso in H0; auto.
rewrite PTree.gro in H0; auto. }
exploit H; eauto. intros (b & A & B).
exists b; split; auto. unfold Genv.find_funct_ptr; simpl.
- destruct gd1; auto. rewrite PTree.gso; auto.
- apply Plt_ne. eapply Genv.genv_symb_range; eauto.
+ destruct gd1; auto. rewrite PTree.gso; auto.
+ apply Plt_ne. eapply Genv.genv_symb_range; eauto.
}
eapply REC. red; intros. rewrite PTree.gempty in H; discriminate.
Qed.
@@ -1002,19 +1002,19 @@ Proof.
set (P := fun (m: PTree.t globdef) res => res = Some id -> m!id = Some(Gfun(External (EF_external name sg)))).
assert (P globs (PTree.fold (lookup_helper_aux name sg) globs None)).
{ apply PTree_Properties.fold_rec; red; intros.
- - rewrite <- H0. apply H1; auto.
+ - rewrite <- H0. apply H1; auto.
- discriminate.
- assert (EITHER: k = id /\ v = Gfun (External (EF_external name sg))
\/ a = Some id).
- { unfold lookup_helper_aux in H3. destruct v; auto. destruct f; auto. destruct e; auto.
+ { unfold lookup_helper_aux in H3. destruct v; auto. destruct f; auto. destruct e; auto.
destruct (String.string_dec name name0); auto.
destruct (signature_eq sg sg0); auto.
inversion H3. left; split; auto. repeat f_equal; auto. }
destruct EITHER as [[X Y] | X].
- subst k v. apply PTree.gss.
+ subst k v. apply PTree.gss.
apply H2 in X. rewrite PTree.gso by congruence. auto.
}
- red in H0. unfold lookup_helper in H.
+ red in H0. unfold lookup_helper in H.
destruct (PTree.fold (lookup_helper_aux name sg) globs None); inv H. auto.
Qed.
@@ -1040,11 +1040,11 @@ Theorem transf_program_correct:
sel_program prog = OK tprog ->
forward_simulation (Cminor.semantics prog) (CminorSel.semantics tprog).
Proof.
- intros. unfold sel_program in H.
+ intros. unfold sel_program in H.
destruct (get_helpers prog) as [hf|] eqn:G; simpl in H; try discriminate.
- apply forward_simulation_opt with (match_states := match_states prog tprog hf) (measure := measure).
+ apply forward_simulation_opt with (match_states := match_states prog tprog hf) (measure := measure).
eapply public_preserved; eauto.
apply sel_initial_states; auto.
apply sel_final_states; auto.
- apply sel_step_correct; auto. eapply get_helpers_correct; eauto.
+ apply sel_step_correct; auto. eapply get_helpers_correct; eauto.
Qed.
diff --git a/backend/Splitting.ml b/backend/Splitting.ml
index 97b26a50..17b8098d 100644
--- a/backend/Splitting.ml
+++ b/backend/Splitting.ml
@@ -39,9 +39,9 @@ let rec repr lr =
| 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 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
diff --git a/backend/Stacking.v b/backend/Stacking.v
index ef96e4b3..ab67e213 100644
--- a/backend/Stacking.v
+++ b/backend/Stacking.v
@@ -71,13 +71,13 @@ Definition save_callee_save_regs
List.fold_right (save_callee_save_reg bound number mkindex ty fe) k csl.
Definition save_callee_save_int (fe: frame_env) :=
- save_callee_save_regs
+ save_callee_save_regs
fe_num_int_callee_save index_int_callee_save FI_saved_int
Tany32 fe int_callee_save_regs.
Definition save_callee_save_float (fe: frame_env) :=
save_callee_save_regs
- fe_num_float_callee_save index_float_callee_save FI_saved_float
+ fe_num_float_callee_save index_float_callee_save FI_saved_float
Tany64 fe float_callee_save_regs.
Definition save_callee_save (fe: frame_env) (k: Mach.code) :=
@@ -101,13 +101,13 @@ Definition restore_callee_save_regs
List.fold_right (restore_callee_save_reg bound number mkindex ty fe) k csl.
Definition restore_callee_save_int (fe: frame_env) :=
- restore_callee_save_regs
+ restore_callee_save_regs
fe_num_int_callee_save index_int_callee_save FI_saved_int
Tany32 fe int_callee_save_regs.
Definition restore_callee_save_float (fe: frame_env) :=
restore_callee_save_regs
- fe_num_float_callee_save index_float_callee_save FI_saved_float
+ fe_num_float_callee_save index_float_callee_save FI_saved_float
Tany64 fe float_callee_save_regs.
Definition restore_callee_save (fe: frame_env) (k: Mach.code) :=
@@ -146,7 +146,7 @@ Fixpoint transl_builtin_arg (fe: frame_env) (a: builtin_arg loc) : builtin_arg m
BA_addrstack (Int.add ofs (Int.repr fe.(fe_stack_data)))
| BA_loadglobal chunk id ofs => BA_loadglobal chunk id ofs
| BA_addrglobal id ofs => BA_addrglobal id ofs
- | BA_splitlong hi lo =>
+ | BA_splitlong hi lo =>
BA_splitlong (transl_builtin_arg fe hi) (transl_builtin_arg fe lo)
end.
diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v
index dce49432..8becb773 100644
--- a/backend/Stackingproof.v
+++ b/backend/Stackingproof.v
@@ -111,7 +111,7 @@ Remark bound_stack_data_stacksize:
f.(Linear.fn_stacksize) <= b.(bound_stack_data).
Proof.
unfold b, function_bounds, bound_stack_data. apply Zmax1.
-Qed.
+Qed.
(** A frame index is valid if it lies within the resource bounds
of the current function. *)
@@ -155,7 +155,7 @@ Definition index_diff (idx1 idx2: frame_index) : Prop :=
Lemma index_diff_sym:
forall idx1 idx2, index_diff idx1 idx2 -> index_diff idx2 idx1.
Proof.
- unfold index_diff; intros.
+ unfold index_diff; intros.
destruct idx1; destruct idx2; intuition.
Qed.
@@ -222,9 +222,9 @@ Lemma offset_of_index_disj_stack_data_2:
offset_of_index fe idx + AST.typesize (type_of_index idx) <= fe.(fe_stack_data)
\/ fe.(fe_stack_data) + f.(Linear.fn_stacksize) <= offset_of_index fe idx.
Proof.
- intros.
+ intros.
exploit offset_of_index_disj_stack_data_1; eauto.
- generalize bound_stack_data_stacksize.
+ generalize bound_stack_data_stacksize.
omega.
Qed.
@@ -240,7 +240,7 @@ Remark aligned_8_4:
forall x, (8 | x) -> (4 | x).
Proof. intros. apply Zdivides_trans with 8; auto. exists 2; auto. Qed.
-Hint Resolve Zdivide_0 Zdivide_refl Zdivide_plus_r
+Hint Resolve Zdivide_0 Zdivide_refl Zdivide_plus_r
aligned_4_4x aligned_4_8x aligned_8_4: align_4.
Hint Extern 4 (?X | ?Y) => (exists (Y/X); reflexivity) : align_4.
@@ -280,8 +280,8 @@ Lemma index_local_valid:
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, slot_valid, index_valid; intros.
- InvBooleans.
+ unfold slot_within_bounds, slot_valid, index_valid; intros.
+ InvBooleans.
split. destruct ty; auto || discriminate. auto.
Qed.
@@ -290,9 +290,9 @@ Lemma index_arg_valid:
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, slot_valid, index_valid; intros.
- InvBooleans.
- split. destruct ty; auto || discriminate. auto.
+ unfold slot_within_bounds, slot_valid, index_valid; intros.
+ InvBooleans.
+ split. destruct ty; auto || discriminate. auto.
Qed.
Lemma index_saved_int_valid:
@@ -301,8 +301,8 @@ Lemma index_saved_int_valid:
index_int_callee_save r < b.(bound_int_callee_save) ->
index_valid (FI_saved_int (index_int_callee_save r)).
Proof.
- intros. red. split.
- apply Zge_le. apply index_int_callee_save_pos; auto.
+ intros. red. split.
+ apply Zge_le. apply index_int_callee_save_pos; auto.
auto.
Qed.
@@ -312,8 +312,8 @@ Lemma index_saved_float_valid:
index_float_callee_save r < b.(bound_float_callee_save) ->
index_valid (FI_saved_float (index_float_callee_save r)).
Proof.
- intros. red. split.
- apply Zge_le. apply index_float_callee_save_pos; auto.
+ intros. red. split.
+ apply Zge_le. apply index_float_callee_save_pos; auto.
auto.
Qed.
@@ -360,7 +360,7 @@ Proof.
generalize (offset_of_index_valid idx H). intros [A B].
apply Int.unsigned_repr.
generalize (AST.typesize_pos (type_of_index idx)).
- generalize size_no_overflow.
+ generalize size_no_overflow.
omega.
Qed.
@@ -369,14 +369,14 @@ Qed.
Lemma shifted_stack_offset_no_overflow:
forall ofs,
0 <= Int.unsigned ofs < Linear.fn_stacksize f ->
- Int.unsigned (Int.add ofs (Int.repr fe.(fe_stack_data)))
+ Int.unsigned (Int.add ofs (Int.repr fe.(fe_stack_data)))
= Int.unsigned ofs + fe.(fe_stack_data).
Proof.
intros. unfold Int.add.
generalize size_no_overflow stack_data_offset_valid bound_stack_data_stacksize; intros.
AddPosProps.
replace (Int.unsigned (Int.repr (fe_stack_data fe))) with (fe_stack_data fe).
- apply Int.unsigned_repr. omega.
+ apply Int.unsigned_repr. omega.
symmetry. apply Int.unsigned_repr. omega.
Qed.
@@ -394,7 +394,7 @@ Lemma index_contains_load_stack:
load_stack m (Vptr sp Int.zero) (type_of_index idx)
(Int.repr (offset_of_index fe idx)) = Some v.
Proof.
- intros. inv H.
+ intros. inv H.
unfold load_stack, Mem.loadv, Val.add. rewrite Int.add_commut. rewrite Int.add_zero.
rewrite offset_of_index_no_overflow; auto.
Qed.
@@ -409,8 +409,8 @@ Lemma gss_index_contains_base:
index_contains m' sp idx v'
/\ decode_encode_val v (chunk_of_type (type_of_index idx)) (chunk_of_type (type_of_index idx)) v'.
Proof.
- intros.
- exploit Mem.load_store_similar. eauto. reflexivity. omega.
+ intros.
+ exploit Mem.load_store_similar. eauto. reflexivity. omega.
intros [v' [A B]].
exists v'; split; auto. constructor; auto.
Qed.
@@ -437,9 +437,9 @@ Lemma gso_index_contains:
index_diff idx idx' ->
index_contains m' sp idx' v'.
Proof.
- intros. inv H1. constructor; auto.
+ intros. inv H1. constructor; auto.
rewrite <- H4. eapply Mem.load_store_other; eauto.
- right. repeat rewrite size_type_chunk.
+ right. repeat rewrite size_type_chunk.
apply offset_of_index_disj; auto. apply index_diff_sym; auto.
Qed.
@@ -451,9 +451,9 @@ Lemma store_other_index_contains:
index_contains m sp idx v ->
index_contains m' sp idx v.
Proof.
- intros. inv H1. constructor; auto. rewrite <- H3.
- eapply Mem.load_store_other; eauto.
- destruct H0. auto. right.
+ intros. inv H1. constructor; auto. rewrite <- H3.
+ eapply Mem.load_store_other; eauto.
+ destruct H0. auto. right.
exploit offset_of_index_disj_stack_data_2; eauto. intros.
rewrite size_type_chunk.
omega.
@@ -487,7 +487,7 @@ Proof.
intros.
destruct (Mem.valid_access_store m (chunk_of_type (type_of_index idx)) sp (offset_of_index fe idx) v) as [m' ST].
constructor.
- rewrite size_type_chunk.
+ rewrite size_type_chunk.
apply Mem.range_perm_implies with Freeable; auto with mem.
apply offset_of_index_perm; auto.
apply offset_of_index_aligned_2; auto.
@@ -535,7 +535,7 @@ Lemma gss_index_contains_inj':
Proof.
intros. exploit gss_index_contains_base; eauto. intros [v'' [A B]].
exists v''; split; auto.
- inv H1; destruct (type_of_index idx); simpl in *; try contradiction; subst; auto.
+ inv H1; destruct (type_of_index idx); simpl in *; try contradiction; subst; auto.
econstructor; eauto.
econstructor; eauto.
econstructor; eauto.
@@ -571,7 +571,7 @@ Lemma index_contains_inj_incr:
inject_incr j j' ->
index_contains_inj j' m sp idx v.
Proof.
- intros. destruct H as [v'' [A B]]. exists v''; split; auto. eauto.
+ intros. destruct H as [v'' [A B]]. exists v''; split; auto. eauto.
Qed.
Lemma index_contains_inj_undef:
@@ -580,15 +580,15 @@ Lemma index_contains_inj_undef:
frame_perm_freeable m sp ->
index_contains_inj j m sp idx Vundef.
Proof.
- intros.
+ intros.
exploit (Mem.valid_access_load m (chunk_of_type (type_of_index idx)) sp (offset_of_index fe idx)).
- constructor.
+ constructor.
rewrite size_type_chunk.
apply Mem.range_perm_implies with Freeable; auto with mem.
- apply offset_of_index_perm; auto.
+ apply offset_of_index_perm; auto.
apply offset_of_index_aligned_2; auto.
- intros [v C].
- exists v; split; auto. constructor; auto.
+ intros [v C].
+ exists v; split; auto. constructor; auto.
Qed.
Hint Resolve store_other_index_contains_inj index_contains_inj_incr: stacking.
@@ -613,21 +613,21 @@ Record agree_frame (j: meminj) (ls ls0: locset)
forall r, ~(mreg_within_bounds b r) -> ls (R r) = ls0 (R r);
(** Local and outgoing stack slots (on the Linear side) have
- the same values as the one loaded from the current Mach frame
+ the same values as the one loaded from the current Mach frame
at the corresponding offsets. *)
agree_locals:
- forall ofs ty,
+ forall 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,
+ forall 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,
+ forall ofs ty,
In (S Incoming ofs ty) (loc_parameters f.(Linear.fn_sig)) ->
ls (S Incoming ofs ty) = ls0 (S Outgoing ofs ty);
@@ -695,7 +695,7 @@ Lemma agree_reg:
forall j ls rs r,
agree_regs j ls rs -> Val.inject j (ls (R r)) (rs r).
Proof.
- intros. auto.
+ intros. auto.
Qed.
Lemma agree_reglist:
@@ -703,7 +703,7 @@ Lemma agree_reglist:
agree_regs j ls rs -> Val.inject_list j (reglist ls rl) (rs##rl).
Proof.
induction rl; simpl; intros.
- auto. constructor. eauto with stacking. auto.
+ auto. constructor. eauto with stacking. auto.
Qed.
Hint Resolve agree_reg agree_reglist: stacking.
@@ -716,8 +716,8 @@ Lemma agree_regs_set_reg:
Val.inject j v v' ->
agree_regs j (Locmap.set (R r) v ls) (Regmap.set r v' rs).
Proof.
- intros; red; intros.
- unfold Regmap.set. destruct (RegEq.eq r0 r). subst r0.
+ intros; red; intros.
+ unfold Regmap.set. destruct (RegEq.eq r0 r). subst r0.
rewrite Locmap.gss; auto.
rewrite Locmap.gso; auto. red. auto.
Qed.
@@ -728,10 +728,10 @@ Lemma agree_regs_set_regs:
Val.inject_list j vl vl' ->
agree_regs j (Locmap.setlist (map R rl) vl ls) (set_regs rl vl' rs).
Proof.
- induction rl; simpl; intros.
+ induction rl; simpl; intros.
auto.
inv H0. auto.
- apply IHrl; auto. apply agree_regs_set_reg; auto.
+ apply IHrl; auto. apply agree_regs_set_reg; auto.
Qed.
Lemma agree_regs_set_res:
@@ -741,9 +741,9 @@ Lemma agree_regs_set_res:
agree_regs j (Locmap.setres res v ls) (set_res res v' rs).
Proof.
induction res; simpl; intros.
-- apply agree_regs_set_reg; auto.
+- apply agree_regs_set_reg; auto.
- auto.
-- apply IHres2. apply IHres1. auto.
+- apply IHres2. apply IHres1. auto.
apply Val.hiword_inject; auto.
apply Val.loword_inject; auto.
Qed.
@@ -755,8 +755,8 @@ Lemma agree_regs_exten:
agree_regs j ls' rs'.
Proof.
intros; red; intros.
- destruct (H0 r) as [A | [A B]].
- rewrite A. constructor.
+ destruct (H0 r) as [A | [A B]].
+ rewrite A. constructor.
rewrite A; rewrite B; auto.
Qed.
@@ -767,7 +767,7 @@ Lemma agree_regs_undef_regs:
Proof.
induction rl; simpl; intros.
auto.
- apply agree_regs_set_reg; auto.
+ apply agree_regs_set_reg; auto.
Qed.
(** Preservation under assignment of stack slot *)
@@ -821,8 +821,8 @@ Lemma agree_frame_set_regs:
agree_frame j (Locmap.setlist (map R rl) vl ls) ls0 m sp m' sp' parent ra.
Proof.
induction rl; destruct vl; simpl; intros; intuition.
- apply IHrl; auto.
- eapply agree_frame_set_reg; eauto.
+ apply IHrl; auto.
+ eapply agree_frame_set_reg; eauto.
Qed.
Lemma agree_frame_set_res:
@@ -866,8 +866,8 @@ Lemma agree_frame_undef_locs:
incl regs destroyed_at_call ->
agree_frame j (LTL.undef_regs regs ls) ls0 m sp m' sp' parent ra.
Proof.
- intros. eapply agree_frame_undef_regs; eauto.
- intros. apply caller_save_reg_within_bounds. auto.
+ intros. eapply agree_frame_undef_regs; eauto.
+ intros. apply caller_save_reg_within_bounds. auto.
Qed.
(** Preservation by assignment to local slot *)
@@ -880,7 +880,7 @@ Lemma agree_frame_set_local:
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.
Proof.
- intros. inv H.
+ intros. inv H.
change (chunk_of_type ty) with (chunk_of_type (type_of_index (FI_local ofs ty))) in H3.
constructor; auto; intros.
(* local *)
@@ -888,19 +888,19 @@ Proof.
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.
+ 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.
- red; auto. red; left; congruence.
+ red; auto. red; left; congruence.
(* parent *)
eapply gso_index_contains; eauto with stacking. red; auto.
(* retaddr *)
eapply gso_index_contains; eauto with stacking. red; auto.
(* int callee save *)
- eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
+ eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
(* float callee save *)
eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
(* valid *)
@@ -919,7 +919,7 @@ Lemma agree_frame_set_outgoing:
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.
Proof.
- intros. inv H.
+ intros. inv H.
change (chunk_of_type ty) with (chunk_of_type (type_of_index (FI_arg ofs ty))) in H3.
constructor; auto; intros.
(* local *)
@@ -930,7 +930,7 @@ Proof.
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.
+ red. red in d. intuition.
apply index_contains_inj_undef. auto with stacking.
red; intros. eapply Mem.perm_store_1; eauto.
(* parent *)
@@ -938,7 +938,7 @@ Proof.
(* retaddr *)
eapply gso_index_contains; eauto with stacking. red; auto.
(* int callee save *)
- eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
+ eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
(* float callee save *)
eapply gso_index_contains_inj; eauto with stacking. simpl; auto.
(* valid *)
@@ -969,11 +969,11 @@ Proof.
assert (IC: forall idx v,
index_contains m' sp' idx v -> index_contains m1' sp' idx v).
intros. inv H5.
- exploit offset_of_index_disj_stack_data_2; eauto. intros.
+ exploit offset_of_index_disj_stack_data_2; eauto. intros.
constructor; eauto. apply H3; auto. rewrite size_type_chunk; auto.
assert (ICI: forall idx v,
index_contains_inj j m' sp' idx v -> index_contains_inj j m1' sp' idx v).
- intros. destruct H5 as [v' [A B]]. exists v'; split; auto.
+ intros. destruct H5 as [v' [A B]]. exists v'; split; auto.
inv H; constructor; auto; intros.
eauto.
red; intros. apply H4; auto.
@@ -995,10 +995,10 @@ Proof.
ofs < fe.(fe_stack_data) \/ fe.(fe_stack_data) + f.(Linear.fn_stacksize) <= ofs ->
loc_out_of_reach j m sp' ofs).
intros; red; intros. exploit agree_inj_unique; eauto. intros [EQ1 EQ2]; subst.
- red; intros. exploit agree_bounds; eauto. omega.
+ red; intros. exploit agree_bounds; eauto. omega.
eapply agree_frame_invariant; eauto.
- intros. eapply Mem.load_unchanged_on; eauto. intros. apply REACH. omega. auto.
- intros. eapply Mem.perm_unchanged_on; eauto with mem. auto.
+ intros. eapply Mem.load_unchanged_on; eauto. intros. apply REACH. omega. auto.
+ intros. eapply Mem.perm_unchanged_on; eauto with mem. auto.
Qed.
(** Preservation by parallel stores in the Linear and Mach codes *)
@@ -1019,20 +1019,20 @@ Opaque Int.add.
eauto with mem.
eauto with mem.
eauto with mem.
- intros. rewrite <- H1. eapply Mem.load_store_other; eauto.
+ intros. rewrite <- H1. eapply Mem.load_store_other; eauto.
destruct (eq_block sp' b2); auto.
subst b2. right.
exploit agree_inj_unique; eauto. intros [P Q]. subst b1 delta.
exploit Mem.store_valid_access_3. eexact STORE1. intros [A B].
rewrite shifted_stack_offset_no_overflow.
- exploit agree_bounds. eauto. apply Mem.perm_cur_max. apply A.
+ exploit agree_bounds. eauto. apply Mem.perm_cur_max. apply A.
instantiate (1 := Int.unsigned ofs1). generalize (size_chunk_pos chunk). omega.
intros C.
- exploit agree_bounds. eauto. apply Mem.perm_cur_max. apply A.
+ exploit agree_bounds. eauto. apply Mem.perm_cur_max. apply A.
instantiate (1 := Int.unsigned ofs1 + size_chunk chunk - 1). generalize (size_chunk_pos chunk). omega.
intros D.
omega.
- eapply agree_bounds. eauto. apply Mem.perm_cur_max. apply A.
+ eapply agree_bounds. eauto. apply Mem.perm_cur_max. apply A.
generalize (size_chunk_pos chunk). omega.
intros; eauto with mem.
Qed.
@@ -1047,8 +1047,8 @@ Lemma agree_frame_inject_incr:
agree_frame j' ls ls0 m sp m' sp' parent retaddr.
Proof.
intros. inv H. constructor; auto; intros; eauto with stacking.
- case_eq (j b0).
- intros [b' delta'] EQ. rewrite (H0 _ _ _ EQ) in H. inv H. auto.
+ case_eq (j b0).
+ intros [b' delta'] EQ. rewrite (H0 _ _ _ EQ) in H. inv H. auto.
intros EQ. exploit H1. eauto. eauto. intros [A B]. contradiction.
Qed.
@@ -1074,7 +1074,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. red; intros; elim H. apply caller_save_reg_within_bounds; auto.
+ rewrite H0; auto. red; intros; elim H. apply caller_save_reg_within_bounds; auto.
rewrite H0; auto.
rewrite H0; auto.
rewrite H0; auto.
@@ -1089,10 +1089,10 @@ 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. red; intros; elim H. apply caller_save_reg_within_bounds; 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.
+ 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]. *)
@@ -1103,7 +1103,7 @@ Lemma agree_callee_save_return_regs:
Proof.
intros; red; intros.
unfold return_regs. destruct l; auto.
- rewrite pred_dec_false; auto.
+ rewrite pred_dec_false; auto.
Qed.
Lemma agree_callee_save_set_result:
@@ -1116,10 +1116,10 @@ Proof.
Opaque destroyed_at_call.
induction l; simpl; intros.
auto.
- destruct vl; auto.
+ destruct vl; auto.
apply IHl; auto.
- red; intros. rewrite Locmap.gso. apply H0; auto.
- destruct l0; simpl; auto.
+ red; intros. rewrite Locmap.gso. apply H0; auto.
+ destruct l0; simpl; auto.
Qed.
(** Properties of destroyed registers. *)
@@ -1129,8 +1129,8 @@ Lemma check_mreg_list_incl:
forallb (fun r => In_dec mreg_eq r l2) l1 = true ->
incl l1 l2.
Proof.
- intros; red; intros.
- rewrite forallb_forall in H. eapply proj_sumbool_true. eauto.
+ intros; red; intros.
+ rewrite forallb_forall in H. eapply proj_sumbool_true. eauto.
Qed.
Remark destroyed_by_op_caller_save:
@@ -1191,7 +1191,7 @@ Hint Resolve destroyed_by_op_caller_save destroyed_by_load_caller_save
Remark destroyed_by_setstack_function_entry:
forall ty, incl (destroyed_by_setstack ty) destroyed_at_function_entry.
Proof.
- destruct ty; apply check_mreg_list_incl; compute; auto.
+ destruct ty; apply check_mreg_list_incl; compute; auto.
Qed.
Remark transl_destroyed_by_op:
@@ -1249,7 +1249,7 @@ Proof.
induction 1; intros. auto. econstructor; eauto.
Qed.
-Hypothesis number_inj:
+Hypothesis number_inj:
forall r1 r2, In r1 csregs -> In r2 csregs -> r1 <> r2 -> number r1 <> number r2.
Hypothesis mkindex_valid:
forall r, In r csregs -> number r < bound fe -> index_valid (mkindex (number r)).
@@ -1275,7 +1275,7 @@ Lemma save_callee_save_regs_correct:
frame_perm_freeable m sp ->
agree_regs j ls rs ->
exists rs', exists m',
- star step tge
+ star step tge
(State cs fb (Vptr sp Int.zero)
(save_callee_save_regs bound number mkindex ty fe l k) rs m)
E0 (State cs fb (Vptr sp Int.zero) k rs' m')
@@ -1294,7 +1294,7 @@ Lemma save_callee_save_regs_correct:
Proof.
induction l; intros; simpl save_callee_save_regs.
(* base case *)
- exists rs; exists m. split. apply star_refl.
+ exists rs; exists m. split. apply star_refl.
split. intros. elim H3.
split. auto.
split. constructor.
@@ -1305,42 +1305,42 @@ Proof.
unfold save_callee_save_reg.
destruct (zlt (number a) (bound fe)).
(* a store takes place *)
- exploit store_index_succeeds. apply (mkindex_valid a); auto with coqlib.
+ exploit store_index_succeeds. apply (mkindex_valid a); auto with coqlib.
eauto. instantiate (1 := rs a). intros [m1 ST].
- exploit (IHl k (undef_regs (destroyed_by_setstack ty) rs) m1). auto with coqlib. auto.
+ exploit (IHl k (undef_regs (destroyed_by_setstack ty) 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_by_setstack ty)).
- left. apply ls_temp_undef; auto.
+ left. apply ls_temp_undef; 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)).
+ exists rs'; exists m'.
+ split. eapply star_left; eauto. econstructor.
+ rewrite <- (mkindex_typ (number a)).
apply store_stack_succeeds; auto with coqlib.
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))).
eapply gss_index_contains_inj; eauto.
- rewrite mkindex_typ. rewrite <- (csregs_typ a). apply wt_ls.
+ rewrite mkindex_typ. rewrite <- (csregs_typ a). apply wt_ls.
auto with coqlib.
destruct H5 as [v' [P Q]].
- exists v'; split; auto. apply C; auto.
- intros. apply mkindex_inj. apply number_inj; auto with coqlib.
+ exists v'; split; auto. apply C; auto.
+ intros. apply mkindex_inj. apply number_inj; auto with coqlib.
inv H0. intuition congruence.
- apply B; auto with coqlib.
+ apply B; auto with coqlib.
intuition congruence.
split. intros.
apply C; auto with coqlib.
- eapply gso_index_contains; eauto with coqlib.
+ eapply gso_index_contains; eauto with coqlib.
split. econstructor; eauto.
rewrite size_type_chunk. apply offset_of_index_disj_stack_data_2; eauto with coqlib.
auto.
(* no store takes place *)
- exploit (IHl k rs m); auto with coqlib.
+ exploit (IHl k rs m); auto with coqlib.
intros [rs' [m' [A [B [C [D [E F]]]]]]].
- exists rs'; exists m'; intuition.
+ exists rs'; exists m'; intuition.
simpl in H3. destruct H3. subst r. omegaContradiction. apply B; auto.
apply C; auto with coqlib.
intros. eapply H4; eauto. auto with coqlib.
@@ -1351,9 +1351,9 @@ 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.
+ 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.
@@ -1361,14 +1361,14 @@ 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.
+ 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.
+ rewrite Locmap.gso. apply IHrl. red; auto.
Qed.
Lemma save_callee_save_correct:
@@ -1378,7 +1378,7 @@ Lemma save_callee_save_correct:
(forall r, Val.has_type (ls (R r)) (mreg_type r)) ->
frame_perm_freeable m sp ->
exists rs', exists m',
- star step tge
+ star step tge
(State cs fb (Vptr sp Int.zero) (save_callee_save fe k) rs m)
E0 (State cs fb (Vptr sp Int.zero) k rs' m')
/\ (forall r,
@@ -1399,12 +1399,12 @@ Proof.
intros.
assert (UNDEF: forall r ty, In r (destroyed_by_setstack ty) -> ls (R r) = Vundef).
intros. unfold ls. apply LTL_undef_regs_same. eapply destroyed_by_setstack_function_entry; eauto.
- exploit (save_callee_save_regs_correct
+ exploit (save_callee_save_regs_correct
fe_num_int_callee_save
index_int_callee_save
FI_saved_int Tany32
j cs fb sp int_callee_save_regs ls).
- intros. apply index_int_callee_save_inj; auto.
+ intros. apply index_int_callee_save_inj; auto.
intros. simpl. split. apply Zge_le. apply index_int_callee_save_pos; auto. assumption.
auto.
intros; congruence.
@@ -1412,41 +1412,41 @@ Proof.
intros. apply int_callee_save_type. auto.
eauto.
auto.
- apply incl_refl.
+ apply incl_refl.
apply int_callee_save_norepet.
eauto.
eauto.
intros [rs1 [m1 [A [B [C [D [E F]]]]]]].
- exploit (save_callee_save_regs_correct
+ exploit (save_callee_save_regs_correct
fe_num_float_callee_save
index_float_callee_save
FI_saved_float Tany64
j cs fb sp float_callee_save_regs ls).
- intros. apply index_float_callee_save_inj; auto.
+ intros. apply index_float_callee_save_inj; auto.
intros. simpl. split. apply Zge_le. apply index_float_callee_save_pos; auto. assumption.
simpl; auto.
intros; congruence.
intros; simpl. destruct idx; auto. congruence.
intros. apply float_callee_save_type. auto.
eauto.
- auto.
- apply incl_refl.
+ auto.
+ apply incl_refl.
apply float_callee_save_norepet.
eexact E.
eexact F.
intros [rs2 [m2 [P [Q [R [S [T U]]]]]]].
exists rs2; exists m2.
split. unfold save_callee_save, save_callee_save_int, save_callee_save_float.
- eapply star_trans; eauto.
+ eapply star_trans; eauto.
split; intros.
destruct (B r H2 H3) as [v [X Y]]. exists v; split; auto. apply R.
- apply index_saved_int_valid; auto.
+ apply index_saved_int_valid; auto.
intros. congruence.
auto.
split. intros. apply Q; auto.
split. intros. apply R. auto.
intros. destruct idx; contradiction||congruence.
- apply C. auto.
+ apply C. auto.
intros. destruct idx; contradiction||congruence.
auto.
split. eapply stores_in_frame_trans; eauto.
@@ -1466,7 +1466,7 @@ Proof.
apply IHstores_in_frame.
intros. eapply Mem.store_outside_inject; eauto.
intros. exploit H; eauto. intros [A B]; subst.
- exploit H0; eauto. omega.
+ exploit H0; eauto. omega.
Qed.
Lemma stores_in_frame_valid:
@@ -1483,10 +1483,10 @@ Qed.
Lemma stores_in_frame_contents:
forall chunk b ofs sp, Plt b sp ->
- forall m m', stores_in_frame sp m m' ->
+ forall m m', stores_in_frame sp m m' ->
Mem.load chunk m' b ofs = Mem.load chunk m b ofs.
Proof.
- induction 2. auto.
+ induction 2. auto.
rewrite IHstores_in_frame. eapply Mem.load_store_other; eauto.
left. apply Plt_ne; auto.
Qed.
@@ -1497,7 +1497,7 @@ Lemma undef_regs_type:
Proof.
induction rl; simpl; intros.
- auto.
-- unfold Locmap.set. destruct (Loc.eq (R a) l). red; auto.
+- unfold Locmap.set. destruct (Loc.eq (R a) l). red; auto.
destruct (Loc.diff_dec (R a) l); auto. red; auto.
Qed.
@@ -1520,7 +1520,7 @@ Lemma function_prologue_correct:
Mem.alloc m1' 0 tf.(fn_stacksize) = (m2', sp')
/\ 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
+ /\ star step tge
(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' ls1 rs'
@@ -1541,13 +1541,13 @@ Proof.
instantiate (1 := sp'). eauto with mem.
instantiate (1 := fe_stack_data fe).
generalize stack_data_offset_valid (bound_stack_data_pos b) size_no_overflow; omega.
- intros; right.
- exploit Mem.perm_alloc_inv. eexact ALLOC'. eauto. rewrite dec_eq_true.
- generalize size_no_overflow. omega.
- intros. apply Mem.perm_implies with Freeable; auto with mem.
+ intros; right.
+ exploit Mem.perm_alloc_inv. eexact ALLOC'. eauto. rewrite dec_eq_true.
+ generalize size_no_overflow. omega.
+ intros. apply Mem.perm_implies with Freeable; auto with mem.
eapply Mem.perm_alloc_2; eauto.
generalize stack_data_offset_valid bound_stack_data_stacksize; omega.
- red. intros. apply Zdivides_trans with 8.
+ red. intros. apply Zdivides_trans with 8.
destruct chunk; simpl; auto with align_4.
apply fe_stack_data_aligned.
intros.
@@ -1558,17 +1558,17 @@ Proof.
assert (PERM: frame_perm_freeable m2' sp').
red; intros. eapply Mem.perm_alloc_2; eauto.
(* Store of parent *)
- exploit (store_index_succeeds m2' sp' FI_link parent). red; auto. auto.
+ exploit (store_index_succeeds m2' sp' FI_link parent). red; auto. auto.
intros [m3' STORE2].
(* Store of retaddr *)
exploit (store_index_succeeds m3' sp' FI_retaddr ra). red; auto. red; eauto with mem.
intros [m4' STORE3].
(* Saving callee-save registers *)
assert (PERM4: frame_perm_freeable m4' sp').
- red; intros. eauto with mem.
- exploit save_callee_save_correct.
+ 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.
+ subst rs1. apply agree_regs_undef_regs.
apply agree_regs_call_regs. eapply agree_regs_inject_incr; eauto.
intros. apply undef_regs_type. simpl. apply WTREGS.
eexact PERM4.
@@ -1576,15 +1576,15 @@ Proof.
intros [rs' [m5' [STEPS [ICS [FCS [OTHERS [STORES [PERM5 AGREGS']]]]]]]].
(* stores in frames *)
assert (SIF: stores_in_frame sp' m2' m5').
- econstructor; eauto.
+ econstructor; eauto.
rewrite size_type_chunk. apply offset_of_index_disj_stack_data_2; auto. red; auto.
econstructor; eauto.
rewrite size_type_chunk. apply offset_of_index_disj_stack_data_2; auto. red; auto.
(* separation *)
assert (SEP: forall b0 delta, j' b0 = Some(sp', delta) -> b0 = sp /\ delta = fe_stack_data fe).
- intros. destruct (eq_block b0 sp).
+ intros. destruct (eq_block b0 sp).
subst b0. rewrite MAP1 in H; inv H; auto.
- rewrite MAP2 in H; auto.
+ rewrite MAP2 in H; auto.
assert (Mem.valid_block m1' sp'). eapply Mem.valid_block_inject_2; eauto.
assert (~Mem.valid_block m1' sp') by eauto with mem.
contradiction.
@@ -1592,11 +1592,11 @@ Proof.
exists j'; exists rs'; exists m2'; exists sp'; exists m3'; exists m4'; exists m5'.
split. auto.
(* store parent *)
- split. change Tint with (type_of_index FI_link).
+ split. change Tint with (type_of_index FI_link).
change (fe_ofs_link fe) with (offset_of_index fe FI_link).
apply store_stack_succeeds; auto. red; auto.
(* store retaddr *)
- split. change Tint with (type_of_index FI_retaddr).
+ split. change Tint with (type_of_index FI_retaddr).
change (fe_ofs_retaddr fe) with (offset_of_index fe FI_retaddr).
apply store_stack_succeeds; auto. red; auto.
(* saving of registers *)
@@ -1606,20 +1606,20 @@ Proof.
(* agree frame *)
split. constructor; intros.
(* unused regs *)
- assert (~In r destroyed_at_call).
+ 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.
+ 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 *)
- rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
+ rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
apply index_contains_inj_undef; auto with stacking.
(* outgoing *)
- rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
+ rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
apply index_contains_inj_undef; auto with stacking.
(* incoming *)
rewrite LS1. rewrite LTL_undef_regs_slot. unfold call_regs.
- apply AGCS; auto.
+ apply AGCS; auto.
(* parent *)
apply OTHERS; auto. red; auto.
eapply gso_index_contains; eauto. red; auto.
@@ -1629,16 +1629,16 @@ Proof.
apply OTHERS; auto. red; auto.
eapply gss_index_contains; eauto. red; auto.
(* int callee save *)
- assert (~In r destroyed_at_call).
+ 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.
+ rewrite AGCS; auto.
red; intros; elim H1. apply destroyed_at_function_entry_caller_save; auto.
(* float callee save *)
- assert (~In r destroyed_at_call).
+ 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.
+ rewrite AGCS; auto.
red; intros; elim H1. apply destroyed_at_function_entry_caller_save; auto.
(* inj *)
auto.
@@ -1700,7 +1700,7 @@ Definition agree_unused (ls0: locset) (rs: regset) : Prop :=
Lemma restore_callee_save_regs_correct:
forall l rs k,
incl l csregs ->
- list_norepet l ->
+ list_norepet l ->
agree_unused ls0 rs ->
exists rs',
star step tge
@@ -1727,9 +1727,9 @@ Proof.
subst r. auto.
auto.
intros [rs' [A [B [C D]]]].
- exists rs'. split.
- eapply star_left.
- constructor. rewrite <- (mkindex_typ (number a)). apply index_contains_load_stack. eauto.
+ exists rs'. split.
+ eapply star_left.
+ constructor. rewrite <- (mkindex_typ (number a)). apply index_contains_load_stack. eauto.
eauto. traceEq.
split. intros. destruct H2.
subst r. rewrite C. unfold rs1. rewrite Regmap.gss. auto. inv H0; auto.
@@ -1742,7 +1742,7 @@ Proof.
intros [rs' [A [B [C D]]]].
exists rs'. split. assumption.
split. intros. destruct H2.
- subst r. apply D.
+ subst r. apply D.
rewrite <- number_within_bounds. auto. auto. auto.
split. intros. simpl in H2. apply C. tauto.
auto.
@@ -1758,16 +1758,16 @@ Lemma restore_callee_save_correct:
star step tge
(State cs fb (Vptr sp' Int.zero) (restore_callee_save fe k) rs m')
E0 (State cs fb (Vptr sp' Int.zero) k rs' m')
- /\ (forall r,
- In r int_callee_save_regs \/ In r float_callee_save_regs ->
+ /\ (forall r,
+ In r int_callee_save_regs \/ In r float_callee_save_regs ->
Val.inject j (ls0 (R r)) (rs' r))
- /\ (forall r,
+ /\ (forall r,
~(In r int_callee_save_regs) ->
~(In r float_callee_save_regs) ->
rs' r = rs r).
Proof.
- intros.
- exploit (restore_callee_save_regs_correct
+ intros.
+ exploit (restore_callee_save_regs_correct
fe_num_int_callee_save
index_int_callee_save
FI_saved_int
@@ -1776,16 +1776,16 @@ Proof.
j cs fb sp' ls0 m'); auto.
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.
+ 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.
+ destruct H2; auto.
+ eapply agree_saved_int; eauto.
apply incl_refl.
apply int_callee_save_norepet.
eauto.
intros [rs1 [A [B [C D]]]].
- exploit (restore_callee_save_regs_correct
+ exploit (restore_callee_save_regs_correct
fe_num_float_callee_save
index_float_callee_save
FI_saved_float
@@ -1794,11 +1794,11 @@ Proof.
j cs fb sp' ls0 m'); auto.
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.
+ 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.
+ destruct H2; auto.
+ eapply agree_saved_float; eauto.
apply incl_refl.
apply float_callee_save_norepet.
eexact D.
@@ -1842,34 +1842,34 @@ Proof.
by omega.
destruct EITHER.
replace ofs with ((ofs - fe_stack_data fe) + fe_stack_data fe) by omega.
- eapply Mem.perm_inject with (f := j). eapply agree_inj; eauto. eauto.
+ eapply Mem.perm_inject with (f := j). eapply agree_inj; eauto. eauto.
eapply Mem.free_range_perm; eauto. omega.
- eapply agree_perm; eauto.
+ eapply agree_perm; eauto.
(* inject after free *)
assert (INJ1: Mem.inject j m1 m1').
eapply Mem.free_inject with (l := (sp, 0, f.(Linear.fn_stacksize)) :: nil); eauto.
simpl. rewrite H2. auto.
intros. exploit agree_inj_unique; eauto. intros [P Q]; subst b1 delta.
exists 0; exists (Linear.fn_stacksize f); split. auto with coqlib.
- eapply agree_bounds. eauto. eapply Mem.perm_max. eauto.
+ eapply agree_bounds. eauto. eapply Mem.perm_max. eauto.
(* can execute epilogue *)
exploit restore_callee_save_correct; eauto.
- instantiate (1 := rs). red; intros.
- rewrite <- (agree_unused_reg _ _ _ _ _ _ _ _ _ H0). auto. auto.
+ instantiate (1 := rs). red; intros.
+ rewrite <- (agree_unused_reg _ _ _ _ _ _ _ _ _ H0). auto. auto.
intros [rs1 [A [B C]]].
(* conclusions *)
exists rs1; exists m1'.
- split. rewrite unfold_transf_function; unfold fn_link_ofs.
+ split. rewrite unfold_transf_function; unfold fn_link_ofs.
eapply index_contains_load_stack with (idx := FI_link); eauto with stacking.
- split. rewrite unfold_transf_function; unfold fn_retaddr_ofs.
+ split. rewrite unfold_transf_function; unfold fn_retaddr_ofs.
eapply index_contains_load_stack with (idx := FI_retaddr); eauto with stacking.
split. auto.
split. eexact A.
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 mreg_eq r destroyed_at_call).
- rewrite C; auto.
- apply B. intuition.
+ destruct (in_dec mreg_eq r destroyed_at_call).
+ rewrite C; auto.
+ apply B. intuition.
split. apply agree_callee_save_return_regs.
auto.
Qed.
@@ -1886,7 +1886,7 @@ Inductive match_globalenvs (j: meminj) (bound: block) : Prop :=
(FUNCTIONS: forall b fd, Genv.find_funct_ptr ge b = Some fd -> Plt b bound)
(VARINFOS: forall b gv, Genv.find_var_info ge b = Some gv -> Plt b bound).
-Inductive match_stacks (j: meminj) (m m': mem):
+Inductive match_stacks (j: meminj) (m m': mem):
list Linear.stackframe -> list stackframe -> signature -> block -> block -> Prop :=
| match_stacks_empty: forall sg hi bound bound',
Ple hi bound -> Ple hi bound' -> match_globalenvs j hi ->
@@ -1919,9 +1919,9 @@ Lemma match_stacks_change_bounds:
Ple bound xbound -> Ple bound' xbound' ->
match_stacks j m1 m' cs cs' sg xbound xbound'.
Proof.
- induction 1; intros.
- apply match_stacks_empty with hi; auto. apply Ple_trans with bound; eauto. apply Ple_trans with bound'; eauto.
- econstructor; eauto. eapply Plt_le_trans; eauto. eapply Plt_le_trans; eauto.
+ induction 1; intros.
+ apply match_stacks_empty with hi; auto. apply Ple_trans with bound; eauto. apply Ple_trans with bound'; eauto.
+ econstructor; eauto. eapply Plt_le_trans; eauto. eapply Plt_le_trans; eauto.
Qed.
(** Invariance with respect to change of [m]. *)
@@ -1936,7 +1936,7 @@ Proof.
induction 1; intros.
econstructor; eauto.
econstructor; eauto.
- eapply agree_frame_invariant; eauto.
+ eapply agree_frame_invariant; eauto.
apply IHmatch_stacks.
intros. apply H0; auto. apply Plt_trans with sp; auto.
intros. apply H1. apply Plt_trans with sp; auto. auto.
@@ -1955,11 +1955,11 @@ Proof.
induction 1; intros.
econstructor; eauto.
econstructor; eauto.
- eapply agree_frame_invariant; eauto.
- apply IHmatch_stacks.
- intros; apply H0; auto. apply Plt_trans with sp'; auto.
- intros; apply H1; auto. apply Plt_trans with sp'; auto.
- intros; apply H2; auto. apply Plt_trans with sp'; auto.
+ eapply agree_frame_invariant; eauto.
+ apply IHmatch_stacks.
+ intros; apply H0; auto. apply Plt_trans with sp'; auto.
+ intros; apply H1; auto. apply Plt_trans with sp'; auto.
+ intros; apply H2; auto. apply Plt_trans with sp'; auto.
Qed.
(** A variant of the latter, for use with external calls *)
@@ -1977,10 +1977,10 @@ Proof.
econstructor; eauto.
econstructor; eauto.
eapply agree_frame_extcall_invariant; eauto.
- apply IHmatch_stacks.
- intros; apply H0; auto. apply Plt_trans with sp; auto.
+ apply IHmatch_stacks.
+ intros; apply H0; auto. apply Plt_trans with sp; auto.
intros; apply H1. apply Plt_trans with sp; auto. auto.
- intros; apply H2; auto. apply Plt_trans with sp'; auto.
+ intros; apply H2; auto. apply Plt_trans with sp'; auto.
auto.
Qed.
@@ -2002,8 +2002,8 @@ Proof.
intros [b' delta'] EQ. rewrite (H _ _ _ EQ) in H3. inv H3. eauto.
intros EQ. exploit H0; eauto. intros [A B]. elim B. red.
apply Plt_le_trans with hi. auto. apply Ple_trans with bound'; auto.
- econstructor; eauto.
- eapply agree_frame_inject_incr; eauto. red. eapply Plt_le_trans; eauto.
+ econstructor; eauto.
+ eapply agree_frame_inject_incr; eauto. red. eapply Plt_le_trans; eauto.
apply IHmatch_stacks. apply Ple_trans with bound'; auto. apply Plt_Ple; auto.
Qed.
@@ -2040,11 +2040,11 @@ Lemma match_stack_change_extcall:
Ple bound (Mem.nextblock m1) -> Ple bound' (Mem.nextblock m1') ->
match_stacks j' m2 m2' cs cs' sg bound bound'.
Proof.
- intros.
- eapply match_stacks_change_meminj; eauto.
+ intros.
+ eapply match_stacks_change_meminj; eauto.
eapply match_stacks_change_mem_extcall; eauto.
intros; eapply external_call_valid_block; eauto.
- intros; eapply external_call_max_perm; eauto. red. eapply Plt_le_trans; eauto.
+ intros; eapply external_call_max_perm; eauto. red. eapply Plt_le_trans; eauto.
intros; eapply external_call_valid_block; eauto.
Qed.
@@ -2057,7 +2057,7 @@ Lemma match_stacks_change_sig:
match_stacks j m m' cs cs' sg1 bound bound'.
Proof.
induction 1; intros.
- econstructor; eauto.
+ econstructor; eauto.
econstructor; eauto. intros. elim (H0 _ H1).
Qed.
@@ -2077,7 +2077,7 @@ Lemma match_stacks_preserves_globals:
meminj_preserves_globals ge j.
Proof.
intros. exploit match_stacks_globalenvs; eauto. intros [hi MG]. inv MG.
- split. eauto. split. eauto. intros. symmetry. eauto.
+ split. eauto. split. eauto. intros. symmetry. eauto.
Qed.
(** Typing properties of [match_stacks]. *)
@@ -2109,7 +2109,7 @@ Remark find_label_fold_right:
(forall x k, Mach.find_label lbl (fn x k) = Mach.find_label lbl k) -> forall (args: list A) k,
Mach.find_label lbl (List.fold_right fn k args) = Mach.find_label lbl k.
Proof.
- induction args; simpl. auto.
+ induction args; simpl. auto.
intros. rewrite H. auto.
Qed.
@@ -2119,10 +2119,10 @@ Remark find_label_save_callee_save:
Proof.
intros. unfold save_callee_save, save_callee_save_int, save_callee_save_float, save_callee_save_regs.
repeat rewrite find_label_fold_right. reflexivity.
- intros. unfold save_callee_save_reg.
+ intros. unfold save_callee_save_reg.
case (zlt (index_float_callee_save x) (fe_num_float_callee_save fe));
intro; reflexivity.
- intros. unfold save_callee_save_reg.
+ intros. unfold save_callee_save_reg.
case (zlt (index_int_callee_save x) (fe_num_int_callee_save fe));
intro; reflexivity.
Qed.
@@ -2133,10 +2133,10 @@ Remark find_label_restore_callee_save:
Proof.
intros. unfold restore_callee_save, restore_callee_save_int, restore_callee_save_float, restore_callee_save_regs.
repeat rewrite find_label_fold_right. reflexivity.
- intros. unfold restore_callee_save_reg.
+ intros. unfold restore_callee_save_reg.
case (zlt (index_float_callee_save x) (fe_num_float_callee_save fe));
intro; reflexivity.
- intros. unfold restore_callee_save_reg.
+ intros. unfold restore_callee_save_reg.
case (zlt (index_int_callee_save x) (fe_num_int_callee_save fe));
intro; reflexivity.
Qed.
@@ -2154,7 +2154,7 @@ Lemma find_label_transl_code:
Proof.
induction c; simpl; intros.
auto.
- rewrite transl_code_eq.
+ rewrite transl_code_eq.
destruct a; unfold transl_instr; auto.
destruct s; simpl; auto.
destruct s; simpl; auto.
@@ -2167,10 +2167,10 @@ Lemma transl_find_label:
forall f tf lbl c,
transf_function f = OK tf ->
Linear.find_label lbl f.(Linear.fn_code) = Some c ->
- Mach.find_label lbl tf.(Mach.fn_code) =
+ Mach.find_label lbl tf.(Mach.fn_code) =
Some (transl_code (make_env (function_bounds f)) c).
Proof.
- intros. rewrite (unfold_transf_function _ _ H). simpl.
+ intros. rewrite (unfold_transf_function _ _ H). simpl.
unfold transl_body. rewrite find_label_save_callee_save.
rewrite find_label_transl_code. rewrite H0. reflexivity.
Qed.
@@ -2180,7 +2180,7 @@ End LABELS.
(** Code tail property for Linear executions. *)
Lemma find_label_tail:
- forall lbl c c',
+ forall lbl c c',
Linear.find_label lbl c = Some c' -> is_tail c' c.
Proof.
induction c; simpl.
@@ -2197,7 +2197,7 @@ Lemma is_tail_save_callee_save_regs:
is_tail k (save_callee_save_regs bound number mkindex ty fe csl k).
Proof.
induction csl; intros; simpl. auto with coqlib.
- unfold save_callee_save_reg. destruct (zlt (number a) (bound fe)).
+ unfold save_callee_save_reg. destruct (zlt (number a) (bound fe)).
constructor; auto. auto.
Qed.
@@ -2214,7 +2214,7 @@ Lemma is_tail_restore_callee_save_regs:
is_tail k (restore_callee_save_regs bound number mkindex ty fe csl k).
Proof.
induction csl; intros; simpl. auto with coqlib.
- unfold restore_callee_save_reg. destruct (zlt (number a) (bound fe)).
+ unfold restore_callee_save_reg. destruct (zlt (number a) (bound fe)).
constructor; auto. auto.
Qed.
@@ -2241,7 +2241,7 @@ Lemma is_tail_transl_code:
forall fe c1 c2, is_tail c1 c2 -> is_tail (transl_code fe c1) (transl_code fe c2).
Proof.
induction 1; simpl. auto with coqlib.
- rewrite transl_code_eq.
+ rewrite transl_code_eq.
eapply is_tail_trans. eauto. apply is_tail_transl_instr.
Qed.
@@ -2251,7 +2251,7 @@ Lemma is_tail_transf_function:
is_tail c (Linear.fn_code f) ->
is_tail (transl_code (make_env (function_bounds f)) c) (fn_code tf).
Proof.
- intros. rewrite (unfold_transf_function _ _ H). simpl.
+ intros. rewrite (unfold_transf_function _ _ H). simpl.
unfold transl_body. eapply is_tail_trans. 2: apply is_tail_save_callee_save.
apply is_tail_transl_code; auto.
Qed.
@@ -2263,25 +2263,25 @@ Qed.
Lemma symbols_preserved:
forall id, Genv.find_symbol tge id = Genv.find_symbol ge id.
Proof.
- intros. unfold ge, tge.
+ intros. unfold ge, tge.
apply Genv.find_symbol_transf_partial with transf_fundef.
- exact TRANSF.
+ exact TRANSF.
Qed.
Lemma public_preserved:
forall id, Genv.public_symbol tge id = Genv.public_symbol ge id.
Proof.
- intros. unfold ge, tge.
+ intros. unfold ge, tge.
apply Genv.public_symbol_transf_partial with transf_fundef.
- exact TRANSF.
+ exact TRANSF.
Qed.
Lemma varinfo_preserved:
forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
Proof.
- intros. unfold ge, tge.
+ intros. unfold ge, tge.
apply Genv.find_var_info_transf_partial with transf_fundef.
- exact TRANSF.
+ exact TRANSF.
Qed.
Lemma functions_translated:
@@ -2305,7 +2305,7 @@ Lemma sig_preserved:
Proof.
intros until tf; unfold transf_fundef, transf_partial_fundef.
destruct f; intros; monadInv H.
- rewrite (unfold_transf_function _ _ EQ). auto.
+ rewrite (unfold_transf_function _ _ EQ). auto.
auto.
Qed.
@@ -2320,17 +2320,17 @@ Lemma find_function_translated:
/\ transf_fundef f = OK tf.
Proof.
intros until f; intros AG MS FF.
- exploit match_stacks_globalenvs; eauto. intros [hi MG].
+ exploit match_stacks_globalenvs; eauto. intros [hi MG].
destruct ros; simpl in FF.
- exploit Genv.find_funct_inv; eauto. intros [b EQ]. rewrite EQ in FF.
- rewrite Genv.find_funct_find_funct_ptr in FF.
+ exploit Genv.find_funct_inv; eauto. intros [b EQ]. rewrite EQ in FF.
+ rewrite Genv.find_funct_find_funct_ptr in FF.
exploit function_ptr_translated; eauto. intros [tf [A B]].
exists b; exists tf; split; auto. simpl.
generalize (AG m0). rewrite EQ. intro INJ. inv INJ.
- inv MG. rewrite DOMAIN in H2. inv H2. simpl. auto. eapply FUNCTIONS; eauto.
- destruct (Genv.find_symbol ge i) as [b|] eqn:?; try discriminate.
+ inv MG. rewrite DOMAIN in H2. inv H2. simpl. auto. eapply FUNCTIONS; eauto.
+ destruct (Genv.find_symbol ge i) as [b|] eqn:?; try discriminate.
exploit function_ptr_translated; eauto. intros [tf [A B]].
- exists b; exists tf; split; auto. simpl.
+ exists b; exists tf; split; auto. simpl.
rewrite symbols_preserved. auto.
Qed.
@@ -2358,21 +2358,21 @@ Proof.
intros.
assert (loc_argument_acceptable l). apply loc_arguments_acceptable with sg; auto.
destruct l; red in H0.
- exists (rs r); split. constructor. auto.
+ exists (rs r); split. constructor. auto.
destruct sl; try contradiction.
inv MS.
elim (H4 _ H).
unfold parent_sp.
assert (slot_valid f Outgoing pos ty = true).
- exploit loc_arguments_acceptable; eauto. intros [A B].
+ exploit loc_arguments_acceptable; eauto. intros [A B].
unfold slot_valid. unfold proj_sumbool. rewrite zle_true by omega.
destruct ty; auto; congruence.
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 pos ty); eauto.
+ constructor.
+ eapply index_contains_load_stack with (idx := FI_arg pos ty); eauto.
red in AGCS. rewrite AGCS; auto.
Qed.
@@ -2394,7 +2394,7 @@ Lemma transl_external_arguments:
extcall_arguments rs m' (parent_sp cs') sg vl /\
Val.inject_list j (ls ## (loc_arguments sg)) vl.
Proof.
- unfold extcall_arguments.
+ unfold extcall_arguments.
apply transl_external_arguments_rec.
auto with coqlib.
Qed.
@@ -2435,34 +2435,34 @@ Local Opaque fe offset_of_index.
- assert (loc_valid f x = true) by auto.
destruct x as [r | [] ofs ty]; try discriminate.
+ exists (rs r); auto with barg.
- + exploit agree_locals; eauto. intros [v [A B]]. inv A.
+ + exploit agree_locals; eauto. intros [v [A B]]. inv A.
exists v; split; auto. constructor. simpl. rewrite Int.add_zero_l.
Local Transparent fe.
- unfold fe, b. erewrite offset_of_index_no_overflow by eauto. exact H1.
+ unfold fe, b. erewrite offset_of_index_no_overflow by eauto. exact H1.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
-- simpl in H. exploit Mem.load_inject; eauto. eapply agree_inj; eauto.
+- simpl in H. exploit Mem.load_inject; eauto. eapply agree_inj; eauto.
intros (v' & A & B). exists v'; split; auto. constructor.
unfold Mem.loadv, Val.add. rewrite <- Int.add_assoc.
unfold fe, b; erewrite shifted_stack_offset_no_overflow; eauto.
- eapply agree_bounds; eauto. eapply Mem.valid_access_perm. eapply Mem.load_valid_access; eauto.
+ eapply agree_bounds; eauto. eapply Mem.valid_access_perm. eapply Mem.load_valid_access; eauto.
- econstructor; split; eauto with barg.
unfold Val.add. rewrite ! Int.add_zero_l. econstructor. eapply agree_inj; eauto. auto.
- assert (Val.inject j (Senv.symbol_address ge id ofs) (Senv.symbol_address ge id ofs)).
- { unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
+ { unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) eqn:FS; auto.
econstructor. eapply (proj1 GINJ); eauto. rewrite Int.add_zero; auto. }
exploit Mem.loadv_inject; eauto. intros (v' & A & B). exists v'; auto with barg.
-- econstructor; split; eauto with barg.
- unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
+- econstructor; split; eauto with barg.
+ unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) eqn:FS; auto.
econstructor. eapply (proj1 GINJ); eauto. rewrite Int.add_zero; auto.
- destruct IHeval_builtin_arg1 as (v1 & A1 & B1); auto using in_or_app.
destruct IHeval_builtin_arg2 as (v2 & A2 & B2); auto using in_or_app.
- exists (Val.longofwords v1 v2); split; auto with barg.
- apply Val.longofwords_inject; auto.
+ exists (Val.longofwords v1 v2); split; auto with barg.
+ apply Val.longofwords_inject; auto.
Qed.
Lemma transl_builtin_args_correct:
@@ -2478,7 +2478,7 @@ Proof.
- exists (@nil val); split; constructor.
- exploit transl_builtin_arg_correct; eauto using in_or_app. intros (v1' & A & B).
exploit IHlist_forall2; eauto using in_or_app. intros (vl' & C & D).
- exists (v1'::vl'); split; constructor; auto.
+ exists (v1'::vl'); split; constructor; auto.
Qed.
End BUILTIN_ARGUMENTS.
@@ -2528,7 +2528,7 @@ Inductive match_states: Linear.state -> Mach.state -> Prop :=
match_states (Linear.Callstate cs f ls m)
(Mach.Callstate cs' fb rs m')
| match_states_return:
- forall cs ls m cs' rs m' j sg
+ forall cs ls m cs' rs m' j sg
(MINJ: Mem.inject j m m')
(STACKS: match_stacks j m m' cs cs' sg (Mem.nextblock m) (Mem.nextblock m'))
(AGREGS: agree_regs j ls rs)
@@ -2546,7 +2546,7 @@ Proof.
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.
+ apply H. eapply is_tail_in; eauto.
*)
induction 1; intros;
try inv MS;
@@ -2562,7 +2562,7 @@ Proof.
+ (* Lgetstack, local *)
exploit agree_locals; eauto. intros [v [A B]].
econstructor; split.
- apply plus_one. apply exec_Mgetstack.
+ apply plus_one. apply exec_Mgetstack.
eapply index_contains_load_stack; eauto.
econstructor; eauto with coqlib.
apply agree_regs_set_reg; auto.
@@ -2575,12 +2575,12 @@ Proof.
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.
+ 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.
+ 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 ofs ty))
@@ -2588,21 +2588,21 @@ Proof.
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 caller_save_reg_within_bounds.
+ 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 caller_save_reg_within_bounds.
apply temp_for_parent_frame_caller_save.
+ (* Lgetstack, outgoing *)
exploit agree_outgoing; eauto. intros [v [A B]].
econstructor; split.
- apply plus_one. apply exec_Mgetstack.
+ apply plus_one. apply exec_Mgetstack.
eapply index_contains_load_stack; eauto.
econstructor; eauto with coqlib.
apply agree_regs_set_reg; auto.
apply agree_frame_set_reg; auto.
- (* Lsetstack *)
- exploit wt_state_setstack; eauto. intros (SV & SW).
+ exploit wt_state_setstack; eauto. intros (SV & SW).
set (idx := match sl with
| Local => FI_local ofs ty
| Incoming => FI_link (*dummy*)
@@ -2619,22 +2619,22 @@ Proof.
apply plus_one. destruct sl; simpl in SW.
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.
+ 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 H2.
exploit offset_of_index_disj_stack_data_2; eauto.
- exploit agree_bounds. eauto. apply Mem.perm_cur_max. eauto.
+ exploit agree_bounds. eauto. apply Mem.perm_cur_max. eauto.
omega.
apply match_stacks_change_mach_mem with m'; auto.
- eauto with mem. eauto with mem. intros. rewrite <- H1; eapply Mem.load_store_other; eauto. left; apply Plt_ne; auto.
- eauto. eauto.
- apply agree_regs_set_slot. apply agree_regs_undef_regs; auto.
+ eauto with mem. eauto with mem. intros. rewrite <- H1; eapply Mem.load_store_other; eauto. left; apply Plt_ne; auto.
+ eauto. eauto.
+ apply agree_regs_set_slot. apply agree_regs_undef_regs; auto.
destruct sl.
+ eapply agree_frame_set_local. eapply agree_frame_undef_locs; eauto.
apply destroyed_by_setstack_caller_save. auto. auto. auto.
- assumption.
+ assumption.
+ simpl in SW; discriminate.
+ eapply agree_frame_set_outgoing. eapply agree_frame_undef_locs; eauto.
apply destroyed_by_setstack_caller_save. auto. auto. auto.
@@ -2649,13 +2649,13 @@ Proof.
eapply match_stacks_preserves_globals; eauto.
eapply agree_inj; eauto. eapply agree_reglist; eauto.
destruct H0 as [v' [A B]].
- econstructor; split.
- apply plus_one. econstructor.
- instantiate (1 := v'). rewrite <- A. apply eval_operation_preserved.
+ econstructor; split.
+ apply plus_one. econstructor.
+ instantiate (1 := v'). rewrite <- A. apply eval_operation_preserved.
exact symbols_preserved. eauto.
econstructor; eauto with coqlib.
apply agree_regs_set_reg; auto.
- rewrite transl_destroyed_by_op. apply agree_regs_undef_regs; 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.
@@ -2663,38 +2663,38 @@ Proof.
assert (exists a',
eval_addressing ge (Vptr sp' Int.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
/\ Val.inject j a a').
- eapply eval_addressing_inject; eauto.
+ 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]].
exploit Mem.loadv_inject; eauto. intros [v' [C D]].
- econstructor; split.
- apply plus_one. econstructor.
+ econstructor; split.
+ apply plus_one. econstructor.
instantiate (1 := a'). rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved.
eexact C. eauto.
econstructor; eauto with coqlib.
- apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto.
+ 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.
+ apply destroyed_by_load_caller_save. auto.
- (* Lstore *)
assert (exists a',
eval_addressing ge (Vptr sp' Int.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
/\ Val.inject j a a').
- eapply eval_addressing_inject; eauto.
+ 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]].
exploit Mem.storev_mapped_inject; eauto. intros [m1' [C D]].
- econstructor; split.
- apply plus_one. econstructor.
+ econstructor; split.
+ apply plus_one. econstructor.
instantiate (1 := a'). rewrite <- A. apply eval_addressing_preserved. exact symbols_preserved.
eexact C. eauto.
econstructor. eauto.
- eapply match_stacks_parallel_stores. eexact MINJ. eexact B. eauto. eauto. auto.
- eauto. eauto.
- rewrite transl_destroyed_by_store.
- apply agree_regs_undef_regs; auto.
+ eapply match_stacks_parallel_stores. eexact MINJ. eexact B. eauto. eauto. auto.
+ eauto. eauto.
+ 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.
@@ -2727,23 +2727,23 @@ Proof.
econstructor; eauto.
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_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; apply Plt_ne; auto.
- intros. rewrite <- H1. eapply Mem.load_free; eauto. left; apply Plt_ne; auto.
- eauto with mem. intros. eapply Mem.perm_free_3; eauto.
- apply Plt_Ple. change (Mem.valid_block m' stk). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_linear; eauto.
- apply Plt_Ple. change (Mem.valid_block m1' sp'). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_mach; eauto.
- apply zero_size_arguments_tailcall_possible. eapply wt_state_tailcall; eauto.
+ auto.
+ eauto with mem. intros. eapply Mem.perm_free_1; eauto. left; apply Plt_ne; auto.
+ intros. rewrite <- H1. eapply Mem.load_free; eauto. left; apply Plt_ne; auto.
+ eauto with mem. intros. eapply Mem.perm_free_3; eauto.
+ apply Plt_Ple. change (Mem.valid_block m' stk). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_linear; eauto.
+ apply Plt_Ple. change (Mem.valid_block m1' sp'). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_mach; eauto.
+ apply zero_size_arguments_tailcall_possible. eapply wt_state_tailcall; eauto.
- (* Lbuiltin *)
destruct BOUND as [BND1 BND2].
exploit transl_builtin_args_correct; eauto.
eapply match_stacks_preserves_globals; eauto.
rewrite <- forallb_forall. eapply wt_state_builtin; eauto.
- intros [vargs' [P Q]].
- exploit external_call_mem_inject; eauto.
+ intros [vargs' [P Q]].
+ 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.
@@ -2756,7 +2756,7 @@ Proof.
apply Plt_Ple. change (Mem.valid_block m sp0). eapply agree_valid_linear; eauto.
apply Plt_Ple. change (Mem.valid_block m'0 sp'). eapply agree_valid_mach; eauto.
apply agree_regs_set_res; auto. apply agree_regs_undef_regs; auto. eapply agree_regs_inject_incr; eauto.
- eapply agree_frame_inject_incr; eauto.
+ eapply agree_frame_inject_incr; eauto.
apply agree_frame_set_res; auto. apply agree_frame_undef_regs; auto.
apply agree_frame_extcall_invariant with m m'0; auto.
eapply external_call_valid_block; eauto.
@@ -2773,7 +2773,7 @@ Proof.
econstructor; split.
apply plus_one; eapply exec_Mgoto; eauto.
apply transl_find_label; eauto.
- econstructor; eauto.
+ econstructor; eauto.
eapply find_label_tail; eauto.
- (* Lcond, true *)
@@ -2782,8 +2782,8 @@ Proof.
eapply eval_condition_inject; eauto. eapply agree_reglist; eauto.
eapply transl_find_label; eauto.
econstructor. eauto. eauto. eauto. eauto.
- apply agree_regs_undef_regs; auto.
- apply agree_frame_undef_locs; auto. apply destroyed_by_cond_caller_save.
+ 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 *)
@@ -2791,15 +2791,15 @@ Proof.
apply plus_one. eapply exec_Mcond_false; eauto.
eapply eval_condition_inject; eauto. eapply agree_reglist; eauto.
econstructor. eauto. eauto. eauto. eauto.
- apply agree_regs_undef_regs; auto.
- apply agree_frame_undef_locs; auto. apply destroyed_by_cond_caller_save.
+ 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).
{ generalize (AGREGS arg). rewrite H. intro IJ; inv IJ; auto. }
econstructor; split.
- apply plus_one; eapply exec_Mjumptable; eauto.
+ apply plus_one; eapply exec_Mjumptable; eauto.
apply transl_find_label; eauto.
econstructor. eauto. eauto. eauto. eauto.
apply agree_regs_undef_regs; auto.
@@ -2814,37 +2814,37 @@ Proof.
traceEq.
econstructor; eauto.
apply match_stacks_change_bounds with stk sp'.
- apply match_stacks_change_linear_mem with m.
+ apply match_stacks_change_linear_mem with m.
apply match_stacks_change_mach_mem with m'0.
- eauto.
- eauto with mem. intros. eapply Mem.perm_free_1; eauto. left; apply Plt_ne; auto.
- intros. rewrite <- H1. eapply Mem.load_free; eauto. left; apply Plt_ne; auto.
- eauto with mem. intros. eapply Mem.perm_free_3; eauto.
- apply Plt_Ple. change (Mem.valid_block m' stk). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_linear; eauto.
- apply Plt_Ple. change (Mem.valid_block m1' sp'). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_mach; eauto.
+ eauto.
+ eauto with mem. intros. eapply Mem.perm_free_1; eauto. left; apply Plt_ne; auto.
+ intros. rewrite <- H1. eapply Mem.load_free; eauto. left; apply Plt_ne; auto.
+ eauto with mem. intros. eapply Mem.perm_free_3; eauto.
+ apply Plt_Ple. change (Mem.valid_block m' stk). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_linear; eauto.
+ apply Plt_Ple. change (Mem.valid_block m1' sp'). eapply Mem.valid_block_free_1; eauto. eapply agree_valid_mach; eauto.
- (* internal function *)
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.
exploit function_prologue_correct; eauto. eapply wt_callstate_wt_regs; eauto.
- eapply match_stacks_type_sp; eauto.
+ eapply match_stacks_type_sp; eauto.
eapply match_stacks_type_retaddr; eauto.
intros [j' [rs' [m2' [sp' [m3' [m4' [m5' [A [B [C [D [E [F [G [J [K L]]]]]]]]]]]]]]]].
econstructor; split.
- eapply plus_left. econstructor; eauto.
- rewrite (unfold_transf_function _ _ TRANSL). unfold fn_code. unfold transl_body.
+ eapply plus_left. econstructor; eauto.
+ rewrite (unfold_transf_function _ _ TRANSL). unfold fn_code. unfold transl_body.
eexact D. traceEq.
- generalize (Mem.alloc_result _ _ _ _ _ H). intro SP_EQ.
+ generalize (Mem.alloc_result _ _ _ _ _ H). intro SP_EQ.
generalize (Mem.alloc_result _ _ _ _ _ A). intro SP'_EQ.
- econstructor; eauto.
+ econstructor; eauto.
apply match_stacks_change_mach_mem with m'0.
apply match_stacks_change_linear_mem with m.
rewrite SP_EQ; rewrite SP'_EQ.
- eapply match_stacks_change_meminj; eauto. apply Ple_refl.
- eauto with mem. intros. exploit Mem.perm_alloc_inv. eexact H. eauto.
- rewrite dec_eq_false; auto. apply Plt_ne; auto.
- intros. eapply stores_in_frame_valid; eauto with mem.
+ eapply match_stacks_change_meminj; eauto. apply Ple_refl.
+ eauto with mem. intros. exploit Mem.perm_alloc_inv. eexact H. eauto.
+ rewrite dec_eq_false; auto. apply Plt_ne; auto.
+ intros. eapply stores_in_frame_valid; eauto with mem.
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.
@@ -2853,7 +2853,7 @@ Proof.
- (* external function *)
simpl in TRANSL. inversion TRANSL; subst tf.
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.
@@ -2862,18 +2862,18 @@ Proof.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
econstructor; eauto.
apply match_stacks_change_bounds with (Mem.nextblock m) (Mem.nextblock m'0).
- inv H0; inv A. eapply match_stack_change_extcall; eauto. apply Ple_refl. apply Ple_refl.
+ inv H0; inv A. eapply match_stack_change_extcall; eauto. apply Ple_refl. apply Ple_refl.
eapply external_call_nextblock'; eauto.
eapply external_call_nextblock'; eauto.
- apply agree_regs_set_regs; auto. apply agree_regs_inject_incr with j; auto.
- apply agree_callee_save_set_result; auto.
+ apply agree_regs_set_regs; auto. apply agree_regs_inject_incr with j; auto.
+ apply agree_callee_save_set_result; auto.
- (* return *)
- inv STACKS. simpl in AGLOCS.
+ inv STACKS. simpl in AGLOCS.
econstructor; split.
- apply plus_one. apply exec_return.
+ apply plus_one. apply exec_return.
econstructor; eauto.
- apply agree_frame_return with rs0; auto.
+ apply agree_frame_return with rs0; auto.
Qed.
Lemma transf_initial_states:
@@ -2883,13 +2883,13 @@ Proof.
intros. inv H.
exploit function_ptr_translated; eauto. intros [tf [FIND TR]].
econstructor; split.
- econstructor.
+ econstructor.
eapply Genv.init_mem_transf_partial; eauto.
- rewrite (transform_partial_program_main _ _ TRANSF).
+ rewrite (transform_partial_program_main _ _ TRANSF).
rewrite symbols_preserved. eauto.
econstructor; eauto.
eapply Genv.initmem_inject; eauto.
- apply match_stacks_empty with (Mem.nextblock m0). apply Ple_refl. apply Ple_refl.
+ apply match_stacks_empty with (Mem.nextblock m0). apply Ple_refl. apply Ple_refl.
constructor.
intros. unfold Mem.flat_inj. apply pred_dec_true; auto.
unfold Mem.flat_inj; intros. destruct (plt b1 (Mem.nextblock m0)); congruence.
@@ -2902,21 +2902,21 @@ Proof.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> Linear.final_state st1 r -> Mach.final_state st2 r.
Proof.
intros. inv H0. inv H. inv STACKS.
- generalize (AGREGS r0). rewrite H2. intros A; inv A.
- econstructor; eauto.
+ generalize (AGREGS r0). rewrite H2. intros A; inv A.
+ econstructor; eauto.
Qed.
Lemma wt_prog:
forall i fd, In (i, Gfun fd) prog.(prog_defs) -> wt_fundef fd.
Proof.
- intros. exploit transform_partial_program_succeeds; eauto.
- intros [tfd TF]. destruct fd; simpl in *.
+ intros. exploit transform_partial_program_succeeds; eauto.
+ intros [tfd TF]. destruct fd; simpl in *.
- monadInv TF. unfold transf_function in EQ.
- destruct (wt_function f). auto. discriminate.
+ destruct (wt_function f). auto. discriminate.
- auto.
Qed.
@@ -2924,16 +2924,16 @@ Theorem transf_program_correct:
forward_simulation (Linear.semantics prog) (Mach.semantics return_address_offset tprog).
Proof.
set (ms := fun s s' => wt_state s /\ match_states s s').
- eapply forward_simulation_plus with (match_states := ms).
+ eapply forward_simulation_plus with (match_states := ms).
- exact public_preserved.
-- intros. exploit transf_initial_states; eauto. intros [st2 [A B]].
+- intros. exploit transf_initial_states; eauto. intros [st2 [A B]].
exists st2; split; auto. split; auto.
- apply wt_initial_state with (prog := prog); auto. exact wt_prog.
-- intros. destruct H. eapply transf_final_states; eauto.
-- intros. destruct H0.
+ apply wt_initial_state with (prog := prog); auto. exact wt_prog.
+- intros. destruct H. eapply transf_final_states; eauto.
+- intros. destruct H0.
exploit transf_step_correct; eauto. intros [s2' [A B]].
exists s2'; split. exact A. split.
- eapply step_type_preservation; eauto. eexact wt_prog. eexact H.
+ eapply step_type_preservation; eauto. eexact wt_prog. eexact H.
auto.
Qed.
diff --git a/backend/Tailcall.v b/backend/Tailcall.v
index db246fec..e8ce9e25 100644
--- a/backend/Tailcall.v
+++ b/backend/Tailcall.v
@@ -27,7 +27,7 @@ Require Import Conventions.
If the current function had a non-empty stack block, it could be that
the called function accesses it, for instance if a pointer into the
-stack block is passed as an argument. In this case, it would be
+stack block is passed as an argument. In this case, it would be
semantically incorrect to deallocate the stack block before the call,
as [Itailcall] does. Therefore, the optimization can only be performed if
the stack block of the current function is empty, in which case it makes
@@ -47,7 +47,7 @@ The general pattern we recognize is therefore:
The [is_return] function below recognizes this pattern.
*)
-Fixpoint is_return (n: nat) (f: function) (pc: node) (rret: reg)
+Fixpoint is_return (n: nat) (f: function) (pc: node) (rret: reg)
{struct n}: bool :=
match n with
| O => false
diff --git a/backend/Tailcallproof.v b/backend/Tailcallproof.v
index 1c25d244..7e7b7b53 100644
--- a/backend/Tailcallproof.v
+++ b/backend/Tailcallproof.v
@@ -59,7 +59,7 @@ Proof.
assert (forall n pc, (return_measure_rec n f pc <= n)%nat).
induction n; intros; simpl.
omega.
- destruct (f!pc); try omega.
+ destruct (f!pc); try omega.
destruct i; try omega.
generalize (IHn n0). omega.
generalize (IHn n0). omega.
@@ -125,28 +125,28 @@ Lemma is_return_charact:
Proof.
induction n; intros.
simpl in H. congruence.
- generalize H. simpl.
+ generalize H. simpl.
caseEq ((fn_code f)!pc); try congruence.
intro i. caseEq i; try congruence.
intros s; intros. eapply is_return_nop; eauto. eapply IHn; eauto. omega.
unfold return_measure.
rewrite <- (is_return_measure_rec f (S n) niter pc rret); auto.
- rewrite <- (is_return_measure_rec f n niter s rret); auto.
+ rewrite <- (is_return_measure_rec f n niter s rret); auto.
simpl. rewrite H2. omega. omega.
- intros op args dst s EQ1 EQ2.
+ intros op args dst s EQ1 EQ2.
caseEq (is_move_operation op args); try congruence.
intros src IMO. destruct (Reg.eq rret src); try congruence.
- subst rret. intro.
- exploit is_move_operation_correct; eauto. intros [A B]. subst.
+ subst rret. intro.
+ exploit is_move_operation_correct; eauto. intros [A B]. subst.
eapply is_return_move; eauto. eapply IHn; eauto. omega.
unfold return_measure.
rewrite <- (is_return_measure_rec f (S n) niter pc src); auto.
- rewrite <- (is_return_measure_rec f n niter s dst); auto.
+ rewrite <- (is_return_measure_rec f n niter s dst); auto.
simpl. rewrite EQ2. omega. omega.
-
- intros or EQ1 EQ2. destruct or; intros.
- assert (r = rret). eapply proj_sumbool_true; eauto. subst r.
+
+ intros or EQ1 EQ2. destruct or; intros.
+ assert (r = rret). eapply proj_sumbool_true; eauto. subst r.
apply is_return_some; auto.
apply is_return_none; auto.
Qed.
@@ -172,7 +172,7 @@ Proof.
opt_typ_eq (sig_res s) (sig_res (fn_sig f))); intros.
destruct (andb_prop _ _ H0). destruct (andb_prop _ _ H1).
eapply transf_instr_tailcall; eauto.
- eapply is_return_charact; eauto.
+ eapply is_return_charact; eauto.
constructor.
Qed.
@@ -183,8 +183,8 @@ Lemma transf_instr_lookup:
Proof.
intros. unfold transf_function.
destruct (zeq (fn_stacksize f) 0).
- simpl. rewrite PTree.gmap. rewrite H. simpl.
- exists (transf_instr f pc i); split. auto. apply transf_instr_charact; auto.
+ simpl. rewrite PTree.gmap. rewrite H. simpl.
+ exists (transf_instr f pc i); split. auto. apply transf_instr_charact; auto.
exists i; split. auto. constructor.
Qed.
@@ -246,14 +246,14 @@ Proof (@Genv.find_funct_ptr_transf _ _ _ transf_fundef prog).
Lemma sig_preserved:
forall f, funsig (transf_fundef f) = funsig f.
Proof.
- destruct f; auto. simpl. unfold transf_function.
- destruct (zeq (fn_stacksize f) 0); auto.
+ destruct f; auto. simpl. unfold transf_function.
+ destruct (zeq (fn_stacksize f) 0); auto.
Qed.
Lemma stacksize_preserved:
forall f, fn_stacksize (transf_function f) = fn_stacksize f.
Proof.
- unfold transf_function. intros.
+ unfold transf_function. intros.
destruct (zeq (fn_stacksize f) 0); auto.
Qed.
@@ -410,33 +410,33 @@ Proof.
induction 1; intros; inv MS; EliminatedInstr.
(* nop *)
- TransfInstr. left. econstructor; split.
+ TransfInstr. left. econstructor; split.
eapply exec_Inop; eauto. constructor; auto.
(* eliminated nop *)
assert (s0 = pc') by congruence. subst s0.
- right. split. simpl. omega. split. auto.
- econstructor; eauto.
+ right. split. simpl. omega. split. auto.
+ econstructor; eauto.
(* op *)
TransfInstr.
- assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
- exploit eval_operation_lessdef; eauto.
- intros [v' [EVAL' VLD]].
+ assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
+ exploit eval_operation_lessdef; eauto.
+ intros [v' [EVAL' VLD]].
left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' (rs'#res <- v') m'); split.
eapply exec_Iop; eauto. rewrite <- EVAL'.
apply eval_operation_preserved. exact symbols_preserved.
econstructor; eauto. apply set_reg_lessdef; auto.
(* eliminated move *)
- rewrite H1 in H. clear H1. inv H.
+ rewrite H1 in H. clear H1. inv H.
right. split. simpl. omega. split. auto.
- econstructor; eauto. simpl in H0. rewrite PMap.gss. congruence.
+ econstructor; eauto. simpl in H0. rewrite PMap.gss. congruence.
(* load *)
TransfInstr.
- assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
- exploit eval_addressing_lessdef; eauto.
+ assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
+ exploit eval_addressing_lessdef; eauto.
intros [a' [ADDR' ALD]].
- exploit Mem.loadv_extends; eauto.
+ exploit Mem.loadv_extends; eauto.
intros [v' [LOAD' VLD]].
left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' (rs'#dst <- v') m'); split.
eapply exec_Iload with (a := a'). eauto. rewrite <- ADDR'.
@@ -445,10 +445,10 @@ Proof.
(* store *)
TransfInstr.
- assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
- exploit eval_addressing_lessdef; eauto.
+ assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
+ exploit eval_addressing_lessdef; eauto.
intros [a' [ADDR' ALD]].
- exploit Mem.storev_extends. 2: eexact H1. eauto. eauto. apply RLD.
+ exploit Mem.storev_extends. 2: eexact H1. eauto. eauto. apply RLD.
intros [m'1 [STORE' MLD']].
left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' rs' m'1); split.
eapply exec_Istore with (a := a'). eauto. rewrite <- ADDR'.
@@ -457,32 +457,32 @@ Proof.
econstructor; eauto.
(* call *)
- exploit find_function_translated; eauto. intro FIND'.
+ exploit find_function_translated; eauto. intro FIND'.
TransfInstr.
(* call turned tailcall *)
assert ({ m'' | Mem.free m' sp0 0 (fn_stacksize (transf_function f)) = Some m''}).
- apply Mem.range_perm_free. rewrite stacksize_preserved. rewrite H7.
+ apply Mem.range_perm_free. rewrite stacksize_preserved. rewrite H7.
red; intros; omegaContradiction.
destruct X as [m'' FREE].
left. exists (Callstate s' (transf_fundef fd) (rs'##args) m''); split.
- eapply exec_Itailcall; eauto. apply sig_preserved.
+ eapply exec_Itailcall; eauto. apply sig_preserved.
constructor. eapply match_stackframes_tail; eauto. apply regs_lessdef_regs; auto.
eapply Mem.free_right_extends; eauto.
rewrite stacksize_preserved. rewrite H7. intros. omegaContradiction.
(* call that remains a call *)
left. exists (Callstate (Stackframe res (transf_function f) (Vptr sp0 Int.zero) pc' rs' :: s')
(transf_fundef fd) (rs'##args) m'); split.
- eapply exec_Icall; eauto. apply sig_preserved.
- constructor. constructor; auto. apply regs_lessdef_regs; auto. auto.
+ eapply exec_Icall; eauto. apply sig_preserved.
+ constructor. constructor; auto. apply regs_lessdef_regs; auto. auto.
-(* tailcall *)
+(* tailcall *)
exploit find_function_translated; eauto. intro FIND'.
exploit Mem.free_parallel_extends; eauto. intros [m'1 [FREE EXT]].
TransfInstr.
left. exists (Callstate s' (transf_fundef fd) (rs'##args) m'1); split.
eapply exec_Itailcall; eauto. apply sig_preserved.
rewrite stacksize_preserved; auto.
- constructor. auto. apply regs_lessdef_regs; auto. auto.
+ constructor. auto. apply regs_lessdef_regs; auto. auto.
(* builtin *)
TransfInstr.
@@ -498,18 +498,18 @@ Proof.
econstructor; eauto. apply set_res_lessdef; auto.
(* cond *)
- TransfInstr.
+ TransfInstr.
left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) (if b then ifso else ifnot) rs' m'); split.
eapply exec_Icond; eauto.
apply eval_condition_lessdef with (rs##args) m; auto. apply regs_lessdef_regs; auto.
- constructor; auto.
+ constructor; auto.
(* jumptable *)
- TransfInstr.
+ TransfInstr.
left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' rs' m'); split.
eapply exec_Ijumptable; eauto.
generalize (RLD arg). rewrite H0. intro. inv H2. auto.
- constructor; auto.
+ constructor; auto.
(* return *)
exploit Mem.free_parallel_extends; eauto. intros [m'1 [FREE EXT]].
@@ -521,11 +521,11 @@ Proof.
auto.
(* eliminated return None *)
- assert (or = None) by congruence. subst or.
- right. split. simpl. omega. split. auto.
+ assert (or = None) by congruence. subst or.
+ right. split. simpl. omega. split. auto.
constructor. auto.
simpl. constructor.
- eapply Mem.free_left_extends; eauto.
+ eapply Mem.free_left_extends; eauto.
(* eliminated return Some *)
assert (or = Some r) by congruence. subst or.
@@ -542,12 +542,12 @@ Proof.
assert (fn_stacksize (transf_function f) = fn_stacksize f /\
fn_entrypoint (transf_function f) = fn_entrypoint f /\
fn_params (transf_function f) = fn_params f).
- unfold transf_function. destruct (zeq (fn_stacksize f) 0); auto.
- destruct H0 as [EQ1 [EQ2 EQ3]].
+ unfold transf_function. destruct (zeq (fn_stacksize f) 0); auto.
+ destruct H0 as [EQ1 [EQ2 EQ3]].
left. econstructor; split.
simpl. eapply exec_function_internal; eauto. rewrite EQ1; eauto.
rewrite EQ2. rewrite EQ3. constructor; auto.
- apply regs_lessdef_init_regs. auto.
+ apply regs_lessdef_init_regs. auto.
(* external call *)
exploit external_call_mem_extends; eauto.
@@ -556,29 +556,29 @@ Proof.
simpl. econstructor; eauto.
eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- constructor; auto.
+ constructor; auto.
(* returnstate *)
- inv H2.
+ inv H2.
(* synchronous return in both programs *)
- left. econstructor; split.
- apply exec_return.
- constructor; auto. apply set_reg_lessdef; auto.
+ left. econstructor; split.
+ apply exec_return.
+ constructor; auto. apply set_reg_lessdef; auto.
(* return instr in source program, eliminated because of tailcall *)
- right. split. unfold measure. simpl length.
+ right. split. unfold measure. simpl length.
change (S (length s) * (niter + 2))%nat
- with ((niter + 2) + (length s) * (niter + 2))%nat.
- generalize (return_measure_bounds (fn_code f) pc). omega.
- split. auto.
+ with ((niter + 2) + (length s) * (niter + 2))%nat.
+ generalize (return_measure_bounds (fn_code f) pc). omega.
+ split. auto.
econstructor; eauto.
- rewrite Regmap.gss. auto.
+ rewrite Regmap.gss. auto.
Qed.
Lemma transf_initial_states:
forall st1, initial_state prog st1 ->
exists st2, initial_state tprog st2 /\ match_states st1 st2.
Proof.
- intros. inv H.
+ intros. inv H.
exploit funct_ptr_translated; eauto. intro FIND.
exists (Callstate nil (transf_fundef f) nil m0); split.
econstructor; eauto. apply Genv.init_mem_transf. auto.
@@ -586,14 +586,14 @@ Proof.
rewrite symbols_preserved. eauto.
reflexivity.
rewrite <- H3. apply sig_preserved.
- constructor. constructor. constructor. apply Mem.extends_refl.
+ constructor. constructor. constructor. apply Mem.extends_refl.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
- intros. inv H0. inv H. inv H5. inv H3. constructor.
+ intros. inv H0. inv H. inv H5. inv H3. constructor.
Qed.
@@ -607,7 +607,7 @@ Proof.
eexact public_preserved.
eexact transf_initial_states.
eexact transf_final_states.
- exact transf_step_correct.
+ exact transf_step_correct.
Qed.
End PRESERVATION.
diff --git a/backend/Tunneling.v b/backend/Tunneling.v
index bdc8117e..fa7ff787 100644
--- a/backend/Tunneling.v
+++ b/backend/Tunneling.v
@@ -55,7 +55,7 @@ Require Import LTL.
<< L1: nop L2;
L2: nop L1;
>>
- Coq warns us of this fact by not accepting the definition
+ Coq warns us of this fact by not accepting the definition
of [branch_target] above.
To handle this problem, we proceed in two passes. The first pass
diff --git a/backend/Tunnelingproof.v b/backend/Tunnelingproof.v
index e9e4856e..22f0521e 100644
--- a/backend/Tunnelingproof.v
+++ b/backend/Tunnelingproof.v
@@ -65,8 +65,8 @@ Proof.
red; intros; simpl. rewrite PTree.gempty. apply U.repr_empty.
(* inductive case *)
- intros m uf pc bb; intros. destruct uf as [u f].
- assert (PC: U.repr u pc = pc).
+ 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 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)).
@@ -75,27 +75,27 @@ Proof.
(* u and f are unchanged *)
rewrite B.
- red. intro pc'. simpl. rewrite PTree.gsspec. destruct (peq pc' pc). subst pc'.
+ red. intro pc'. simpl. rewrite PTree.gsspec. destruct (peq pc' pc). subst pc'.
destruct bb; auto. destruct i; auto.
- apply H1.
+ apply H1.
(* 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.
(* The new instruction *)
- rewrite (U.repr_union_2 u pc s); auto. rewrite U.repr_union_3.
+ rewrite (U.repr_union_2 u pc s); auto. rewrite U.repr_union_3.
unfold measure_edge. destruct (peq (U.repr u s) pc). auto. right. split. auto.
rewrite PC. rewrite peq_true. omega.
(* 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.
+ intro. rewrite <- H2 at 2. apply U.repr_union_1. congruence.
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 s0) pc). omega. auto.
+ rewrite P. destruct (peq (U.repr u s0) pc). omega. auto.
Qed.
Definition record_gotos' (f: function) :=
@@ -104,12 +104,12 @@ Definition record_gotos' (f: function) :=
Lemma record_gotos_gotos':
forall f, fst (record_gotos' f) = record_gotos f.
Proof.
- intros. unfold record_gotos', record_gotos.
+ intros. unfold record_gotos', record_gotos.
repeat rewrite PTree.fold_spec.
generalize (PTree.elements (fn_code f)) (U.empty) (fun _ : node => O).
induction l; intros; simpl.
auto.
- unfold record_goto' at 2. unfold record_goto at 2.
+ unfold record_goto' at 2. unfold record_goto at 2.
destruct (snd a). apply IHl. destruct i; apply IHl.
Qed.
@@ -128,7 +128,7 @@ Theorem record_gotos_correct:
| _ => branch_target f pc = pc
end.
Proof.
- intros.
+ intros.
generalize (record_gotos'_correct f.(fn_code) pc). simpl.
fold (record_gotos' f). unfold branch_map_correct, branch_target, count_gotos.
rewrite record_gotos_gotos'. auto.
@@ -284,14 +284,14 @@ Proof.
(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.
+ 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.
+ 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.
@@ -302,7 +302,7 @@ Proof.
econstructor; eauto.
(* Lload *)
left; simpl; econstructor; split.
- eapply exec_Lload with (a := a).
+ eapply exec_Lload with (a := a).
rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
eauto. eauto.
econstructor; eauto.
@@ -321,24 +321,24 @@ Proof.
eauto. eauto.
econstructor; eauto.
(* Lcall *)
- left; simpl; econstructor; split.
+ 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.
constructor; auto.
(* Ltailcall *)
- left; simpl; econstructor; split.
+ left; simpl; econstructor; split.
eapply exec_Ltailcall with (fd := tunnel_fundef fd); eauto.
- erewrite match_parent_locset; eauto.
+ erewrite match_parent_locset; eauto.
apply find_function_translated; auto.
apply sig_preserved.
erewrite <- match_parent_locset; eauto.
econstructor; eauto.
(* Lbuiltin *)
left; simpl; econstructor; split.
- eapply exec_Lbuiltin; eauto.
+ eapply exec_Lbuiltin; eauto.
eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
eapply external_call_symbols_preserved. eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
@@ -346,10 +346,10 @@ Proof.
(* Lbranch (preserved) *)
left; simpl; econstructor; split.
- eapply exec_Lbranch; eauto.
+ eapply exec_Lbranch; eauto.
fold (branch_target f pc). econstructor; eauto.
(* Lbranch (eliminated) *)
- right; split. simpl. omega. split. auto. constructor; auto.
+ right; split. simpl. omega. split. auto. constructor; auto.
(* Lcond *)
left; simpl; econstructor; split.
@@ -357,9 +357,9 @@ Proof.
destruct b; econstructor; eauto.
(* Ljumptable *)
left; simpl; econstructor; split.
- eapply exec_Ljumptable.
+ eapply exec_Ljumptable.
eauto. rewrite list_nth_z_map. change U.elt with node. rewrite H0. reflexivity. eauto.
- econstructor; eauto.
+ econstructor; eauto.
(* Lreturn *)
left; simpl; econstructor; split.
eapply exec_Lreturn; eauto.
@@ -368,13 +368,13 @@ Proof.
(* internal function *)
left; simpl; econstructor; split.
eapply exec_function_internal; eauto.
- simpl. econstructor; eauto.
+ simpl. econstructor; eauto.
(* external function *)
left; simpl; econstructor; split.
eapply exec_function_external; eauto.
eapply external_call_symbols_preserved'; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
- simpl. econstructor; eauto.
+ simpl. econstructor; eauto.
(* return *)
inv H3. inv H1.
left; econstructor; split.
@@ -386,22 +386,22 @@ Lemma transf_initial_states:
forall st1, initial_state prog st1 ->
exists st2, initial_state tprog st2 /\ match_states st1 st2.
Proof.
- intros. inversion H.
+ intros. inversion H.
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).
rewrite symbols_preserved. eauto.
apply function_ptr_translated; auto.
- rewrite <- H3. apply sig_preserved.
+ rewrite <- H3. apply sig_preserved.
constructor. constructor.
Qed.
Lemma transf_final_states:
- forall st1 st2 r,
+ forall st1 st2 r,
match_states st1 st2 -> final_state st1 r -> final_state st2 r.
Proof.
- intros. inv H0. inv H. inv H6. econstructor; eauto.
+ intros. inv H0. inv H. inv H6. econstructor; eauto.
Qed.
Theorem transf_program_correct:
@@ -411,7 +411,7 @@ Proof.
eexact public_preserved.
eexact transf_initial_states.
eexact transf_final_states.
- eexact tunnel_step_correct.
+ eexact tunnel_step_correct.
Qed.
End PRESERVATION.
diff --git a/backend/Unusedglobproof.v b/backend/Unusedglobproof.v
index 4d7547f0..7c1b60a9 100644
--- a/backend/Unusedglobproof.v
+++ b/backend/Unusedglobproof.v
@@ -37,7 +37,7 @@ Module ISP := FSetProperties.Properties(IS).
(** Monotonic evolution of the workset. *)
-Inductive workset_incl (w1 w2: workset) : Prop :=
+Inductive workset_incl (w1 w2: workset) : Prop :=
workset_incl_intro:
forall (SEEN: IS.Subset w1.(w_seen) w2.(w_seen))
(TODO: List.incl w1.(w_todo) w2.(w_todo))
@@ -53,7 +53,7 @@ Qed.
Lemma workset_incl_refl: forall w, workset_incl w w.
Proof.
- intros; split. red; auto. red; auto. auto.
+ intros; split. red; auto. red; auto. auto.
Qed.
Lemma workset_incl_trans:
@@ -62,7 +62,7 @@ Proof.
intros. destruct H, H0; split.
red; eauto.
red; eauto.
- intros. edestruct TRACK0; eauto. edestruct TRACK; eauto.
+ intros. edestruct TRACK0; eauto. edestruct TRACK; eauto.
Qed.
Lemma add_workset_incl:
@@ -73,13 +73,13 @@ Proof.
- split; simpl.
+ red; intros. apply IS.add_2; auto.
+ red; simpl; auto.
- + intros. destruct (ident_eq id id0); auto. apply IS.add_3 in H; auto.
+ + intros. destruct (ident_eq id id0); auto. apply IS.add_3 in H; auto.
Qed.
Lemma addlist_workset_incl:
forall l w, workset_incl w (addlist_workset l w).
Proof.
- induction l; simpl; intros.
+ induction l; simpl; intros.
apply workset_incl_refl.
eapply workset_incl_trans. apply add_workset_incl. eauto.
Qed.
@@ -90,14 +90,14 @@ Proof.
unfold add_ref_function; intros. apply PTree_Properties.fold_rec.
- auto.
- apply workset_incl_refl.
-- intros. apply workset_incl_trans with a; auto.
+- intros. apply workset_incl_trans with a; auto.
unfold add_ref_instruction. apply addlist_workset_incl.
Qed.
Lemma add_ref_globvar_incl:
forall gv w, workset_incl w (add_ref_globvar gv w).
Proof.
- unfold add_ref_globvar; intros.
+ unfold add_ref_globvar; intros.
revert w. induction (gvar_init gv); simpl; intros.
apply workset_incl_refl.
eapply workset_incl_trans; [ | eauto ].
@@ -108,7 +108,7 @@ Qed.
Lemma add_ref_definition_incl:
forall pm id w, workset_incl w (add_ref_definition pm id w).
Proof.
- unfold add_ref_definition; intros.
+ unfold add_ref_definition; intros.
destruct (pm!id) as [[[] | ? ] | ].
apply add_ref_function_incl.
apply workset_incl_refl.
@@ -120,10 +120,10 @@ Lemma initial_workset_incl:
forall p, workset_incl {| w_seen := IS.empty; w_todo := nil |} (initial_workset p).
Proof.
unfold initial_workset; intros.
- eapply workset_incl_trans. 2: apply add_workset_incl.
+ eapply workset_incl_trans. 2: apply add_workset_incl.
generalize {| w_seen := IS.empty; w_todo := nil |}. induction (prog_public p); simpl; intros.
apply workset_incl_refl.
- eapply workset_incl_trans. eapply add_workset_incl. eapply IHl.
+ eapply workset_incl_trans. eapply add_workset_incl. eapply IHl.
Qed.
(** Soundness properties for functions that add identifiers to the workset *)
@@ -141,7 +141,7 @@ Lemma seen_addlist_workset:
forall id l (w: workset),
In id l -> IS.In id (addlist_workset l w).
Proof.
- induction l; simpl; intros.
+ induction l; simpl; intros.
tauto.
destruct H. subst a.
eapply seen_workset_incl. apply addlist_workset_incl. apply seen_add_workset.
@@ -159,9 +159,9 @@ Proof.
- destruct H1 as (pc & i & A & B). apply H0; auto. exists pc, i; split; auto. rewrite H; auto.
- destruct H as (pc & i & A & B). rewrite PTree.gempty in A; discriminate.
- destruct H2 as (pc & i & A & B). rewrite PTree.gsspec in A. destruct (peq pc k).
- + inv A. unfold add_ref_instruction. apply seen_addlist_workset; auto.
- + unfold add_ref_instruction. eapply seen_workset_incl. apply addlist_workset_incl.
- apply H1. exists pc, i; auto.
+ + inv A. unfold add_ref_instruction. apply seen_addlist_workset; auto.
+ + unfold add_ref_instruction. eapply seen_workset_incl. apply addlist_workset_incl.
+ apply H1. exists pc, i; auto.
Qed.
Definition ref_fundef (fd: fundef) (id: ident) : Prop :=
@@ -186,35 +186,35 @@ Proof.
IS.In id' w \/ In (Init_addrof id' ofs) l ->
IS.In id' (fold_left add_ref_init_data l w)).
{
- induction l; simpl; intros.
+ induction l; simpl; intros.
tauto.
apply IHl. intuition auto.
- left. destruct a; simpl; auto. eapply seen_workset_incl. apply add_workset_incl. auto.
+ left. destruct a; simpl; auto. eapply seen_workset_incl. apply add_workset_incl. auto.
subst; left; simpl. apply seen_add_workset.
}
- apply H0; auto.
+ apply H0; auto.
Qed.
Lemma seen_main_initial_workset:
forall p, IS.In p.(prog_main) (initial_workset p).
Proof.
- unfold initial_workset; intros. apply seen_add_workset.
+ unfold initial_workset; intros. apply seen_add_workset.
Qed.
Lemma seen_public_initial_workset:
forall p id, In id p.(prog_public) -> IS.In id (initial_workset p).
Proof.
- intros. unfold initial_workset. eapply seen_workset_incl. apply add_workset_incl.
+ intros. unfold initial_workset. eapply seen_workset_incl. apply add_workset_incl.
assert (forall l (w: workset),
IS.In id w \/ In id l -> IS.In id (fold_left (fun w id => add_workset id w) l w)).
{
induction l; simpl; intros.
tauto.
apply IHl. intuition auto; left.
- eapply seen_workset_incl. apply add_workset_incl. auto.
- subst a. apply seen_add_workset.
+ eapply seen_workset_incl. apply add_workset_incl. auto.
+ subst a. apply seen_add_workset.
}
- apply H0. auto.
+ apply H0. auto.
Qed.
(** * Semantic preservation *)
@@ -248,16 +248,16 @@ Lemma iter_step_invariant:
| inr w' => workset_invariant w'
end.
Proof.
- unfold iter_step, workset_invariant, used_set_closed; intros.
+ unfold iter_step, workset_invariant, used_set_closed; intros.
destruct (w_todo w) as [ | id rem ]; intros.
- eapply H; eauto.
- set (w' := {| w_seen := w.(w_seen); w_todo := rem |}) in *.
destruct (add_ref_definition_incl pm id w').
destruct (ident_eq id id0).
- + subst id0. eapply seen_add_ref_definition; eauto.
+ + subst id0. eapply seen_add_ref_definition; eauto.
+ exploit TRACK; eauto. intros [A|A].
- * apply SEEN. eapply H; eauto. simpl.
- assert (~ In id0 rem).
+ * apply SEEN. eapply H; eauto. simpl.
+ assert (~ In id0 rem).
{ change rem with (w_todo w'). red; intros. elim H1; auto. }
tauto.
* contradiction.
@@ -267,10 +267,10 @@ Theorem used_globals_sound:
forall u, used_globals p = Some u -> used_set_closed u.
Proof.
unfold used_globals; intros. eapply PrimIter.iterate_prop with (P := workset_invariant); eauto.
-- intros. apply iter_step_invariant; auto.
-- destruct (initial_workset_incl p).
- red; intros. edestruct TRACK; eauto.
- simpl in H4. eelim IS.empty_1; eauto.
+- intros. apply iter_step_invariant; auto.
+- destruct (initial_workset_incl p).
+ red; intros. edestruct TRACK; eauto.
+ simpl in H4. eelim IS.empty_1; eauto.
contradiction.
Qed.
@@ -286,7 +286,7 @@ Proof.
- red; auto.
Qed.
-Definition used: IS.t :=
+Definition used: IS.t :=
match used_globals p with Some u => u | None => IS.empty end.
Remark USED_GLOBALS: used_globals p = Some used.
@@ -322,7 +322,7 @@ Remark filter_globdefs_accu:
Proof.
induction defs; simpl; intros.
auto.
- destruct a as [id gd]. destruct (IS.mem id u); auto.
+ destruct a as [id gd]. destruct (IS.mem id u); auto.
rewrite <- IHdefs. auto.
Qed.
@@ -330,28 +330,28 @@ Remark filter_globdefs_nil:
forall u accu defs,
filter_globdefs u accu defs = filter_globdefs u nil defs ++ accu.
Proof.
- intros. rewrite <- filter_globdefs_accu. auto.
+ intros. rewrite <- filter_globdefs_accu. auto.
Qed.
Theorem transform_program_charact:
forall id, (program_map tp)!id = if IS.mem id used then pm!id else None.
Proof.
- intros.
+ intros.
assert (X: forall l u m1 m2,
IS.In id u ->
m1!id = m2!id ->
(fold_left add_def_prog_map (filter_globdefs u nil l) m1)!id =
(fold_left add_def_prog_map (List.rev l) m2)!id).
{
- induction l; simpl; intros.
+ induction l; simpl; intros.
auto.
- destruct a as [id1 gd1]. rewrite fold_left_app. simpl.
+ destruct a as [id1 gd1]. rewrite fold_left_app. simpl.
destruct (IS.mem id1 u) eqn:MEM.
- rewrite filter_globdefs_nil. rewrite fold_left_app. simpl.
- unfold add_def_prog_map at 1 3. simpl.
+ rewrite filter_globdefs_nil. rewrite fold_left_app. simpl.
+ unfold add_def_prog_map at 1 3. simpl.
rewrite ! PTree.gsspec. destruct (peq id id1). auto.
- apply IHl; auto. apply IS.remove_2; auto.
- unfold add_def_prog_map at 2. simpl. rewrite PTree.gso. apply IHl; auto.
+ apply IHl; auto. apply IS.remove_2; auto.
+ unfold add_def_prog_map at 2. simpl. rewrite PTree.gso. apply IHl; auto.
red; intros; subst id1.
assert (IS.mem id u = true) by (apply IS.mem_1; auto). congruence.
}
@@ -364,9 +364,9 @@ Proof.
auto.
destruct a as [id1 gd1].
destruct (IS.mem id1 u) eqn:MEM.
- rewrite filter_globdefs_nil. rewrite fold_left_app. simpl.
+ rewrite filter_globdefs_nil. rewrite fold_left_app. simpl.
unfold add_def_prog_map at 1. simpl. rewrite PTree.gso by congruence. eapply IHl; eauto.
- rewrite ISF.remove_b. rewrite H; auto.
+ rewrite ISF.remove_b. rewrite H; auto.
eapply IHl; eauto.
}
unfold pm, program_map.
@@ -393,8 +393,8 @@ Definition genv_progmap_match (ge: genv) (pm: prog_map) : Prop :=
Lemma genv_program_map:
forall p, genv_progmap_match (Genv.globalenv p) (program_map p).
Proof.
- intros. unfold Genv.globalenv, program_map.
- assert (REC: forall defs g m,
+ intros. unfold Genv.globalenv, program_map.
+ assert (REC: forall defs g m,
genv_progmap_match g m ->
genv_progmap_match (Genv.add_globals g defs) (fold_left add_def_prog_map defs m)).
{
@@ -407,12 +407,12 @@ Proof.
- rewrite PTree.gsspec. destruct (peq id id1); subst.
+ rewrite ! PTree.gss. auto.
+ destruct (Genv.genv_symb g)!id as [b|] eqn:S; rewrite PTree.gso by auto.
- * rewrite PTree.gso. auto. apply Plt_ne. eapply Genv.genv_symb_range; eauto.
+ * rewrite PTree.gso. auto. apply Plt_ne. eapply Genv.genv_symb_range; eauto.
* auto.
- rewrite PTree.gsspec. destruct (peq id id1); subst.
+ rewrite ! PTree.gss. auto.
+ destruct (Genv.genv_symb g)!id as [b|] eqn:S; rewrite PTree.gso by auto.
- * rewrite PTree.gso. auto. apply Plt_ne. eapply Genv.genv_symb_range; eauto.
+ * rewrite PTree.gso. auto. apply Plt_ne. eapply Genv.genv_symb_range; eauto.
* auto.
}
apply REC. red; intros. unfold Genv.find_symbol, Genv.empty_genv; simpl. rewrite ! PTree.gempty; auto.
@@ -421,7 +421,7 @@ Qed.
Lemma transform_program_kept:
forall id b, Genv.find_symbol tge id = Some b -> kept id.
Proof.
- intros. generalize (genv_program_map tp id). fold tge; rewrite H.
+ intros. generalize (genv_program_map tp id). fold tge; rewrite H.
rewrite transform_program_charact. intros. destruct (IS.mem id used) eqn:U.
unfold kept; apply IS.mem_2; auto.
contradiction.
@@ -454,7 +454,7 @@ Record meminj_preserves_globals (f: meminj) : Prop := {
Definition init_meminj : meminj :=
fun b =>
match Genv.invert_symbol ge b with
- | Some id =>
+ | Some id =>
match Genv.find_symbol tge id with
| Some b' => Some (b', 0)
| None => None
@@ -471,7 +471,7 @@ Proof.
destruct (Genv.invert_symbol ge b) as [id|] eqn:S; try discriminate.
destruct (Genv.find_symbol tge id) as [b''|] eqn:F; inv H.
split. auto. exists id. split. apply Genv.invert_find_symbol; auto. auto.
-Qed.
+Qed.
Lemma init_meminj_preserves_globals:
meminj_preserves_globals init_meminj.
@@ -487,35 +487,35 @@ Proof.
exists b'; auto. rewrite H2 in H1; contradiction.
- generalize (genv_program_map tp id). fold tge. rewrite H. intros.
destruct (program_map tp)!id as [gd|] eqn:PM; try contradiction.
- generalize (transform_program_charact id). rewrite PM.
+ generalize (transform_program_charact id). rewrite PM.
destruct (IS.mem id used) eqn:USED; intros; try discriminate.
generalize (genv_program_map p id). fold ge; fold pm.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; intros; try congruence.
- exists b; split; auto. unfold init_meminj.
+ exists b; split; auto. unfold init_meminj.
erewrite Genv.find_invert_symbol by eauto. rewrite H. auto.
- exploit init_meminj_invert; eauto. intros (A & id & B & C).
generalize (genv_program_map p id) (genv_program_map tp id). fold ge; fold tge; fold pm.
rewrite transform_program_charact. rewrite B, C. intros.
destruct (IS.mem id used) eqn:KEPT; try contradiction.
- destruct (pm!id) as [gd|] eqn:PM; try contradiction.
+ destruct (pm!id) as [gd|] eqn:PM; try contradiction.
destruct gd as [fd'|gv'].
- + assert (fd' = fd) by congruence. subst fd'. split. auto. split. auto.
- intros. eapply kept_closed; eauto. red; apply IS.mem_2; auto.
+ + assert (fd' = fd) by congruence. subst fd'. split. auto. split. auto.
+ intros. eapply kept_closed; eauto. red; apply IS.mem_2; auto.
+ assert (b <> b) by (eapply Genv.genv_funs_vars; eassumption). congruence.
- exploit init_meminj_invert; eauto. intros (A & id & B & C). split; auto.
generalize (genv_program_map p id) (genv_program_map tp id). fold ge; fold tge; fold pm.
rewrite transform_program_charact. rewrite B, C. intros.
destruct (IS.mem id used); try contradiction.
- destruct (pm!id) as [gd|]; try contradiction.
- destruct gd as [fd'|gv'].
+ destruct (pm!id) as [gd|]; try contradiction.
+ destruct gd as [fd'|gv'].
+ assert (b <> b) by (eapply Genv.genv_funs_vars; eassumption). congruence.
+ congruence.
- exploit init_meminj_invert; eauto. intros (A & id & B & C). split; auto.
generalize (genv_program_map p id) (genv_program_map tp id). fold ge; fold tge; fold pm.
rewrite transform_program_charact. rewrite B, C. intros.
destruct (IS.mem id used); try contradiction.
- destruct (pm!id) as [gd|]; try contradiction.
- destruct gd as [fd'|gv'].
+ destruct (pm!id) as [gd|]; try contradiction.
+ destruct gd as [fd'|gv'].
+ assert (b' <> b') by (eapply Genv.genv_funs_vars; eassumption). congruence.
+ congruence.
Qed.
@@ -523,28 +523,28 @@ Qed.
Lemma globals_symbols_inject:
forall j, meminj_preserves_globals j -> symbols_inject j ge tge.
Proof.
- intros.
+ intros.
assert (E1: Genv.genv_public ge = p.(prog_public)).
{ apply Genv.globalenv_public. }
assert (E2: Genv.genv_public tge = p.(prog_public)).
- { unfold tge; rewrite Genv.globalenv_public.
+ { unfold tge; rewrite Genv.globalenv_public.
unfold transform_program in TRANSF. rewrite USED_GLOBALS in TRANSF. inversion TRANSF. auto. }
split; [|split;[|split]]; intros.
+ simpl; unfold Genv.public_symbol; rewrite E1, E2.
destruct (Genv.find_symbol tge id) as [b'|] eqn:TFS.
- exploit symbols_inject_3; eauto. intros (b & FS & INJ). rewrite FS. auto.
+ exploit symbols_inject_3; eauto. intros (b & FS & INJ). rewrite FS. auto.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
destruct (in_dec ident_eq id (prog_public p)); simpl; auto.
- exploit symbols_inject_2; eauto. apply kept_public; auto.
+ exploit symbols_inject_2; eauto. apply kept_public; auto.
intros (b' & TFS' & INJ). congruence.
- + eapply symbols_inject_1; eauto.
- + simpl in *; unfold Genv.public_symbol in H0.
+ + eapply symbols_inject_1; eauto.
+ + simpl in *; unfold Genv.public_symbol in H0.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; try discriminate.
- rewrite E1 in H0.
+ rewrite E1 in H0.
destruct (in_dec ident_eq id (prog_public p)); try discriminate. inv H1.
- exploit symbols_inject_2; eauto. apply kept_public; auto.
+ exploit symbols_inject_2; eauto. apply kept_public; auto.
intros (b' & A & B); exists b'; auto.
- + simpl. unfold Genv.block_is_volatile.
+ + simpl. unfold Genv.block_is_volatile.
destruct (Genv.find_var_info ge b1) as [gv|] eqn:V1.
exploit var_info_inject; eauto. intros [A B]. rewrite A. auto.
destruct (Genv.find_var_info tge b2) as [gv|] eqn:V2; auto.
@@ -552,14 +552,14 @@ Proof.
Qed.
Lemma symbol_address_inject:
- forall j id ofs,
+ forall j id ofs,
meminj_preserves_globals j -> kept id ->
Val.inject j (Genv.symbol_address ge id ofs) (Genv.symbol_address tge id ofs).
Proof.
intros. unfold Genv.symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
- exploit symbols_inject_2; eauto. intros (b' & TFS & INJ). rewrite TFS.
+ exploit symbols_inject_2; eauto. intros (b' & TFS & INJ). rewrite TFS.
econstructor; eauto. rewrite Int.add_zero; auto.
-Qed.
+Qed.
(** Semantic preservation *)
@@ -569,7 +569,7 @@ Definition regset_inject (f: meminj) (rs rs': regset): Prop :=
Lemma regs_inject:
forall f rs rs', regset_inject f rs rs' -> forall l, Val.inject_list f rs##l rs'##l.
Proof.
- induction l; simpl. constructor. constructor; auto.
+ induction l; simpl. constructor. constructor; auto.
Qed.
Lemma set_reg_inject:
@@ -577,7 +577,7 @@ Lemma set_reg_inject:
regset_inject f rs rs' -> Val.inject f v v' ->
regset_inject f (rs#r <- v) (rs'#r <- v').
Proof.
- intros; red; intros. rewrite ! Regmap.gsspec. destruct (peq r0 r); auto.
+ intros; red; intros. rewrite ! Regmap.gsspec. destruct (peq r0 r); auto.
Qed.
Lemma set_res_inject:
@@ -585,13 +585,13 @@ Lemma set_res_inject:
regset_inject f rs rs' -> Val.inject f v v' ->
regset_inject f (regmap_setres res v rs) (regmap_setres res v' rs').
Proof.
- intros. destruct res; auto. apply set_reg_inject; auto.
+ intros. destruct res; auto. apply set_reg_inject; auto.
Qed.
Lemma regset_inject_incr:
forall f f' rs rs', regset_inject f rs rs' -> inject_incr f f' -> regset_inject f' rs rs'.
Proof.
- intros; red; intros. apply val_inject_incr with f; auto.
+ intros; red; intros. apply val_inject_incr with f; auto.
Qed.
Lemma regset_undef_inject:
@@ -606,7 +606,7 @@ Lemma init_regs_inject:
regset_inject f (init_regs args params) (init_regs args' params).
Proof.
induction 1; intros; destruct params; simpl; try (apply regset_undef_inject).
- apply set_reg_inject; auto.
+ apply set_reg_inject; auto.
Qed.
Inductive match_stacks (j: meminj):
@@ -631,7 +631,7 @@ Lemma match_stacks_preserves_globals:
match_stacks j s ts bound tbound ->
meminj_preserves_globals j.
Proof.
- induction 1; auto.
+ induction 1; auto.
Qed.
Lemma match_stacks_incr:
@@ -645,7 +645,7 @@ Proof.
- assert (SAME: forall b b' delta, Plt b (Genv.genv_next ge) ->
j' b = Some(b', delta) -> j b = Some(b', delta)).
{ intros. destruct (j b) as [[b1 delta1] | ] eqn: J.
- exploit H; eauto. congruence.
+ exploit H; eauto. congruence.
exploit H3; eauto. intros [A B]. elim (Plt_strict b).
eapply Plt_Ple_trans. eauto. eapply Ple_trans; eauto. }
assert (SAME': forall b b' delta, Plt b' (Genv.genv_next tge) ->
@@ -655,19 +655,19 @@ Proof.
exploit H3; eauto. intros [A B]. elim (Plt_strict b').
eapply Plt_Ple_trans. eauto. eapply Ple_trans; eauto. }
constructor; auto. constructor; intros.
- + exploit symbols_inject_1; eauto. apply SAME; auto.
- eapply Genv.genv_symb_range; eauto.
- + exploit symbols_inject_2; eauto. intros (b' & A & B).
+ + exploit symbols_inject_1; eauto. apply SAME; auto.
+ eapply Genv.genv_symb_range; eauto.
+ + exploit symbols_inject_2; eauto. intros (b' & A & B).
exists b'; auto.
+ exploit symbols_inject_3; eauto. intros (b & A & B).
exists b; auto.
- + eapply funct_ptr_inject; eauto. apply SAME; auto.
+ + eapply funct_ptr_inject; eauto. apply SAME; auto.
eapply Genv.genv_funs_range; eauto.
- + eapply var_info_inject; eauto. apply SAME; auto.
+ + eapply var_info_inject; eauto. apply SAME; auto.
eapply Genv.genv_vars_range; eauto.
- + eapply var_info_rev_inject; eauto. apply SAME'; auto.
+ + eapply var_info_rev_inject; eauto. apply SAME'; auto.
eapply Genv.genv_vars_range; eauto.
-- econstructor; eauto.
+- econstructor; eauto.
apply IHmatch_stacks.
intros. exploit H1; eauto. intros [A B]. split; eapply Ple_trans; eauto.
apply Plt_Ple; auto. apply Plt_Ple; auto.
@@ -681,8 +681,8 @@ Lemma match_stacks_bound:
match_stacks j s ts bound' tbound'.
Proof.
induction 1; intros.
-- constructor; auto. eapply Ple_trans; eauto. eapply Ple_trans; eauto.
-- econstructor; eauto. eapply Plt_Ple_trans; eauto. eapply Plt_Ple_trans; eauto.
+- constructor; auto. eapply Ple_trans; eauto. eapply Ple_trans; eauto.
+- econstructor; eauto. eapply Plt_Ple_trans; eauto. eapply Plt_Ple_trans; eauto.
Qed.
Inductive match_states: state -> state -> Prop :=
@@ -735,14 +735,14 @@ Lemma find_function_inject:
find_function tge ros trs = Some fd /\ (forall id, ref_fundef fd id -> kept id).
Proof.
intros. destruct ros as [r|id]; simpl in *.
-- exploit Genv.find_funct_inv; eauto. intros (b & R). rewrite R in H0.
- rewrite Genv.find_funct_find_funct_ptr in H0.
- specialize (H1 r). rewrite R in H1. inv H1.
- exploit funct_ptr_inject; eauto. intros (A & B & C).
+- exploit Genv.find_funct_inv; eauto. intros (b & R). rewrite R in H0.
+ rewrite Genv.find_funct_find_funct_ptr in H0.
+ specialize (H1 r). rewrite R in H1. inv H1.
+ exploit funct_ptr_inject; eauto. intros (A & B & C).
rewrite B; auto.
- destruct (Genv.find_symbol ge id) as [b|] eqn:FS; try discriminate.
- exploit symbols_inject_2; eauto. intros (tb & P & Q). rewrite P.
- exploit funct_ptr_inject; eauto. intros (A & B & C).
+ exploit symbols_inject_2; eauto. intros (tb & P & Q). rewrite P.
+ exploit funct_ptr_inject; eauto. intros (A & B & C).
auto.
Qed.
@@ -759,28 +759,28 @@ Lemma eval_builtin_arg_inject:
/\ Val.inject j v v'.
Proof.
induction 1; intros SP GL RS MI K; simpl in K.
-- exists rs'#x; split; auto. constructor.
+- exists rs'#x; split; auto. constructor.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
-- simpl in H. exploit Mem.load_inject; eauto. rewrite Zplus_0_r.
+- simpl in H. exploit Mem.load_inject; eauto. rewrite Zplus_0_r.
intros (v' & A & B). exists v'; auto with barg.
- econstructor; split; eauto with barg. simpl. econstructor; eauto. rewrite Int.add_zero; auto.
- assert (Val.inject j (Senv.symbol_address ge id ofs) (Senv.symbol_address tge id ofs)).
- { unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
+ { unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
- exploit symbols_inject_2; eauto. intros (b' & A & B). rewrite A.
+ exploit symbols_inject_2; eauto. intros (b' & A & B). rewrite A.
econstructor; eauto. rewrite Int.add_zero; auto. }
exploit Mem.loadv_inject; eauto. intros (v' & A & B). exists v'; auto with barg.
- econstructor; split; eauto with barg.
- unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
+ unfold Senv.symbol_address; simpl; unfold Genv.symbol_address.
destruct (Genv.find_symbol ge id) as [b|] eqn:FS; auto.
- exploit symbols_inject_2; eauto. intros (b' & A & B). rewrite A.
+ exploit symbols_inject_2; eauto. intros (b' & A & B). rewrite A.
econstructor; eauto. rewrite Int.add_zero; auto.
- destruct IHeval_builtin_arg1 as (v1' & A1 & B1); eauto using in_or_app.
destruct IHeval_builtin_arg2 as (v2' & A2 & B2); eauto using in_or_app.
- exists (Val.longofwords v1' v2'); split; auto with barg.
+ exists (Val.longofwords v1' v2'); split; auto with barg.
apply Val.longofwords_inject; auto.
Qed.
@@ -796,11 +796,11 @@ Lemma eval_builtin_args_inject:
eval_builtin_args tge (fun r => rs'#r) (Vptr sp' Int.zero) m' al vl'
/\ Val.inject_list j vl vl'.
Proof.
- induction 1; intros.
+ induction 1; intros.
- exists (@nil val); split; constructor.
-- simpl in H5.
+- simpl in H5.
exploit eval_builtin_arg_inject; eauto using in_or_app. intros (v1' & A & B).
- destruct IHlist_forall2 as (vl' & C & D); eauto using in_or_app.
+ destruct IHlist_forall2 as (vl' & C & D); eauto using in_or_app.
exists (v1' :: vl'); split; constructor; auto.
Qed.
@@ -812,36 +812,36 @@ Proof.
induction 1; intros; inv MS.
- (* nop *)
- econstructor; split.
- eapply exec_Inop; eauto.
- econstructor; eauto.
+ econstructor; split.
+ eapply exec_Inop; eauto.
+ econstructor; eauto.
- (* op *)
- assert (A: exists tv,
+ assert (A: exists tv,
eval_operation tge (Vptr tsp Int.zero) op trs##args tm = Some tv
/\ Val.inject j v tv).
- { apply eval_operation_inj with (ge1 := ge) (m1 := m) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
+ { apply eval_operation_inj with (ge1 := ge) (m1 := m) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
intros; eapply Mem.valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
intros; eapply Mem.different_pointers_inject; eauto.
intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
apply KEPT. red. exists pc, (Iop op args res pc'); auto.
- econstructor; eauto.
+ econstructor; eauto.
apply regs_inject; auto.
assumption. }
destruct A as (tv & B & C).
econstructor; split. eapply exec_Iop; eauto.
- econstructor; eauto. apply set_reg_inject; auto.
+ econstructor; eauto. apply set_reg_inject; auto.
- (* load *)
- assert (A: exists ta,
+ assert (A: exists ta,
eval_addressing tge (Vptr tsp Int.zero) addr trs##args = Some ta
/\ Val.inject j a ta).
- { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
+ { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
apply KEPT. red. exists pc, (Iload chunk addr args dst pc'); auto.
- econstructor; eauto.
+ econstructor; eauto.
apply regs_inject; auto.
assumption. }
destruct A as (ta & B & C).
@@ -850,13 +850,13 @@ Proof.
econstructor; eauto. apply set_reg_inject; auto.
- (* store *)
- assert (A: exists ta,
+ assert (A: exists ta,
eval_addressing tge (Vptr tsp Int.zero) addr trs##args = Some ta
/\ Val.inject j a ta).
- { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
+ { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Int.zero) (vl1 := rs##args).
intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
apply KEPT. red. exists pc, (Istore chunk addr args src pc'); auto.
- econstructor; eauto.
+ econstructor; eauto.
apply regs_inject; auto.
assumption. }
destruct A as (ta & B & C).
@@ -865,45 +865,45 @@ Proof.
econstructor; eauto.
- (* call *)
- exploit find_function_inject.
- eapply match_stacks_preserves_globals; eauto. eauto.
+ exploit find_function_inject.
+ eapply match_stacks_preserves_globals; eauto. eauto.
destruct ros as [r|id]. eauto. apply KEPT. red. econstructor; econstructor; split; eauto. simpl; auto.
intros (A & B).
- econstructor; split. eapply exec_Icall; eauto.
- econstructor; eauto.
+ econstructor; split. eapply exec_Icall; eauto.
+ econstructor; eauto.
econstructor; eauto.
- change (Mem.valid_block m sp0). eapply Mem.valid_block_inject_1; eauto.
- change (Mem.valid_block tm tsp). eapply Mem.valid_block_inject_2; eauto.
+ change (Mem.valid_block m sp0). eapply Mem.valid_block_inject_1; eauto.
+ change (Mem.valid_block tm tsp). eapply Mem.valid_block_inject_2; eauto.
apply regs_inject; auto.
- (* tailcall *)
- exploit find_function_inject.
- eapply match_stacks_preserves_globals; eauto. eauto.
+ exploit find_function_inject.
+ eapply match_stacks_preserves_globals; eauto. eauto.
destruct ros as [r|id]. eauto. apply KEPT. red. econstructor; econstructor; split; eauto. simpl; auto.
intros (A & B).
exploit Mem.free_parallel_inject; eauto. rewrite ! Zplus_0_r. intros (tm' & C & D).
econstructor; split.
- eapply exec_Itailcall; eauto.
- econstructor; eauto.
+ eapply exec_Itailcall; eauto.
+ econstructor; eauto.
apply match_stacks_bound with stk tsp; auto.
- apply Plt_Ple.
+ apply Plt_Ple.
change (Mem.valid_block m' stk). eapply Mem.valid_block_inject_1; eauto.
- apply Plt_Ple.
+ apply Plt_Ple.
change (Mem.valid_block tm' tsp). eapply Mem.valid_block_inject_2; eauto.
apply regs_inject; auto.
- (* builtin *)
- exploit eval_builtin_args_inject; eauto.
+ exploit eval_builtin_args_inject; eauto.
eapply match_stacks_preserves_globals; eauto.
- intros. apply KEPT. red. econstructor; econstructor; eauto.
+ intros. apply KEPT. red. econstructor; econstructor; eauto.
intros (vargs' & P & Q).
- exploit external_call_inject; eauto.
+ exploit external_call_inject; eauto.
eapply match_stacks_preserves_globals; eauto.
intros (j' & tv & tm' & A & B & C & D & E & F & G).
econstructor; split.
eapply exec_Ibuiltin; eauto.
- eapply match_states_regular with (j := j'); eauto.
- apply match_stacks_incr with j; auto.
+ eapply match_states_regular with (j := j'); eauto.
+ apply match_stacks_incr with j; auto.
intros. exploit G; eauto. intros [U V].
assert (Mem.valid_block m sp0) by (eapply Mem.valid_block_inject_1; eauto).
assert (Mem.valid_block tm tsp) by (eapply Mem.valid_block_inject_2; eauto).
@@ -914,52 +914,52 @@ Proof.
assert (C: eval_condition cond trs##args tm = Some b).
{ eapply eval_condition_inject; eauto. apply regs_inject; auto. }
econstructor; split.
- eapply exec_Icond with (pc' := if b then ifso else ifnot); eauto.
+ eapply exec_Icond with (pc' := if b then ifso else ifnot); eauto.
econstructor; eauto.
- (* jumptbl *)
generalize (REGINJ arg); rewrite H0; intros INJ; inv INJ.
econstructor; split.
- eapply exec_Ijumptable; eauto.
+ eapply exec_Ijumptable; eauto.
econstructor; eauto.
- (* return *)
exploit Mem.free_parallel_inject; eauto. rewrite ! Zplus_0_r. intros (tm' & C & D).
econstructor; split.
- eapply exec_Ireturn; eauto.
- econstructor; eauto.
+ eapply exec_Ireturn; eauto.
+ econstructor; eauto.
apply match_stacks_bound with stk tsp; auto.
- apply Plt_Ple.
+ apply Plt_Ple.
change (Mem.valid_block m' stk). eapply Mem.valid_block_inject_1; eauto.
- apply Plt_Ple.
+ apply Plt_Ple.
change (Mem.valid_block tm' tsp). eapply Mem.valid_block_inject_2; eauto.
- destruct or; simpl; auto.
+ destruct or; simpl; auto.
- (* internal function *)
- exploit Mem.alloc_parallel_inject. eauto. eauto. apply Zle_refl. apply Zle_refl.
+ exploit Mem.alloc_parallel_inject. eauto. eauto. apply Zle_refl. apply Zle_refl.
intros (j' & tm' & tstk & C & D & E & F & G).
- assert (STK: stk = Mem.nextblock m) by (eapply Mem.alloc_result; eauto).
- assert (TSTK: tstk = Mem.nextblock tm) by (eapply Mem.alloc_result; eauto).
+ assert (STK: stk = Mem.nextblock m) by (eapply Mem.alloc_result; eauto).
+ assert (TSTK: tstk = Mem.nextblock tm) by (eapply Mem.alloc_result; eauto).
assert (STACKS': match_stacks j' s ts stk tstk).
- { rewrite STK, TSTK.
- apply match_stacks_incr with j; auto.
- intros. destruct (eq_block b1 stk).
- subst b1. rewrite F in H1; inv H1. subst b2. split; apply Ple_refl.
+ { rewrite STK, TSTK.
+ apply match_stacks_incr with j; auto.
+ intros. destruct (eq_block b1 stk).
+ subst b1. rewrite F in H1; inv H1. subst b2. split; apply Ple_refl.
rewrite G in H1 by auto. congruence. }
- econstructor; split.
+ econstructor; split.
eapply exec_function_internal; eauto.
eapply match_states_regular with (j := j'); eauto.
- apply init_regs_inject; auto. apply val_inject_list_incr with j; auto.
+ apply init_regs_inject; auto. apply val_inject_list_incr with j; auto.
- (* external function *)
- exploit external_call_inject; eauto.
+ exploit external_call_inject; eauto.
eapply match_stacks_preserves_globals; eauto.
intros (j' & tres & tm' & A & B & C & D & E & F & G).
econstructor; split.
eapply exec_function_external; eauto.
eapply match_states_return with (j := j'); eauto.
apply match_stacks_bound with (Mem.nextblock m) (Mem.nextblock tm).
- apply match_stacks_incr with j; auto.
+ apply match_stacks_incr with j; auto.
intros. exploit G; eauto. intros [P Q].
unfold Mem.valid_block in *; xomega.
eapply external_call_nextblock; eauto.
@@ -967,8 +967,8 @@ Proof.
- (* return *)
inv STACKS. econstructor; split.
- eapply exec_return.
- econstructor; eauto. apply set_reg_inject; auto.
+ eapply exec_return.
+ econstructor; eauto. apply set_reg_inject; auto.
Qed.
(** Relating initial memory states *)
@@ -976,11 +976,11 @@ Qed.
Remark init_meminj_no_overlap:
forall m, Mem.meminj_no_overlap init_meminj m.
Proof.
- intros; red; intros.
+ intros; red; intros.
exploit init_meminj_invert. eexact H0. intros (A1 & id1 & B1 & C1).
exploit init_meminj_invert. eexact H1. intros (A2 & id2 & B2 & C2).
- left; red; intros; subst b2'.
- assert (id1 = id2) by (eapply Genv.genv_vars_inj; eauto).
+ left; red; intros; subst b2'.
+ assert (id1 = id2) by (eapply Genv.genv_vars_inj; eauto).
congruence.
Qed.
@@ -994,7 +994,7 @@ Lemma store_zeros_unmapped_inj:
Proof.
intros until m1'. functional induction (store_zeros m1 b1 i n); intros.
inv H. auto.
- eapply IHo; eauto. eapply Mem.store_unmapped_inj; eauto.
+ eapply IHo; eauto. eapply Mem.store_unmapped_inj; eauto.
discriminate.
Qed.
@@ -1007,10 +1007,10 @@ Lemma store_zeros_mapped_inj:
exists m2', store_zeros m2 b2 i n = Some m2' /\ Mem.mem_inj init_meminj m1' m2'.
Proof.
intros until m1'. functional induction (store_zeros m1 b1 i n); intros.
- inv H. exists m2; split; auto. rewrite store_zeros_equation, e; auto.
- exploit Mem.store_mapped_inj; eauto. apply init_meminj_no_overlap. instantiate (1 := Vzero); constructor.
+ inv H. exists m2; split; auto. rewrite store_zeros_equation, e; auto.
+ exploit Mem.store_mapped_inj; eauto. apply init_meminj_no_overlap. instantiate (1 := Vzero); constructor.
intros (m2' & A & B). rewrite Zplus_0_r in A.
- exploit IHo; eauto. intros (m3' & C & D).
+ exploit IHo; eauto. intros (m3' & C & D).
exists m3'; split; auto. rewrite store_zeros_equation, e, A, C; auto.
discriminate.
Qed.
@@ -1022,7 +1022,7 @@ Lemma store_init_data_unmapped_inj:
init_meminj b1 = None ->
Mem.mem_inj init_meminj m1' m2.
Proof.
- intros. destruct id; simpl in H; try (eapply Mem.store_unmapped_inj; now eauto).
+ intros. destruct id; simpl in H; try (eapply Mem.store_unmapped_inj; now eauto).
inv H; auto.
destruct (Genv.find_symbol ge i0); try discriminate. eapply Mem.store_unmapped_inj; now eauto.
Qed.
@@ -1039,8 +1039,8 @@ Proof.
destruct init; simpl in *; try (eapply Mem.store_mapped_inj; now eauto).
inv H. exists m2; auto.
destruct (Genv.find_symbol ge i0) as [bi|] eqn:FS1; try discriminate.
- exploit symbols_inject_2. eapply init_meminj_preserves_globals. eapply H2; eauto. eauto.
- intros (bi' & A & B). rewrite A. eapply Mem.store_mapped_inj; eauto.
+ exploit symbols_inject_2. eapply init_meminj_preserves_globals. eapply H2; eauto. eauto.
+ intros (bi' & A & B). rewrite A. eapply Mem.store_mapped_inj; eauto.
econstructor; eauto. rewrite Int.add_zero; auto.
Qed.
@@ -1051,7 +1051,7 @@ Qed.
init_meminj b1 = None ->
Mem.mem_inj init_meminj m1' m2.
Proof.
- induction initlist; simpl; intros.
+ induction initlist; simpl; intros.
- inv H; auto.
- destruct (Genv.store_init_data ge m1 b1 i a) as [m1''|] eqn:ST; try discriminate.
eapply IHinitlist; eauto. eapply store_init_data_unmapped_inj; eauto.
@@ -1070,7 +1070,7 @@ Proof.
- destruct (Genv.store_init_data ge m1 b1 i a) as [m1''|] eqn:ST; try discriminate.
exploit store_init_data_mapped_inj; eauto. intros (m2'' & A & B).
exploit IHinitlist; eauto. intros (m2' & C & D).
- exists m2'; split; auto. rewrite A; auto.
+ exists m2'; split; auto. rewrite A; auto.
Qed.
Lemma alloc_global_unmapped_inj:
@@ -1082,16 +1082,16 @@ Lemma alloc_global_unmapped_inj:
Proof.
unfold Genv.alloc_global; intros. destruct g as [fd|gv].
- destruct (Mem.alloc m1 0 1) as (m1a, b) eqn:ALLOC.
- exploit Mem.alloc_result; eauto. intros EQ. rewrite <- EQ in H1.
- eapply Mem.drop_unmapped_inj with (m1 := m1a); eauto.
+ exploit Mem.alloc_result; eauto. intros EQ. rewrite <- EQ in H1.
+ eapply Mem.drop_unmapped_inj with (m1 := m1a); eauto.
eapply Mem.alloc_left_unmapped_inj; eauto.
- set (sz := Genv.init_data_list_size (gvar_init gv)) in *.
destruct (Mem.alloc m1 0 sz) as (m1a, b) eqn:ALLOC.
destruct (store_zeros m1a b 0 sz) as [m1b|] eqn: STZ; try discriminate.
destruct (Genv.store_init_data_list ge m1b b 0 (gvar_init gv)) as [m1c|] eqn:ST; try discriminate.
- exploit Mem.alloc_result; eauto. intros EQ. rewrite <- EQ in H1.
+ exploit Mem.alloc_result; eauto. intros EQ. rewrite <- EQ in H1.
eapply Mem.drop_unmapped_inj with (m1 := m1c); eauto.
- eapply store_init_data_list_unmapped_inj with (m1 := m1b); eauto.
+ eapply store_init_data_list_unmapped_inj with (m1 := m1b); eauto.
eapply store_zeros_unmapped_inj with (m1 := m1a); eauto.
eapply Mem.alloc_left_unmapped_inj; eauto.
Qed.
@@ -1114,10 +1114,10 @@ Proof.
{ eapply Mem.alloc_left_mapped_inj with (b1 := b1) (b2 := b2) (delta := 0).
eapply Mem.alloc_right_inj; eauto.
eauto.
- eauto with mem.
+ eauto with mem.
red; intros; apply Z.divide_0_r.
- intros. apply Mem.perm_implies with Freeable; auto with mem.
- eapply Mem.perm_alloc_2; eauto. omega.
+ intros. apply Mem.perm_implies with Freeable; auto with mem.
+ eapply Mem.perm_alloc_2; eauto. omega.
auto.
}
exploit Mem.drop_mapped_inj; eauto. apply init_meminj_no_overlap.
@@ -1132,14 +1132,14 @@ Proof.
{ eapply Mem.alloc_left_mapped_inj with (b1 := b1) (b2 := b2) (delta := 0).
eapply Mem.alloc_right_inj; eauto.
eauto.
- eauto with mem.
+ eauto with mem.
red; intros; apply Z.divide_0_r.
- intros. apply Mem.perm_implies with Freeable; auto with mem.
- eapply Mem.perm_alloc_2; eauto. omega.
+ intros. apply Mem.perm_implies with Freeable; auto with mem.
+ eapply Mem.perm_alloc_2; eauto. omega.
auto.
}
exploit store_zeros_mapped_inj; eauto. intros (m2b & A & B).
- exploit store_init_data_list_mapped_inj; eauto.
+ exploit store_init_data_list_mapped_inj; eauto.
intros. apply H2. red. exists ofs; auto. intros (m2c & C & D).
exploit Mem.drop_mapped_inj; eauto. apply init_meminj_no_overlap. intros (m2' & E & F).
exists m2'; split; auto. rewrite ! Zplus_0_r in E. rewrite A, C, E. auto.
@@ -1153,8 +1153,8 @@ Lemma alloc_globals_app:
| Some m1 => Genv.alloc_globals g m1 defs2
end.
Proof.
- induction defs1; simpl; intros. auto.
- destruct (Genv.alloc_global g m a); auto.
+ induction defs1; simpl; intros. auto.
+ destruct (Genv.alloc_global g m a); auto.
Qed.
Lemma alloc_globals_snoc:
@@ -1166,7 +1166,7 @@ Lemma alloc_globals_snoc:
end.
Proof.
intros. rewrite alloc_globals_app.
- destruct (Genv.alloc_globals g m defs1); auto. unfold Genv.alloc_globals.
+ destruct (Genv.alloc_globals g m defs1); auto. unfold Genv.alloc_globals.
destruct (Genv.alloc_global g m0 id_def); auto.
Qed.
@@ -1187,20 +1187,20 @@ Lemma alloc_globals_inj:
Proof.
induction defs; simpl; intros until g2; intros ALLOC GE1 GE2 NEXT1 SYMB1 SYMB2 SYMB3 PROGMAP.
- inv ALLOC. exists Mem.empty. intuition auto. constructor; intros.
- eelim Mem.perm_empty; eauto.
- exploit init_meminj_invert; eauto. intros [A B]. subst delta. apply Z.divide_0_r.
eelim Mem.perm_empty; eauto.
-- rewrite Genv.add_globals_app in GE1. simpl in GE1.
+ exploit init_meminj_invert; eauto. intros [A B]. subst delta. apply Z.divide_0_r.
+ eelim Mem.perm_empty; eauto.
+- rewrite Genv.add_globals_app in GE1. simpl in GE1.
set (g1' := Genv.add_globals (Genv.empty_genv fundef unit pubs) (rev defs)) in *.
rewrite alloc_globals_snoc in ALLOC.
destruct (Genv.alloc_globals ge Mem.empty (rev defs)) as [m1'|] eqn:ALLOC1'; try discriminate.
exploit Genv.alloc_global_nextblock; eauto. intros NEXTBLOCK1.
- assert (NEXTGE1: Genv.genv_next g1 = Pos.succ (Genv.genv_next g1')) by (rewrite GE1; reflexivity).
+ assert (NEXTGE1: Genv.genv_next g1 = Pos.succ (Genv.genv_next g1')) by (rewrite GE1; reflexivity).
assert (NEXT1': Mem.nextblock m1' = Genv.genv_next g1') by (unfold block in *; xomega).
rewrite fold_left_app in PROGMAP. simpl in PROGMAP.
destruct a as [id gd]. unfold add_def_prog_map at 1 in PROGMAP. simpl in PROGMAP.
destruct (IS.mem id u) eqn:MEM.
- + rewrite filter_globdefs_nil in *. rewrite alloc_globals_snoc.
+ + rewrite filter_globdefs_nil in *. rewrite alloc_globals_snoc.
rewrite Genv.add_globals_app in GE2. simpl in GE2.
set (g2' := Genv.add_globals (Genv.empty_genv fundef unit pubs) (filter_globdefs (IS.remove id u) nil defs)) in *.
assert (NEXTGE2: Genv.genv_next g2 = Pos.succ (Genv.genv_next g2')) by (rewrite GE2; reflexivity).
@@ -1209,15 +1209,15 @@ Proof.
rewrite GE1. unfold Genv.find_symbol; simpl. rewrite PTree.gss. congruence. }
exploit (IHdefs m1' (IS.remove id u) g1' g2'); eauto.
intros. rewrite ISF.remove_iff in H; destruct H.
- rewrite <- SYMB1 by auto. rewrite GE1. unfold Genv.find_symbol; simpl.
+ rewrite <- SYMB1 by auto. rewrite GE1. unfold Genv.find_symbol; simpl.
rewrite PTree.gso; auto.
intros. rewrite ISF.remove_iff in H; destruct H.
- rewrite <- SYMB2 by auto. rewrite GE2. unfold Genv.find_symbol; simpl.
+ rewrite <- SYMB2 by auto. rewrite GE2. unfold Genv.find_symbol; simpl.
rewrite PTree.gso; auto.
intros. rewrite ISF.remove_iff. destruct (ident_eq id id0).
subst id0. rewrite FS1 in H. inv H. eelim Plt_strict; eauto.
- exploit SYMB3. eexact H. unfold block in *; xomega. auto. tauto.
- intros. rewrite ISF.remove_iff in H; destruct H.
+ exploit SYMB3. eexact H. unfold block in *; xomega. auto. tauto.
+ intros. rewrite ISF.remove_iff in H; destruct H.
rewrite <- PROGMAP by auto. rewrite PTree.gso by auto. auto.
intros (m2' & A & B & C). fold g2' in B.
assert (FS2: Genv.find_symbol tge id = Some (Mem.nextblock m2')).
@@ -1226,16 +1226,16 @@ Proof.
assert (INJ: init_meminj (Mem.nextblock m1') = Some (Mem.nextblock m2', 0)).
{ apply Genv.find_invert_symbol in FS1. unfold init_meminj. rewrite FS1, FS2. auto. }
exploit alloc_global_mapped_inj. eexact ALLOC. eexact C. exact INJ.
- intros. apply kept_closed with id gd. eapply transform_program_kept; eauto.
- rewrite <- PROGMAP by (apply IS.mem_2; auto). apply PTree.gss. auto.
+ intros. apply kept_closed with id gd. eapply transform_program_kept; eauto.
+ rewrite <- PROGMAP by (apply IS.mem_2; auto). apply PTree.gss. auto.
intros (m2 & D & E).
- exploit Genv.alloc_global_nextblock; eauto. intros NEXTBLOCK2.
- exists m2; split. rewrite A, D. auto.
+ exploit Genv.alloc_global_nextblock; eauto. intros NEXTBLOCK2.
+ exists m2; split. rewrite A, D. auto.
split. unfold block in *; xomega.
auto.
- + exploit (IHdefs m1' u g1' g2); auto.
- intros. rewrite <- SYMB1 by auto. rewrite GE1.
- unfold Genv.find_symbol; simpl. rewrite PTree.gso; auto.
+ + exploit (IHdefs m1' u g1' g2); auto.
+ intros. rewrite <- SYMB1 by auto. rewrite GE1.
+ unfold Genv.find_symbol; simpl. rewrite PTree.gso; auto.
red; intros; subst id0. apply IS.mem_1 in H. congruence.
intros. eapply SYMB3; eauto. unfold block in *; xomega.
intros. rewrite <- PROGMAP by auto. rewrite PTree.gso; auto.
@@ -1248,8 +1248,8 @@ Proof.
eapply transform_program_kept; eauto.
intros P.
revert V. rewrite <- SYMB1, GE1 by auto. unfold Genv.find_symbol; simpl.
- rewrite PTree.gsspec. rewrite NEXT1'. destruct (peq id1 id); intros Q.
- subst id1. apply IS.mem_1 in P. congruence.
+ rewrite PTree.gsspec. rewrite NEXT1'. destruct (peq id1 id); intros Q.
+ subst id1. apply IS.mem_1 in P. congruence.
eelim Plt_strict. eapply Genv.genv_symb_range; eauto. }
exists m2; intuition auto. eapply alloc_global_unmapped_inj; eauto.
Qed.
@@ -1262,15 +1262,15 @@ Proof.
intros.
unfold transform_program in TRANSF; rewrite USED_GLOBALS in TRANSF; injection TRANSF. intros EQ.
destruct (alloc_globals_inj (prog_public p) (List.rev (prog_defs p)) m used ge tge) as (tm & A & B & C).
- rewrite rev_involutive; auto.
+ rewrite rev_involutive; auto.
rewrite rev_involutive; auto.
unfold tge; rewrite <- EQ; auto.
- symmetry. apply Genv.init_mem_genv_next; auto.
+ symmetry. apply Genv.init_mem_genv_next; auto.
auto. auto. auto.
intros. rewrite rev_involutive. auto.
assert (D: Genv.init_mem tp = Some tm).
{ unfold Genv.init_mem. fold tge. rewrite <- EQ. exact A. }
- pose proof (init_meminj_preserves_globals).
+ pose proof (init_meminj_preserves_globals).
exists init_meminj, tm; intuition auto.
constructor; intros.
+ auto.
@@ -1280,7 +1280,7 @@ Proof.
+ exploit init_meminj_invert; eauto. intros (P & id & Q & R).
eapply Genv.find_symbol_not_fresh; eauto.
+ apply init_meminj_no_overlap.
- + exploit init_meminj_invert; eauto. intros (P & id & Q & R).
+ + exploit init_meminj_invert; eauto. intros (P & id & Q & R).
split. omega. generalize (Int.unsigned_range_2 ofs). omega.
Qed.
@@ -1289,28 +1289,28 @@ Lemma transf_initial_states:
Proof.
intros. inv H. exploit init_mem_inject; eauto. intros (j & tm & A & B & C).
exploit symbols_inject_2. eauto. apply kept_main. eexact H1. intros (tb & P & Q).
- exploit funct_ptr_inject. eauto. eexact Q. exact H2.
+ exploit funct_ptr_inject. eauto. eexact Q. exact H2.
intros (R & S & T).
exists (Callstate nil f nil tm); split.
econstructor; eauto.
- fold tge. unfold transform_program in TRANSF; rewrite USED_GLOBALS in TRANSF; inversion TRANSF; auto.
+ fold tge. unfold transform_program in TRANSF; rewrite USED_GLOBALS in TRANSF; inversion TRANSF; auto.
econstructor; eauto.
- constructor. auto.
- erewrite <- Genv.init_mem_genv_next by eauto. apply Ple_refl.
- erewrite <- Genv.init_mem_genv_next by eauto. apply Ple_refl.
+ constructor. auto.
+ erewrite <- Genv.init_mem_genv_next by eauto. apply Ple_refl.
+ erewrite <- Genv.init_mem_genv_next by eauto. apply Ple_refl.
Qed.
Lemma transf_final_states:
forall S1 S2 r,
match_states S1 S2 -> final_state S1 r -> final_state S2 r.
Proof.
- intros. inv H0. inv H. inv STACKS. inv RESINJ. constructor.
+ intros. inv H0. inv H. inv STACKS. inv RESINJ. constructor.
Qed.
Theorem transf_program_correct:
forward_simulation (semantics p) (semantics tp).
Proof.
- intros.
+ intros.
eapply forward_simulation_step.
exploit globals_symbols_inject. apply init_meminj_preserves_globals. intros [A B]. exact A.
eexact transf_initial_states.
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index 22121075..979f8c0e 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -231,7 +231,7 @@ Definition definitive_initializer (init: list init_data) : bool :=
Definition alloc_global (rm: romem) (idg: ident * globdef fundef unit): romem :=
match idg with
- | (id, Gfun f) =>
+ | (id, Gfun f) =>
PTree.remove id rm
| (id, Gvar v) =>
if v.(gvar_readonly) && negb v.(gvar_volatile) && definitive_initializer v.(gvar_init)
@@ -255,26 +255,26 @@ Lemma analyze_entrypoint:
/\ ematch bc (init_regs vl (fn_params f)) ae
/\ mmatch bc m am.
Proof.
- intros.
- unfold analyze.
+ intros.
+ unfold analyze.
set (lu := Liveness.last_uses f).
set (entry := VA.State (einit_regs f.(fn_params)) mfunction_entry).
destruct (DS.fixpoint (fn_code f) successors_instr (transfer' f lu rm)
(fn_entrypoint f) entry) as [res|] eqn:FIX.
- assert (A: VA.ge res!!(fn_entrypoint f) entry) by (eapply DS.fixpoint_entry; eauto).
destruct (res!!(fn_entrypoint f)) as [ | ae am ]; simpl in A. contradiction.
- destruct A as [A1 A2].
- exists ae, am.
- split. auto.
- split. eapply ematch_ge; eauto. apply ematch_init; auto.
+ destruct A as [A1 A2].
+ exists ae, am.
+ split. auto.
+ split. eapply ematch_ge; eauto. apply ematch_init; auto.
auto.
- exists AE.top, mtop.
- split. apply PMap.gi.
- split. apply ematch_ge with (einit_regs (fn_params f)).
- apply ematch_init; auto. apply AE.ge_top.
- eapply mmatch_top'; eauto.
+ split. apply PMap.gi.
+ split. apply ematch_ge with (einit_regs (fn_params f)).
+ apply ematch_init; auto. apply AE.ge_top.
+ eapply mmatch_top'; eauto.
Qed.
-
+
Lemma analyze_successor:
forall f n ae am instr s rm ae' am',
(analyze rm f)!!n = VA.State ae am ->
@@ -291,9 +291,9 @@ Proof.
- assert (A: VA.ge res!!s (transfer' f lu rm n res#n)).
{ eapply DS.fixpoint_solution; eauto with coqlib.
intros. unfold transfer'. simpl. auto. }
- rewrite H in A. unfold transfer' in A. rewrite H2 in A. rewrite H2.
+ rewrite H in A. unfold transfer' in A. rewrite H2 in A. rewrite H2.
destruct lu!n.
- eapply VA.ge_trans. eauto. split; auto. apply eforget_ge.
+ eapply VA.ge_trans. eauto. split; auto. apply eforget_ge.
auto.
- rewrite H2. rewrite PMap.gi. split; intros. apply AE.ge_top. eapply mmatch_top'; eauto.
Qed.
@@ -311,11 +311,11 @@ Lemma analyze_succ:
/\ ematch bc e ae''
/\ mmatch bc m am''.
Proof.
- intros. exploit analyze_successor; eauto. rewrite H2.
+ intros. exploit analyze_successor; eauto. rewrite H2.
destruct (analyze rm f)#s as [ | ae'' am'']; simpl; try tauto. intros [A B].
exists ae'', am''.
- split. auto.
- split. eapply ematch_ge; eauto. eauto.
+ split. auto.
+ split. eapply ematch_ge; eauto. eauto.
Qed.
(** ** Analysis of registers and builtin arguments *)
@@ -323,7 +323,7 @@ Qed.
Lemma areg_sound:
forall bc e ae r, ematch bc e ae -> vmatch bc (e#r) (areg ae r).
Proof.
- intros. apply H.
+ intros. apply H.
Qed.
Lemma aregs_sound:
@@ -347,8 +347,8 @@ Lemma abuiltin_arg_sound:
Proof.
intros until am; intros EM RM MM GM SP.
induction 1; simpl; eauto with va.
-- eapply loadv_sound; eauto. simpl. rewrite Int.add_zero_l. auto with va.
-- simpl. rewrite Int.add_zero_l. auto with va.
+- eapply loadv_sound; eauto. simpl. rewrite Int.add_zero_l. auto with va.
+- simpl. rewrite Int.add_zero_l. auto with va.
- eapply loadv_sound; eauto. apply symbol_address_sound; auto.
- apply symbol_address_sound; auto.
Qed.
@@ -367,7 +367,7 @@ Proof.
intros until am; intros EM RM MM GM SP.
induction 1; simpl.
- constructor.
-- constructor; auto. eapply abuiltin_arg_sound; eauto.
+- constructor; auto. eapply abuiltin_arg_sound; eauto.
Qed.
Lemma set_builtin_res_sound:
@@ -376,7 +376,7 @@ Lemma set_builtin_res_sound:
vmatch bc v av ->
ematch bc (regmap_setres res v rs) (set_builtin_res res av ae).
Proof.
- intros. destruct res; simpl; auto. apply ematch_update; auto.
+ intros. destruct res; simpl; auto. apply ematch_update; auto.
Qed.
(** ** Constructing block classifications *)
@@ -396,24 +396,24 @@ Qed.
Lemma vmatch_no_stack: forall v x, vmatch bc v x -> vmatch bc v (Ifptr Nonstack).
Proof.
- induction 1; constructor; auto; eapply pmatch_no_stack; eauto.
+ induction 1; constructor; auto; eapply pmatch_no_stack; eauto.
Qed.
Lemma smatch_no_stack: forall m b p, smatch bc m b p -> smatch bc m b Nonstack.
Proof.
- intros. destruct H as [A B]. split; intros.
- eapply vmatch_no_stack; eauto.
- eapply pmatch_no_stack; eauto.
+ intros. destruct H as [A B]. split; intros.
+ eapply vmatch_no_stack; eauto.
+ eapply pmatch_no_stack; eauto.
Qed.
Lemma mmatch_no_stack: forall m am astk,
mmatch bc m am -> mmatch bc m {| am_stack := astk; am_glob := PTree.empty _; am_nonstack := Nonstack; am_top := Nonstack |}.
Proof.
intros. destruct H. constructor; simpl; intros.
-- elim (NOSTACK b); auto.
+- elim (NOSTACK b); auto.
- rewrite PTree.gempty in H0; discriminate.
-- eapply smatch_no_stack; eauto.
-- eapply smatch_no_stack; eauto.
+- eapply smatch_no_stack; eauto.
+- eapply smatch_no_stack; eauto.
- auto.
Qed.
@@ -439,8 +439,8 @@ Theorem allocate_stack:
/\ (forall b, Plt b sp -> bc' b = bc b)
/\ (forall v x, vmatch bc v x -> vmatch bc' v (Ifptr Nonstack)).
Proof.
- intros until am; intros ALLOC GENV RO MM NOSTACK.
- exploit Mem.nextblock_alloc; eauto. intros NB.
+ intros until am; intros ALLOC GENV RO MM NOSTACK.
+ exploit Mem.nextblock_alloc; eauto. intros NB.
exploit Mem.alloc_result; eauto. intros SP.
assert (SPINVALID: bc sp = BCinvalid).
{ rewrite SP. eapply bc_below_invalid. apply Plt_strict. eapply mmatch_below; eauto. }
@@ -450,13 +450,13 @@ Proof.
{
assert (forall b, f b = BCstack -> b = sp).
{ unfold f; intros. destruct (eq_block b sp); auto. eelim NOSTACK; eauto. }
- intros. transitivity sp; auto. symmetry; auto.
+ intros. transitivity sp; auto. symmetry; auto.
}
assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2).
{
assert (forall b id, f b = BCglob id -> bc b = BCglob id).
{ unfold f; intros. destruct (eq_block b sp). congruence. auto. }
- intros. eapply (bc_glob bc); eauto.
+ intros. eapply (bc_glob bc); eauto.
}
set (bc' := BC f F_stack F_glob). unfold f in bc'.
assert (BC'EQ: forall b, bc b <> BCinvalid -> bc' b = bc b).
@@ -466,15 +466,15 @@ Proof.
(* Part 2: invariance properties *)
assert (SM: forall b p, bc b <> BCinvalid -> smatch bc m b p -> smatch bc' m' b Nonstack).
{
- intros.
+ intros.
apply smatch_incr with bc; auto.
- apply smatch_inv with m.
+ apply smatch_inv with m.
apply smatch_no_stack with p; auto.
- intros. eapply Mem.loadbytes_alloc_unchanged; eauto. eapply mmatch_below; eauto.
+ intros. eapply Mem.loadbytes_alloc_unchanged; eauto. eapply mmatch_below; eauto.
}
assert (SMSTACK: smatch bc' m' sp Pbot).
{
- split; intros.
+ split; intros.
exploit Mem.load_alloc_same; eauto. intros EQ. subst v. constructor.
exploit Mem.loadbytes_alloc_same; eauto with coqlib. congruence.
}
@@ -483,15 +483,15 @@ Proof.
- (* incr *)
assumption.
- (* sp is BCstack *)
- simpl; apply dec_eq_true.
+ simpl; apply dec_eq_true.
- (* genv match *)
eapply genv_match_exten; eauto.
simpl; intros. destruct (eq_block b sp); intuition congruence.
- simpl; intros. destruct (eq_block b sp); congruence.
+ simpl; intros. destruct (eq_block b sp); congruence.
- (* romatch *)
- apply romatch_exten with bc.
- eapply romatch_alloc; eauto. eapply mmatch_below; eauto.
- simpl; intros. destruct (eq_block b sp); intuition.
+ apply romatch_exten with bc.
+ eapply romatch_alloc; eauto. eapply mmatch_below; eauto.
+ simpl; intros. destruct (eq_block b sp); intuition.
- (* mmatch *)
constructor; simpl; intros.
+ (* stack *)
@@ -504,16 +504,16 @@ Proof.
destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto.
+ (* top *)
destruct (eq_block b sp).
- subst b. apply smatch_ge with Pbot. apply SMSTACK. constructor.
+ subst b. apply smatch_ge with Pbot. apply SMSTACK. constructor.
eapply SM; auto. eapply mmatch_top; eauto.
+ (* below *)
- red; simpl; intros. rewrite NB. destruct (eq_block b sp).
- subst b; rewrite SP; xomega.
- exploit mmatch_below; eauto. xomega.
+ red; simpl; intros. rewrite NB. destruct (eq_block b sp).
+ subst b; rewrite SP; xomega.
+ exploit mmatch_below; eauto. xomega.
- (* unchanged *)
simpl; intros. apply dec_eq_false. apply Plt_ne. auto.
- (* values *)
- intros. apply vmatch_incr with bc; auto. eapply vmatch_no_stack; eauto.
+ intros. apply vmatch_incr with bc; auto. eapply vmatch_no_stack; eauto.
Qed.
(** Construction 2: turn the stack into an "other" block, at public calls or function returns *)
@@ -540,15 +540,15 @@ Proof.
{
unfold f; intros.
destruct (eq_block b1 sp); try discriminate.
- destruct (eq_block b2 sp); try discriminate.
- eapply bc_stack; eauto.
+ destruct (eq_block b2 sp); try discriminate.
+ eapply bc_stack; eauto.
}
assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2).
{
unfold f; intros.
destruct (eq_block b1 sp); try discriminate.
- destruct (eq_block b2 sp); try discriminate.
- eapply bc_glob; eauto.
+ destruct (eq_block b2 sp); try discriminate.
+ eapply bc_glob; eauto.
}
set (bc' := BC f F_stack F_glob). unfold f in bc'.
@@ -556,7 +556,7 @@ Proof.
assert (PM: forall b ofs p, pmatch bc b ofs p -> pmatch bc' b ofs Ptop).
{
intros. assert (pmatch bc b ofs Ptop) by (eapply pmatch_top'; eauto).
- inv H0. constructor; simpl. destruct (eq_block b sp); congruence.
+ inv H0. constructor; simpl. destruct (eq_block b sp); congruence.
}
assert (VM: forall v x, vmatch bc v x -> vmatch bc' v Vtop).
{
@@ -571,32 +571,32 @@ Proof.
(* Conclusions *)
exists bc'; splitall.
- (* nostack *)
- red; simpl; intros. destruct (eq_block b sp). congruence.
- red; intros. elim n. eapply bc_stack; eauto.
+ red; simpl; intros. destruct (eq_block b sp). congruence.
+ red; intros. elim n. eapply bc_stack; eauto.
- (* bc' sp is BCother *)
- simpl; apply dec_eq_true.
+ simpl; apply dec_eq_true.
- (* other blocks *)
- intros; simpl; apply dec_eq_false; auto.
+ intros; simpl; apply dec_eq_false; auto.
- (* values *)
auto.
- (* genv *)
- apply genv_match_exten with bc; auto.
+ apply genv_match_exten with bc; auto.
simpl; intros. destruct (eq_block b sp); intuition congruence.
simpl; intros. destruct (eq_block b sp); auto.
- (* romatch *)
- apply romatch_exten with bc; auto.
- simpl; intros. destruct (eq_block b sp); intuition.
+ apply romatch_exten with bc; auto.
+ simpl; intros. destruct (eq_block b sp); intuition.
- (* mmatch top *)
- constructor; simpl; intros.
+ constructor; simpl; intros.
+ destruct (eq_block b sp). congruence. elim n. eapply bc_stack; eauto.
+ rewrite PTree.gempty in H0; discriminate.
+ destruct (eq_block b sp).
subst b. eapply SM. eapply mmatch_stack; eauto.
- eapply SM. eapply mmatch_nonstack; eauto.
+ eapply SM. eapply mmatch_nonstack; eauto.
+ destruct (eq_block b sp).
subst b. eapply SM. eapply mmatch_stack; eauto.
eapply SM. eapply mmatch_top; eauto.
- + red; simpl; intros. destruct (eq_block b sp).
+ + red; simpl; intros. destruct (eq_block b sp).
subst b. eapply mmatch_below; eauto. congruence.
eapply mmatch_below; eauto.
Qed.
@@ -626,15 +626,15 @@ Proof.
{
unfold f; intros.
destruct (eq_block b1 sp); try discriminate.
- destruct (eq_block b2 sp); try discriminate.
- eapply bc_stack; eauto.
+ destruct (eq_block b2 sp); try discriminate.
+ eapply bc_stack; eauto.
}
assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2).
{
unfold f; intros.
destruct (eq_block b1 sp); try discriminate.
- destruct (eq_block b2 sp); try discriminate.
- eapply bc_glob; eauto.
+ destruct (eq_block b2 sp); try discriminate.
+ eapply bc_glob; eauto.
}
set (bc' := BC f F_stack F_glob). unfold f in bc'.
@@ -642,11 +642,11 @@ Proof.
assert (PM: forall b ofs p, pge Nonstack p -> pmatch bc b ofs p -> pmatch bc' b ofs Ptop).
{
intros. assert (pmatch bc b ofs Nonstack) by (eapply pmatch_ge; eauto).
- inv H1. constructor; simpl; destruct (eq_block b sp); congruence.
+ inv H1. constructor; simpl; destruct (eq_block b sp); congruence.
}
assert (VM: forall v x, vge (Ifptr Nonstack) x -> vmatch bc v x -> vmatch bc' v Vtop).
{
- intros. apply vmatch_ifptr; intros. subst v.
+ intros. apply vmatch_ifptr; intros. subst v.
inv H0; inv H; eapply PM; eauto.
}
assert (SM: forall b p, pge Nonstack p -> smatch bc m b p -> smatch bc' m b Ptop).
@@ -658,31 +658,31 @@ Proof.
(* Conclusions *)
exists bc'; splitall.
- (* nostack *)
- red; simpl; intros. destruct (eq_block b sp). congruence.
- red; intros. elim n. eapply bc_stack; eauto.
+ red; simpl; intros. destruct (eq_block b sp). congruence.
+ red; intros. elim n. eapply bc_stack; eauto.
- (* bc' sp is BCinvalid *)
- simpl; apply dec_eq_true.
+ simpl; apply dec_eq_true.
- (* other blocks *)
- intros; simpl; apply dec_eq_false; auto.
+ intros; simpl; apply dec_eq_false; auto.
- (* values *)
auto.
- (* genv *)
- apply genv_match_exten with bc; auto.
+ apply genv_match_exten with bc; auto.
simpl; intros. destruct (eq_block b sp); intuition congruence.
simpl; intros. destruct (eq_block b sp); congruence.
- (* romatch *)
- apply romatch_exten with bc; auto.
- simpl; intros. destruct (eq_block b sp); intuition.
+ apply romatch_exten with bc; auto.
+ simpl; intros. destruct (eq_block b sp); intuition.
- (* mmatch top *)
- constructor; simpl; intros.
+ constructor; simpl; intros.
+ destruct (eq_block b sp). congruence. elim n. eapply bc_stack; eauto.
+ rewrite PTree.gempty in H0; discriminate.
+ destruct (eq_block b sp). congruence.
- eapply SM. eauto. eapply mmatch_nonstack; eauto.
+ eapply SM. eauto. eapply mmatch_nonstack; eauto.
+ destruct (eq_block b sp). congruence.
- eapply SM. eauto. eapply mmatch_nonstack; eauto.
- red; intros; elim n. eapply bc_stack; eauto.
- + red; simpl; intros. destruct (eq_block b sp). congruence.
+ eapply SM. eauto. eapply mmatch_nonstack; eauto.
+ red; intros; elim n. eapply bc_stack; eauto.
+ + red; simpl; intros. destruct (eq_block b sp). congruence.
eapply mmatch_below; eauto.
Qed.
@@ -718,29 +718,29 @@ Proof.
{
assert (forall b, f b = BCstack -> b = sp).
{ unfold f; intros. destruct (eq_block b sp); auto. eelim NOSTACK; eauto. }
- intros. transitivity sp; auto. symmetry; auto.
+ intros. transitivity sp; auto. symmetry; auto.
}
assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2).
{
assert (forall b id, f b = BCglob id -> callee b = BCglob id).
{ unfold f; intros. destruct (eq_block b sp). congruence. auto. }
- intros. eapply (bc_glob callee); eauto.
+ intros. eapply (bc_glob callee); eauto.
}
set (bc := BC f F_stack F_glob). unfold f in bc.
assert (INCR: bc_incr caller bc).
{
- red; simpl; intros. destruct (eq_block b sp). congruence.
- symmetry; apply SAME; auto.
+ red; simpl; intros. destruct (eq_block b sp). congruence.
+ symmetry; apply SAME; auto.
}
(* Invariance properties *)
- assert (PM: forall b ofs p, pmatch callee b ofs p -> pmatch bc b ofs Ptop).
+ assert (PM: forall b ofs p, pmatch callee b ofs p -> pmatch bc b ofs Ptop).
{
intros. assert (pmatch callee b ofs Ptop) by (eapply pmatch_top'; eauto).
inv H0. constructor; simpl. destruct (eq_block b sp); congruence.
}
assert (VM: forall v x, vmatch callee v x -> vmatch bc v Vtop).
{
- intros. assert (vmatch callee v0 Vtop) by (eapply vmatch_top; eauto).
+ intros. assert (vmatch callee v0 Vtop) by (eapply vmatch_top; eauto).
inv H0; constructor; eauto.
}
assert (SM: forall b p, smatch callee m b p -> smatch bc m b Ptop).
@@ -752,38 +752,38 @@ Proof.
- (* result value *)
eapply VM; eauto.
- (* environment *)
- eapply ematch_incr; eauto.
+ eapply ematch_incr; eauto.
- (* romem *)
apply romatch_exten with callee; auto.
- intros; simpl. destruct (eq_block b sp); intuition.
+ intros; simpl. destruct (eq_block b sp); intuition.
- (* mmatch *)
constructor; simpl; intros.
+ (* stack *)
apply ablock_init_sound. destruct (eq_block b sp).
- subst b. eapply SM. eapply mmatch_nonstack; eauto. congruence.
+ subst b. eapply SM. eapply mmatch_nonstack; eauto. congruence.
elim (NOSTACK b); auto.
+ (* globals *)
rewrite PTree.gempty in H0; discriminate.
+ (* nonstack *)
destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto.
+ (* top *)
- eapply SM. eapply mmatch_top; eauto.
+ eapply SM. eapply mmatch_top; eauto.
destruct (eq_block b sp); congruence.
+ (* below *)
- red; simpl; intros. destruct (eq_block b sp).
+ red; simpl; intros. destruct (eq_block b sp).
subst b. eapply mmatch_below; eauto. congruence.
eapply mmatch_below; eauto.
- (* genv *)
eapply genv_match_exten with caller; eauto.
- simpl; intros. destruct (eq_block b sp). intuition congruence.
+ simpl; intros. destruct (eq_block b sp). intuition congruence.
split; intros. rewrite SAME in H by eauto with va. auto.
apply <- (proj1 GE2) in H. apply (proj1 GE1) in H. auto.
- simpl; intros. destruct (eq_block b sp). congruence.
+ simpl; intros. destruct (eq_block b sp). congruence.
rewrite <- SAME; eauto with va.
- (* sp *)
simpl. apply dec_eq_true.
- (* unchanged *)
- simpl; intros. destruct (eq_block b sp). congruence.
+ simpl; intros. destruct (eq_block b sp). congruence.
symmetry. apply SAME; auto. eapply Plt_trans. eauto. apply BELOW. congruence.
Qed.
@@ -820,19 +820,19 @@ Proof.
{
assert (forall b, f b = BCstack -> b = sp).
{ unfold f; intros. destruct (eq_block b sp); auto. eelim NOSTACK; eauto. }
- intros. transitivity sp; auto. symmetry; auto.
+ intros. transitivity sp; auto. symmetry; auto.
}
assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2).
{
assert (forall b id, f b = BCglob id -> callee b = BCglob id).
{ unfold f; intros. destruct (eq_block b sp). congruence. auto. }
- intros. eapply (bc_glob callee); eauto.
+ intros. eapply (bc_glob callee); eauto.
}
set (bc := BC f F_stack F_glob). unfold f in bc.
assert (INCR1: bc_incr caller bc).
{
- red; simpl; intros. destruct (eq_block b sp). congruence.
- symmetry; apply SAME; auto.
+ red; simpl; intros. destruct (eq_block b sp). congruence.
+ symmetry; apply SAME; auto.
}
assert (INCR2: bc_incr callee bc).
{
@@ -840,14 +840,14 @@ Proof.
}
(* Invariance properties *)
- assert (PM: forall b ofs p, pmatch callee b ofs p -> pmatch bc b ofs Nonstack).
+ assert (PM: forall b ofs p, pmatch callee b ofs p -> pmatch bc b ofs Nonstack).
{
intros. assert (pmatch callee b ofs Ptop) by (eapply pmatch_top'; eauto).
inv H0. constructor; simpl; destruct (eq_block b sp); congruence.
}
assert (VM: forall v x, vmatch callee v x -> vmatch bc v (Ifptr Nonstack)).
{
- intros. assert (vmatch callee v0 Vtop) by (eapply vmatch_top; eauto).
+ intros. assert (vmatch callee v0 Vtop) by (eapply vmatch_top; eauto).
inv H0; constructor; eauto.
}
assert (SM: forall b p, smatch callee m b p -> smatch bc m b Nonstack).
@@ -856,21 +856,21 @@ Proof.
}
assert (BSTK: bmatch bc m sp (am_stack am)).
{
- apply bmatch_incr with caller; eauto.
+ apply bmatch_incr with caller; eauto.
}
(* Conclusions *)
exists bc; splitall.
- (* result value *)
eapply VM; eauto.
- (* environment *)
- eapply ematch_incr; eauto.
+ eapply ematch_incr; eauto.
- (* romem *)
apply romatch_exten with callee; auto.
- intros; simpl. destruct (eq_block b sp); intuition.
+ intros; simpl. destruct (eq_block b sp); intuition.
- (* mmatch *)
constructor; simpl; intros.
+ (* stack *)
- destruct (eq_block b sp).
+ destruct (eq_block b sp).
subst b. exact BSTK.
elim (NOSTACK b); auto.
+ (* globals *)
@@ -878,12 +878,12 @@ Proof.
+ (* nonstack *)
destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto.
+ (* top *)
- destruct (eq_block b sp).
- subst. apply smatch_ge with (ab_summary (am_stack am)). apply BSTK. apply pge_lub_l.
+ destruct (eq_block b sp).
+ subst. apply smatch_ge with (ab_summary (am_stack am)). apply BSTK. apply pge_lub_l.
apply smatch_ge with Nonstack. eapply SM. eapply mmatch_top; eauto. apply pge_lub_r.
+ (* below *)
- red; simpl; intros. destruct (eq_block b sp).
- subst b. apply Plt_le_trans with bound. apply BELOW. congruence. auto.
+ red; simpl; intros. destruct (eq_block b sp).
+ subst b. apply Plt_le_trans with bound. apply BELOW. congruence. auto.
eapply mmatch_below; eauto.
- (* genv *)
eapply genv_match_exten; eauto.
@@ -892,7 +892,7 @@ Proof.
- (* sp *)
simpl. apply dec_eq_true.
- (* unchanged *)
- simpl; intros. destruct (eq_block b sp). congruence.
+ simpl; intros. destruct (eq_block b sp). congruence.
symmetry. apply SAME; auto. eapply Plt_trans. eauto. apply BELOW. congruence.
Qed.
@@ -919,19 +919,19 @@ Proof.
intros until am; intros EC GENV ARGS RO MM NOSTACK.
(* Part 1: using ec_mem_inject *)
exploit (@external_call_mem_inject ef _ _ ge vargs m t vres m' (inj_of_bc bc) m vargs).
- apply inj_of_bc_preserves_globals; auto.
+ apply inj_of_bc_preserves_globals; auto.
exact EC.
eapply mmatch_inj; eauto. eapply mmatch_below; eauto.
- revert ARGS. generalize vargs.
+ revert ARGS. generalize vargs.
induction vargs0; simpl; intros; constructor.
- eapply vmatch_inj; eauto. auto.
+ eapply vmatch_inj; eauto. auto.
intros (j' & vres' & m'' & EC' & IRES & IMEM & UNCH1 & UNCH2 & IINCR & ISEP).
assert (JBELOW: forall b, Plt b (Mem.nextblock m) -> j' b = inj_of_bc bc b).
{
intros. destruct (inj_of_bc bc b) as [[b' delta] | ] eqn:EQ.
- eapply IINCR; eauto.
+ eapply IINCR; eauto.
destruct (j' b) as [[b'' delta'] | ] eqn:EQ'; auto.
- exploit ISEP; eauto. tauto.
+ exploit ISEP; eauto. tauto.
}
(* Part 2: constructing bc' from j' *)
set (f := fun b => if plt b (Mem.nextblock m)
@@ -941,13 +941,13 @@ Proof.
{
assert (forall b, f b = BCstack -> bc b = BCstack).
{ unfold f; intros. destruct (plt b (Mem.nextblock m)); auto. destruct (j' b); discriminate. }
- intros. apply (bc_stack bc); auto.
+ intros. apply (bc_stack bc); auto.
}
assert (F_glob: forall b1 b2 id, f b1 = BCglob id -> f b2 = BCglob id -> b1 = b2).
{
assert (forall b id, f b = BCglob id -> bc b = BCglob id).
{ unfold f; intros. destruct (plt b (Mem.nextblock m)); auto. destruct (j' b); discriminate. }
- intros. eapply (bc_glob bc); eauto.
+ intros. eapply (bc_glob bc); eauto.
}
set (bc' := BC f F_stack F_glob). unfold f in bc'.
assert (INCR: bc_incr bc bc').
@@ -956,34 +956,34 @@ Proof.
}
assert (BC'INV: forall b, bc' b <> BCinvalid -> exists b' delta, j' b = Some(b', delta)).
{
- simpl; intros. destruct (plt b (Mem.nextblock m)).
+ simpl; intros. destruct (plt b (Mem.nextblock m)).
exists b, 0. rewrite JBELOW by auto. apply inj_of_bc_valid; auto.
- destruct (j' b) as [[b' delta] | ].
- exists b', delta; auto.
+ destruct (j' b) as [[b' delta] | ].
+ exists b', delta; auto.
congruence.
}
(* Part 3: injection wrt j' implies matching with top wrt bc' *)
assert (PMTOP: forall b b' delta ofs, j' b = Some (b', delta) -> pmatch bc' b ofs Ptop).
{
- intros. constructor. simpl; unfold f.
+ intros. constructor. simpl; unfold f.
destruct (plt b (Mem.nextblock m)).
- rewrite JBELOW in H by auto. eapply inj_of_bc_inv; eauto.
+ rewrite JBELOW in H by auto. eapply inj_of_bc_inv; eauto.
rewrite H; congruence.
}
assert (VMTOP: forall v v', Val.inject j' v v' -> vmatch bc' v Vtop).
{
- intros. inv H; constructor. eapply PMTOP; eauto.
+ intros. inv H; constructor. eapply PMTOP; eauto.
}
assert (SMTOP: forall b, bc' b <> BCinvalid -> smatch bc' m' b Ptop).
{
intros; split; intros.
- - exploit BC'INV; eauto. intros (b' & delta & J').
- exploit Mem.load_inject. eexact IMEM. eauto. eauto. intros (v' & A & B).
+ - exploit BC'INV; eauto. intros (b' & delta & J').
+ exploit Mem.load_inject. eexact IMEM. eauto. eauto. intros (v' & A & B).
eapply VMTOP; eauto.
- - exploit BC'INV; eauto. intros (b'' & delta & J').
+ - exploit BC'INV; eauto. intros (b'' & delta & J').
exploit Mem.loadbytes_inject. eexact IMEM. eauto. eauto. intros (bytes & A & B).
- inv B. inv H3. inv H7. eapply PMTOP; eauto.
+ inv B. inv H3. inv H7. eapply PMTOP; eauto.
}
(* Conclusions *)
exists bc'; splitall.
@@ -1004,29 +1004,29 @@ Proof.
exploit RO; eauto. intros (R & P & Q).
split; auto.
split. apply bmatch_incr with bc; auto. apply bmatch_inv with m; auto.
- intros. eapply Mem.loadbytes_unchanged_on_1. eapply external_call_readonly; eauto.
- auto. intros; red. apply Q.
- intros; red; intros; elim (Q ofs).
+ intros. eapply Mem.loadbytes_unchanged_on_1. eapply external_call_readonly; eauto.
+ auto. intros; red. apply Q.
+ intros; red; intros; elim (Q ofs).
eapply external_call_max_perm with (m2 := m'); eauto.
destruct (j' b); congruence.
- (* mmatch top *)
constructor; simpl; intros.
- + apply ablock_init_sound. apply SMTOP. simpl; congruence.
+ + apply ablock_init_sound. apply SMTOP. simpl; congruence.
+ rewrite PTree.gempty in H0; discriminate.
+ apply SMTOP; auto.
- + apply SMTOP; auto.
- + red; simpl; intros. destruct (plt b (Mem.nextblock m)).
- eapply Plt_le_trans. eauto. eapply external_call_nextblock; eauto.
- destruct (j' b) as [[bx deltax] | ] eqn:J'.
- eapply Mem.valid_block_inject_1; eauto.
+ + apply SMTOP; auto.
+ + red; simpl; intros. destruct (plt b (Mem.nextblock m)).
+ eapply Plt_le_trans. eauto. eapply external_call_nextblock; eauto.
+ destruct (j' b) as [[bx deltax] | ] eqn:J'.
+ eapply Mem.valid_block_inject_1; eauto.
congruence.
- (* nostack *)
- red; simpl; intros. destruct (plt b (Mem.nextblock m)).
+ red; simpl; intros. destruct (plt b (Mem.nextblock m)).
apply NOSTACK; auto.
destruct (j' b); congruence.
- (* unmapped blocks are invariant *)
intros. eapply Mem.loadbytes_unchanged_on_1; auto.
- apply UNCH1; auto. intros; red. unfold inj_of_bc; rewrite H0; auto.
+ apply UNCH1; auto. intros; red. unfold inj_of_bc; rewrite H0; auto.
Qed.
Remark list_forall2_in_l:
@@ -1036,8 +1036,8 @@ Proof.
induction 1; simpl; intros.
- contradiction.
- destruct H1.
- + subst. exists b1; auto.
- + exploit IHlist_forall2; eauto. intros (x2 & U & V). exists x2; auto.
+ + subst. exists b1; auto.
+ + exploit IHlist_forall2; eauto. intros (x2 & U & V). exists x2; auto.
Qed.
(** ** Semantic invariant *)
@@ -1122,10 +1122,10 @@ Lemma sound_stack_ext:
Proof.
induction 1; intros INV.
- constructor.
-- assert (Plt sp bound') by eauto with va.
+- assert (Plt sp bound') by eauto with va.
eapply sound_stack_public_call; eauto. apply IHsound_stack; intros.
apply INV. xomega. rewrite SAME; auto. xomega. auto. auto.
-- assert (Plt sp bound') by eauto with va.
+- assert (Plt sp bound') by eauto with va.
eapply sound_stack_private_call; eauto. apply IHsound_stack; intros.
apply INV. xomega. rewrite SAME; auto. xomega. auto. auto.
apply bmatch_ext with m; auto. intros. apply INV. xomega. auto. auto. auto.
@@ -1137,7 +1137,7 @@ Lemma sound_stack_inv:
(forall b ofs n, Plt b bound -> bc b = BCinvalid -> n >= 0 -> Mem.loadbytes m' b ofs n = Mem.loadbytes m b ofs n) ->
sound_stack bc stk m' bound.
Proof.
- intros. eapply sound_stack_ext; eauto. intros. rewrite <- H0; auto.
+ intros. eapply sound_stack_ext; eauto. intros. rewrite <- H0; auto.
Qed.
Lemma sound_stack_storev:
@@ -1147,13 +1147,13 @@ Lemma sound_stack_storev:
sound_stack bc stk m bound ->
sound_stack bc stk m' bound.
Proof.
- intros. apply sound_stack_inv with m; auto.
+ intros. apply sound_stack_inv with m; auto.
destruct addr; simpl in H; try discriminate.
assert (A: pmatch bc b i Ptop).
{ inv H0; eapply pmatch_top'; eauto. }
- inv A.
- intros. eapply Mem.loadbytes_store_other; eauto. left; congruence.
-Qed.
+ inv A.
+ intros. eapply Mem.loadbytes_store_other; eauto. left; congruence.
+Qed.
Lemma sound_stack_storebytes:
forall m b ofs bytes m' bc aaddr stk bound,
@@ -1162,12 +1162,12 @@ Lemma sound_stack_storebytes:
sound_stack bc stk m bound ->
sound_stack bc stk m' bound.
Proof.
- intros. apply sound_stack_inv with m; auto.
+ intros. apply sound_stack_inv with m; auto.
assert (A: pmatch bc b ofs Ptop).
{ inv H0; eapply pmatch_top'; eauto. }
- inv A.
- intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence.
-Qed.
+ inv A.
+ intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence.
+Qed.
Lemma sound_stack_free:
forall m b lo hi m' bc stk bound,
@@ -1185,10 +1185,10 @@ Lemma sound_stack_new_bound:
Ple bound bound' ->
sound_stack bc stk m bound'.
Proof.
- intros. inv H.
+ intros. inv H.
- constructor.
-- eapply sound_stack_public_call with (bound' := bound'0); eauto. xomega.
-- eapply sound_stack_private_call with (bound' := bound'0); eauto. xomega.
+- eapply sound_stack_public_call with (bound' := bound'0); eauto. xomega.
+- eapply sound_stack_private_call with (bound' := bound'0); eauto. xomega.
Qed.
Lemma sound_stack_exten:
@@ -1197,15 +1197,15 @@ Lemma sound_stack_exten:
(forall b, Plt b bound -> bc1 b = bc b) ->
sound_stack bc1 stk m bound.
Proof.
- intros. inv H.
+ intros. inv H.
- constructor.
-- assert (Plt sp bound') by eauto with va.
+- assert (Plt sp bound') by eauto with va.
eapply sound_stack_public_call; eauto.
- rewrite H0; auto. xomega.
+ rewrite H0; auto. xomega.
intros. rewrite H0; auto. xomega.
-- assert (Plt sp bound') by eauto with va.
+- assert (Plt sp bound') by eauto with va.
eapply sound_stack_private_call; eauto.
- rewrite H0; auto. xomega.
+ rewrite H0; auto. xomega.
intros. rewrite H0; auto. xomega.
Qed.
@@ -1226,7 +1226,7 @@ Lemma sound_succ_state:
sound_state (State s f (Vptr sp Int.zero) pc' e' m').
Proof.
intros. exploit analyze_succ; eauto. intros (ae'' & am'' & AN & EM & MM).
- econstructor; eauto.
+ econstructor; eauto.
Qed.
Theorem sound_step:
@@ -1235,72 +1235,72 @@ Proof.
induction 1; intros SOUND; inv SOUND.
- (* nop *)
- eapply sound_succ_state; eauto. simpl; auto.
+ eapply sound_succ_state; eauto. simpl; auto.
unfold transfer; rewrite H. auto.
- (* op *)
- eapply sound_succ_state; eauto. simpl; auto.
- unfold transfer; rewrite H. eauto.
+ eapply sound_succ_state; eauto. simpl; auto.
+ unfold transfer; rewrite H. eauto.
apply ematch_update; auto. eapply eval_static_operation_sound; eauto with va.
- (* load *)
- eapply sound_succ_state; eauto. simpl; auto.
- unfold transfer; rewrite H. eauto.
- apply ematch_update; auto. eapply loadv_sound; eauto with va.
+ eapply sound_succ_state; eauto. simpl; auto.
+ unfold transfer; rewrite H. eauto.
+ apply ematch_update; auto. eapply loadv_sound; eauto with va.
eapply eval_static_addressing_sound; eauto with va.
- (* store *)
exploit eval_static_addressing_sound; eauto with va. intros VMADDR.
- eapply sound_succ_state; eauto. simpl; auto.
- unfold transfer; rewrite H. eauto.
- eapply storev_sound; eauto.
- destruct a; simpl in H1; try discriminate. eapply romatch_store; eauto.
- eapply sound_stack_storev; eauto.
+ eapply sound_succ_state; eauto. simpl; auto.
+ unfold transfer; rewrite H. eauto.
+ eapply storev_sound; eauto.
+ destruct a; simpl in H1; try discriminate. eapply romatch_store; eauto.
+ eapply sound_stack_storev; eauto.
- (* call *)
assert (TR: transfer f rm pc ae am = transfer_call ae am args res).
{ unfold transfer; rewrite H; auto. }
- unfold transfer_call, analyze_call in TR.
- destruct (pincl (am_nonstack am) Nonstack &&
+ unfold transfer_call, analyze_call in TR.
+ destruct (pincl (am_nonstack am) Nonstack &&
forallb (fun av => vpincl av Nonstack) (aregs ae args)) eqn:NOLEAK.
+ (* private call *)
InvBooleans.
- exploit analyze_successor; eauto. simpl; eauto. rewrite TR. intros SUCC.
+ exploit analyze_successor; eauto. simpl; eauto. rewrite TR. intros SUCC.
exploit hide_stack; eauto. apply pincl_ge; auto.
intros (bc' & A & B & C & D & E & F & G).
apply sound_call_state with bc'; auto.
* eapply sound_stack_private_call with (bound' := Mem.nextblock m) (bc' := bc); eauto.
apply Ple_refl.
eapply mmatch_below; eauto.
- eapply mmatch_stack; eauto.
- * intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v.
- apply D with (areg ae r).
+ eapply mmatch_stack; eauto.
+ * intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v.
+ apply D with (areg ae r).
rewrite forallb_forall in H2. apply vpincl_ge.
apply H2. apply in_map; auto.
auto with va.
+ (* public call *)
- exploit analyze_successor; eauto. simpl; eauto. rewrite TR. intros SUCC.
+ exploit analyze_successor; eauto. simpl; eauto. rewrite TR. intros SUCC.
exploit anonymize_stack; eauto. intros (bc' & A & B & C & D & E & F & G).
apply sound_call_state with bc'; auto.
* eapply sound_stack_public_call with (bound' := Mem.nextblock m) (bc' := bc); eauto.
apply Ple_refl.
eapply mmatch_below; eauto.
- * intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v.
+ * intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v.
apply D with (areg ae r). auto with va.
- (* tailcall *)
exploit anonymize_stack; eauto. intros (bc' & A & B & C & D & E & F & G).
apply sound_call_state with bc'; auto.
- erewrite Mem.nextblock_free by eauto.
+ erewrite Mem.nextblock_free by eauto.
apply sound_stack_new_bound with stk.
apply sound_stack_exten with bc.
eapply sound_stack_free; eauto.
intros. apply C. apply Plt_ne; auto.
- apply Plt_Ple. eapply mmatch_below; eauto. congruence.
- intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v.
+ apply Plt_Ple. eapply mmatch_below; eauto. congruence.
+ intros. exploit list_in_map_inv; eauto. intros (r & P & Q). subst v.
apply D with (areg ae r). auto with va.
- eapply romatch_free; eauto.
- eapply mmatch_free; eauto.
+ eapply romatch_free; eauto.
+ eapply mmatch_free; eauto.
- (* builtin *)
assert (SPVALID: Plt sp0 (Mem.nextblock m)) by (eapply mmatch_below; eauto with va).
@@ -1314,72 +1314,72 @@ Proof.
{ unfold transfer_builtin_default, analyze_call; intros TR'.
set (aargs := map (abuiltin_arg ae am rm) args) in *.
assert (ARGS: list_forall2 (vmatch bc) vargs aargs) by (eapply abuiltin_args_sound; eauto).
- destruct (pincl (am_nonstack am) Nonstack &&
+ destruct (pincl (am_nonstack am) Nonstack &&
forallb (fun av => vpincl av Nonstack) aargs)
eqn: NOLEAK.
* (* private builtin call *)
InvBooleans. rewrite forallb_forall in H3.
exploit hide_stack; eauto. apply pincl_ge; auto.
intros (bc1 & A & B & C & D & E & F & G).
- exploit external_call_match; eauto.
+ exploit external_call_match; eauto.
intros. exploit list_forall2_in_l; eauto. intros (av & U & V).
- eapply D; eauto with va. apply vpincl_ge. apply H3; auto.
+ eapply D; eauto with va. apply vpincl_ge. apply H3; auto.
intros (bc2 & J & K & L & M & N & O & P & Q).
exploit (return_from_private_call bc bc2); eauto.
eapply mmatch_below; eauto.
rewrite K; auto.
intros. rewrite K; auto. rewrite C; auto.
- apply bmatch_inv with m. eapply mmatch_stack; eauto.
+ apply bmatch_inv with m. eapply mmatch_stack; eauto.
intros. apply Q; auto.
- eapply external_call_nextblock; eauto.
+ eapply external_call_nextblock; eauto.
intros (bc3 & U & V & W & X & Y & Z & AA).
- eapply sound_succ_state with (bc := bc3); eauto. simpl; auto.
- apply set_builtin_res_sound; auto.
- apply sound_stack_exten with bc.
+ eapply sound_succ_state with (bc := bc3); eauto. simpl; auto.
+ apply set_builtin_res_sound; auto.
+ apply sound_stack_exten with bc.
apply sound_stack_inv with m. auto.
intros. apply Q. red. eapply Plt_trans; eauto.
rewrite C; auto.
exact AA.
* (* public builtin call *)
- exploit anonymize_stack; eauto.
+ exploit anonymize_stack; eauto.
intros (bc1 & A & B & C & D & E & F & G).
- exploit external_call_match; eauto.
+ exploit external_call_match; eauto.
intros. exploit list_forall2_in_l; eauto. intros (av & U & V). eapply D; eauto with va.
intros (bc2 & J & K & L & M & N & O & P & Q).
exploit (return_from_public_call bc bc2); eauto.
eapply mmatch_below; eauto.
rewrite K; auto.
intros. rewrite K; auto. rewrite C; auto.
- eapply external_call_nextblock; eauto.
+ eapply external_call_nextblock; eauto.
intros (bc3 & U & V & W & X & Y & Z & AA).
- eapply sound_succ_state with (bc := bc3); eauto. simpl; auto.
- apply set_builtin_res_sound; auto.
- apply sound_stack_exten with bc.
+ eapply sound_succ_state with (bc := bc3); eauto. simpl; auto.
+ apply set_builtin_res_sound; auto.
+ apply sound_stack_exten with bc.
apply sound_stack_inv with m. auto.
intros. apply Q. red. eapply Plt_trans; eauto.
rewrite C; auto.
exact AA.
}
- unfold transfer_builtin in TR.
+ unfold transfer_builtin in TR.
destruct ef; auto.
+ (* volatile load *)
inv H0; auto. inv H3; auto. inv H1.
exploit abuiltin_arg_sound; eauto. intros VM1.
eapply sound_succ_state; eauto. simpl; auto.
- apply set_builtin_res_sound; auto.
+ apply set_builtin_res_sound; auto.
inv H3.
* (* true volatile access *)
assert (V: vmatch bc v (Ifptr Glob)).
{ inv H4; simpl in *; constructor. econstructor. eapply GE; eauto. }
- destruct (va_strict tt). apply vmatch_lub_r. apply vnormalize_sound. auto.
+ destruct (va_strict tt). apply vmatch_lub_r. apply vnormalize_sound. auto.
apply vnormalize_sound. eapply vmatch_ge; eauto. constructor. constructor.
* (* normal memory access *)
exploit loadv_sound; eauto. simpl; eauto. intros V.
- destruct (va_strict tt).
+ destruct (va_strict tt).
apply vmatch_lub_l. auto.
- eapply vnormalize_cast; eauto. eapply vmatch_top; eauto.
+ eapply vnormalize_cast; eauto. eapply vmatch_top; eauto.
+ (* volatile store *)
- inv H0; auto. inv H3; auto. inv H4; auto. inv H1.
+ inv H0; auto. inv H3; auto. inv H4; auto. inv H1.
exploit abuiltin_arg_sound. eauto. eauto. eauto. eauto. eauto. eexact H0. intros VM1.
exploit abuiltin_arg_sound. eauto. eauto. eauto. eauto. eauto. eexact H2. intros VM2.
inv H9.
@@ -1394,84 +1394,84 @@ Proof.
eapply romatch_store; eauto.
eapply sound_stack_storev; eauto. simpl; eauto.
+ (* memcpy *)
- inv H0; auto. inv H3; auto. inv H4; auto. inv H1.
+ inv H0; auto. inv H3; auto. inv H4; auto. inv H1.
exploit abuiltin_arg_sound. eauto. eauto. eauto. eauto. eauto. eexact H0. intros VM1.
exploit abuiltin_arg_sound. eauto. eauto. eauto. eauto. eauto. eexact H2. intros VM2.
- eapply sound_succ_state; eauto. simpl; auto.
+ eapply sound_succ_state; eauto. simpl; auto.
apply set_builtin_res_sound; auto. constructor.
- eapply storebytes_sound; eauto.
- apply match_aptr_of_aval; auto.
- eapply Mem.loadbytes_length; eauto.
- intros. eapply loadbytes_sound; eauto. apply match_aptr_of_aval; auto.
- eapply romatch_storebytes; eauto.
- eapply sound_stack_storebytes; eauto.
+ eapply storebytes_sound; eauto.
+ apply match_aptr_of_aval; auto.
+ eapply Mem.loadbytes_length; eauto.
+ intros. eapply loadbytes_sound; eauto. apply match_aptr_of_aval; auto.
+ eapply romatch_storebytes; eauto.
+ eapply sound_stack_storebytes; eauto.
+ (* annot *)
- inv H1. eapply sound_succ_state; eauto. simpl; auto. apply set_builtin_res_sound; auto. constructor.
+ inv H1. eapply sound_succ_state; eauto. simpl; auto. apply set_builtin_res_sound; auto. constructor.
+ (* annot val *)
inv H0; auto. inv H3; auto. inv H1.
eapply sound_succ_state; eauto. simpl; auto.
apply set_builtin_res_sound; auto. eapply abuiltin_arg_sound; eauto.
+ (* debug *)
- inv H1. eapply sound_succ_state; eauto. simpl; auto. apply set_builtin_res_sound; auto. constructor.
+ inv H1. eapply sound_succ_state; eauto. simpl; auto. apply set_builtin_res_sound; auto. constructor.
- (* cond *)
- eapply sound_succ_state; eauto.
- simpl. destruct b; auto.
- unfold transfer; rewrite H; auto.
+ eapply sound_succ_state; eauto.
+ simpl. destruct b; auto.
+ unfold transfer; rewrite H; auto.
- (* jumptable *)
- eapply sound_succ_state; eauto.
- simpl. eapply list_nth_z_in; eauto.
+ eapply sound_succ_state; eauto.
+ simpl. eapply list_nth_z_in; eauto.
unfold transfer; rewrite H; auto.
- (* return *)
exploit anonymize_stack; eauto. intros (bc' & A & B & C & D & E & F & G).
apply sound_return_state with bc'; auto.
- erewrite Mem.nextblock_free by eauto.
+ erewrite Mem.nextblock_free by eauto.
apply sound_stack_new_bound with stk.
apply sound_stack_exten with bc.
eapply sound_stack_free; eauto.
intros. apply C. apply Plt_ne; auto.
apply Plt_Ple. eapply mmatch_below; eauto with va.
- destruct or; simpl. eapply D; eauto. constructor.
- eapply romatch_free; eauto.
+ destruct or; simpl. eapply D; eauto. constructor.
+ eapply romatch_free; eauto.
eapply mmatch_free; eauto.
- (* internal function *)
- exploit allocate_stack; eauto.
+ exploit allocate_stack; eauto.
intros (bc' & A & B & C & D & E & F & G).
- exploit (analyze_entrypoint rm f args m' bc'); eauto.
+ exploit (analyze_entrypoint rm f args m' bc'); eauto.
intros (ae & am & AN & EM & MM').
- econstructor; eauto.
- erewrite Mem.alloc_result by eauto.
+ econstructor; eauto.
+ erewrite Mem.alloc_result by eauto.
apply sound_stack_exten with bc; auto.
- apply sound_stack_inv with m; auto.
+ apply sound_stack_inv with m; auto.
intros. eapply Mem.loadbytes_alloc_unchanged; eauto.
intros. apply F. erewrite Mem.alloc_result by eauto. auto.
- (* external function *)
exploit external_call_match; eauto with va.
intros (bc' & A & B & C & D & E & F & G & K).
- econstructor; eauto.
+ econstructor; eauto.
apply sound_stack_new_bound with (Mem.nextblock m).
apply sound_stack_exten with bc; auto.
- apply sound_stack_inv with m; auto.
+ apply sound_stack_inv with m; auto.
eapply external_call_nextblock; eauto.
- (* return *)
inv STK.
+ (* from public call *)
- exploit return_from_public_call; eauto.
+ exploit return_from_public_call; eauto.
intros; rewrite SAME; auto.
- intros (bc1 & A & B & C & D & E & F & G).
+ intros (bc1 & A & B & C & D & E & F & G).
destruct (analyze rm f)#pc as [ |ae' am'] eqn:EQ; simpl in AN; try contradiction. destruct AN as [A1 A2].
eapply sound_regular_state with (bc := bc1); eauto.
apply sound_stack_exten with bc'; auto.
- eapply ematch_ge; eauto. apply ematch_update. auto. auto.
+ eapply ematch_ge; eauto. apply ematch_update. auto. auto.
+ (* from private call *)
- exploit return_from_private_call; eauto.
+ exploit return_from_private_call; eauto.
intros; rewrite SAME; auto.
- intros (bc1 & A & B & C & D & E & F & G).
+ intros (bc1 & A & B & C & D & E & F & G).
destruct (analyze rm f)#pc as [ |ae' am'] eqn:EQ; simpl in AN; try contradiction. destruct AN as [A1 A2].
eapply sound_regular_state with (bc := bc1); eauto.
apply sound_stack_exten with bc'; auto.
@@ -1498,8 +1498,8 @@ Lemma initial_block_classification:
/\ (forall b id, bc b = BCglob id -> Genv.find_symbol ge id = Some b)
/\ (forall b, Mem.valid_block m b -> bc b <> BCinvalid).
Proof.
- intros.
- set (f := fun b =>
+ intros.
+ set (f := fun b =>
if plt b (Genv.genv_next ge) then
match Genv.invert_symbol ge b with None => BCother | Some id => BCglob id end
else
@@ -1511,8 +1511,8 @@ Proof.
destruct (Genv.invert_symbol ge b1) as [id1|] eqn:I1; inv H0.
destruct (plt b2 (Genv.genv_next ge)); try discriminate.
destruct (Genv.invert_symbol ge b2) as [id2|] eqn:I2; inv H1.
- exploit Genv.invert_find_symbol. eexact I1.
- exploit Genv.invert_find_symbol. eexact I2.
+ exploit Genv.invert_find_symbol. eexact I1.
+ exploit Genv.invert_find_symbol. eexact I2.
congruence.
}
assert (F_stack: forall b1 b2, f b1 = BCstack -> f b2 = BCstack -> b1 = b2).
@@ -1523,19 +1523,19 @@ Proof.
}
set (bc := BC f F_stack F_glob). unfold f in bc.
exists bc; splitall.
-- split; simpl; intros.
+- split; simpl; intros.
+ split; intros.
* rewrite pred_dec_true by (eapply Genv.genv_symb_range; eauto).
erewrite Genv.find_invert_symbol; eauto.
- * apply Genv.invert_find_symbol.
+ * apply Genv.invert_find_symbol.
destruct (plt b (Genv.genv_next ge)); try discriminate.
destruct (Genv.invert_symbol ge b); congruence.
- + rewrite ! pred_dec_true by assumption.
+ + rewrite ! pred_dec_true by assumption.
destruct (Genv.invert_symbol); split; congruence.
- red; simpl; intros. destruct (plt b (Genv.genv_next ge)); try congruence.
- erewrite <- Genv.init_mem_genv_next by eauto. auto.
-- red; simpl; intros.
- destruct (plt b (Genv.genv_next ge)).
+ erewrite <- Genv.init_mem_genv_next by eauto. auto.
+- red; simpl; intros.
+ destruct (plt b (Genv.genv_next ge)).
destruct (Genv.invert_symbol ge b); congruence.
congruence.
- simpl; intros. destruct (plt b (Genv.genv_next ge)); try discriminate.
@@ -1561,13 +1561,13 @@ Proof.
vge (Ifptr Glob) av ->
pge Glob (ab_summary (ablock_store chunk ab p av))).
{
- intros. simpl. unfold vplub; destruct av; auto.
+ intros. simpl. unfold vplub; destruct av; auto.
inv H0. apply plub_least; auto.
inv H0. apply plub_least; auto.
}
destruct id; auto.
- simpl. destruct (propagate_float_constants tt); auto.
- simpl. destruct (propagate_float_constants tt); auto.
+ simpl. destruct (propagate_float_constants tt); auto.
+ simpl. destruct (propagate_float_constants tt); auto.
apply DFL. constructor. constructor.
Qed.
@@ -1576,7 +1576,7 @@ Lemma store_init_data_list_summary:
pge Glob (ab_summary ab) ->
pge Glob (ab_summary (store_init_data_list ab p idl)).
Proof.
- induction idl; simpl; intros. auto. apply IHidl. apply store_init_data_summary; auto.
+ induction idl; simpl; intros. auto. apply IHidl. apply store_init_data_summary; auto.
Qed.
Lemma store_init_data_sound:
@@ -1588,7 +1588,7 @@ Proof.
intros. destruct id; try (eapply ablock_store_sound; eauto; constructor).
simpl. destruct (propagate_float_constants tt); eapply ablock_store_sound; eauto; constructor.
simpl. destruct (propagate_float_constants tt); eapply ablock_store_sound; eauto; constructor.
- simpl in H. inv H. auto.
+ simpl in H. inv H. auto.
simpl in H. destruct (Genv.find_symbol ge i) as [b'|] eqn:FS; try discriminate.
eapply ablock_store_sound; eauto. constructor. constructor. apply GMATCH; auto.
Qed.
@@ -1602,7 +1602,7 @@ Proof.
induction idl; simpl; intros.
- inv H; auto.
- destruct (Genv.store_init_data ge m b p a) as [m1|] eqn:SI; try discriminate.
- eapply IHidl; eauto. eapply store_init_data_sound; eauto.
+ eapply IHidl; eauto. eapply store_init_data_sound; eauto.
Qed.
Lemma store_init_data_other:
@@ -1614,7 +1614,7 @@ Lemma store_init_data_other:
Proof.
intros. eapply bmatch_inv; eauto.
intros. destruct id; try (eapply Mem.loadbytes_store_other; eauto; fail); simpl in H.
- inv H; auto.
+ inv H; auto.
destruct (Genv.find_symbol ge i); try discriminate.
eapply Mem.loadbytes_store_other; eauto.
Qed.
@@ -1626,10 +1626,10 @@ Lemma store_init_data_list_other:
bmatch bc m b' ab ->
bmatch bc m' b' ab.
Proof.
- induction idl; simpl; intros.
+ induction idl; simpl; intros.
inv H; auto.
destruct (Genv.store_init_data ge m b p a) as [m1|] eqn:SI; try discriminate.
- eapply IHidl; eauto. eapply store_init_data_other; eauto.
+ eapply IHidl; eauto. eapply store_init_data_other; eauto.
Qed.
Lemma store_zeros_same:
@@ -1640,8 +1640,8 @@ Lemma store_zeros_same:
Proof.
intros until n. functional induction (store_zeros m b pos n); intros.
- inv H. auto.
-- eapply IHo; eauto. change p with (vplub (I Int.zero) p).
- eapply smatch_store; eauto. constructor.
+- eapply IHo; eauto. change p with (vplub (I Int.zero) p).
+ eapply smatch_store; eauto. constructor.
- discriminate.
Qed.
@@ -1654,8 +1654,8 @@ Lemma store_zeros_other:
Proof.
intros until n. functional induction (store_zeros m b p n); intros.
- inv H. auto.
-- eapply IHo; eauto. eapply bmatch_inv; eauto.
- intros. eapply Mem.loadbytes_store_other; eauto.
+- eapply IHo; eauto. eapply bmatch_inv; eauto.
+ intros. eapply Mem.loadbytes_store_other; eauto.
- discriminate.
Qed.
@@ -1673,16 +1673,16 @@ Lemma alloc_global_match:
initial_mem_match bc m' (Genv.add_global g idg).
Proof.
intros; red; intros. destruct idg as [id [fd | gv]]; simpl in *.
-- destruct (Mem.alloc m 0 1) as [m1 b1] eqn:ALLOC.
+- destruct (Mem.alloc m 0 1) as [m1 b1] eqn:ALLOC.
unfold Genv.find_var_info, Genv.add_global in H2; simpl in H2.
- assert (Plt b (Mem.nextblock m)).
+ assert (Plt b (Mem.nextblock m)).
{ rewrite <- H. eapply Genv.genv_vars_range; eauto. }
- assert (b <> b1).
+ assert (b <> b1).
{ apply Plt_ne. erewrite Mem.alloc_result by eauto. auto. }
- apply bmatch_inv with m.
- eapply H0; eauto.
- intros. transitivity (Mem.loadbytes m1 b ofs n).
- eapply Mem.loadbytes_drop; eauto.
+ apply bmatch_inv with m.
+ eapply H0; eauto.
+ intros. transitivity (Mem.loadbytes m1 b ofs n).
+ eapply Mem.loadbytes_drop; eauto.
eapply Mem.loadbytes_alloc_unchanged; eauto.
- set (sz := Genv.init_data_list_size (gvar_init gv)) in *.
destruct (Mem.alloc m 0 sz) as [m1 b1] eqn:ALLOC.
@@ -1690,28 +1690,28 @@ Proof.
destruct (Genv.store_init_data_list ge m2 b1 0 (gvar_init gv)) as [m3 | ] eqn:SIDL; try discriminate.
unfold Genv.find_var_info, Genv.add_global in H2; simpl in H2.
rewrite PTree.gsspec in H2. destruct (peq b (Genv.genv_next g)).
-+ inversion H2; clear H2; subst v.
++ inversion H2; clear H2; subst v.
assert (b = b1). { erewrite Mem.alloc_result by eauto. congruence. }
- clear e. subst b.
- apply bmatch_inv with m3.
- eapply store_init_data_list_sound; eauto.
+ clear e. subst b.
+ apply bmatch_inv with m3.
+ eapply store_init_data_list_sound; eauto.
apply ablock_init_sound.
- eapply store_zeros_same; eauto.
- split; intros.
- exploit Mem.load_alloc_same; eauto. intros EQ; subst v; constructor.
+ eapply store_zeros_same; eauto.
+ split; intros.
+ exploit Mem.load_alloc_same; eauto. intros EQ; subst v; constructor.
exploit Mem.loadbytes_alloc_same; eauto with coqlib. congruence.
- intros. eapply Mem.loadbytes_drop; eauto.
- right; right; right. unfold Genv.perm_globvar. rewrite H3, H4. constructor.
-+ assert (Plt b (Mem.nextblock m)).
+ intros. eapply Mem.loadbytes_drop; eauto.
+ right; right; right. unfold Genv.perm_globvar. rewrite H3, H4. constructor.
++ assert (Plt b (Mem.nextblock m)).
{ rewrite <- H. eapply Genv.genv_vars_range; eauto. }
- assert (b <> b1).
+ assert (b <> b1).
{ apply Plt_ne. erewrite Mem.alloc_result by eauto. auto. }
apply bmatch_inv with m3.
- eapply store_init_data_list_other; eauto.
- eapply store_zeros_other; eauto.
- apply bmatch_inv with m.
- eapply H0; eauto.
- intros. eapply Mem.loadbytes_alloc_unchanged; eauto.
+ eapply store_init_data_list_other; eauto.
+ eapply store_zeros_other; eauto.
+ apply bmatch_inv with m.
+ eapply H0; eauto.
+ intros. eapply Mem.loadbytes_alloc_unchanged; eauto.
intros. eapply Mem.loadbytes_drop; eauto.
Qed.
@@ -1722,12 +1722,12 @@ Lemma alloc_globals_match:
Genv.alloc_globals ge m gl = Some m' ->
initial_mem_match bc m' (Genv.add_globals g gl).
Proof.
- induction gl; simpl; intros.
-- inv H1; auto.
+ induction gl; simpl; intros.
+- inv H1; auto.
- destruct (Genv.alloc_global ge m a) as [m1|] eqn:AG; try discriminate.
- eapply IHgl; eauto.
- erewrite Genv.alloc_global_nextblock; eauto. simpl. congruence.
- eapply alloc_global_match; eauto.
+ eapply IHgl; eauto.
+ erewrite Genv.alloc_global_nextblock; eauto. simpl. congruence.
+ eapply alloc_global_match; eauto.
Qed.
Definition romem_consistent (g: genv) (rm: romem) :=
@@ -1747,19 +1747,19 @@ Proof.
intros; red; intros. destruct idg as [id1 [fd1 | v1]];
unfold Genv.add_global, Genv.find_symbol, Genv.find_var_info, alloc_global in *; simpl in *.
- rewrite PTree.gsspec in H0. rewrite PTree.grspec in H1. unfold PTree.elt_eq in *.
- destruct (peq id id1). congruence. eapply H; eauto.
+ destruct (peq id id1). congruence. eapply H; eauto.
- rewrite PTree.gsspec in H0. destruct (peq id id1).
-+ inv H0. rewrite PTree.gss.
++ inv H0. rewrite PTree.gss.
destruct (gvar_readonly v1 && negb (gvar_volatile v1) && definitive_initializer (gvar_init v1)) eqn:RO.
InvBooleans. rewrite negb_true_iff in H4.
rewrite PTree.gss in H1.
- exists v1. intuition congruence.
+ exists v1. intuition congruence.
rewrite PTree.grs in H1. discriminate.
-+ rewrite PTree.gso. eapply H; eauto.
++ rewrite PTree.gso. eapply H; eauto.
destruct (gvar_readonly v1 && negb (gvar_volatile v1) && definitive_initializer (gvar_init v1)).
- rewrite PTree.gso in H1; auto.
- rewrite PTree.gro in H1; auto.
- apply Plt_ne. eapply Genv.genv_symb_range; eauto.
+ rewrite PTree.gso in H1; auto.
+ rewrite PTree.gro in H1; auto.
+ apply Plt_ne. eapply Genv.genv_symb_range; eauto.
Qed.
Lemma alloc_globals_consistent:
@@ -1767,14 +1767,14 @@ Lemma alloc_globals_consistent:
romem_consistent g rm ->
romem_consistent (Genv.add_globals g gl) (List.fold_left alloc_global gl rm).
Proof.
- induction gl; simpl; intros. auto. apply IHgl. apply alloc_global_consistent; auto.
+ induction gl; simpl; intros. auto. apply IHgl. apply alloc_global_consistent; auto.
Qed.
End INIT.
Theorem initial_mem_matches:
forall m,
- Genv.init_mem prog = Some m ->
+ Genv.init_mem prog = Some m ->
exists bc,
genv_match bc ge
/\ bc_below bc (Mem.nextblock m)
@@ -1783,7 +1783,7 @@ Theorem initial_mem_matches:
/\ (forall b, Mem.valid_block m b -> bc b <> BCinvalid).
Proof.
intros.
- exploit initial_block_classification; eauto. intros (bc & GE & BELOW & NOSTACK & INV & VALID).
+ exploit initial_block_classification; eauto. intros (bc & GE & BELOW & NOSTACK & INV & VALID).
exists bc; splitall; auto.
assert (A: initial_mem_match bc m ge).
{
@@ -1796,34 +1796,34 @@ Proof.
red; intros. rewrite PTree.gempty in H1; discriminate.
}
red; intros.
- exploit B; eauto. intros (v & FV & RO & NVOL & EQ).
- split. subst ab. apply store_init_data_list_summary. constructor.
- split. subst ab. eapply A; eauto.
- unfold ge in FV; exploit Genv.init_mem_characterization; eauto.
- intros (P & Q & R).
- intros; red; intros. exploit Q; eauto. intros [U V].
- unfold Genv.perm_globvar in V; rewrite RO, NVOL in V. inv V.
+ exploit B; eauto. intros (v & FV & RO & NVOL & EQ).
+ split. subst ab. apply store_init_data_list_summary. constructor.
+ split. subst ab. eapply A; eauto.
+ unfold ge in FV; exploit Genv.init_mem_characterization; eauto.
+ intros (P & Q & R).
+ intros; red; intros. exploit Q; eauto. intros [U V].
+ unfold Genv.perm_globvar in V; rewrite RO, NVOL in V. inv V.
Qed.
-End INITIAL.
+End INITIAL.
Require Import Axioms.
Theorem sound_initial:
forall prog st, initial_state prog st -> sound_state prog st.
Proof.
- destruct 1.
+ destruct 1.
exploit initial_mem_matches; eauto. intros (bc & GE & BELOW & NOSTACK & RM & VALID).
- apply sound_call_state with bc.
-- constructor.
-- simpl; tauto.
+ apply sound_call_state with bc.
+- constructor.
+- simpl; tauto.
- exact RM.
- apply mmatch_inj_top with m0.
replace (inj_of_bc bc) with (Mem.flat_inj (Mem.nextblock m0)).
eapply Genv.initmem_inject; eauto.
symmetry; apply extensionality; unfold Mem.flat_inj; intros x.
- destruct (plt x (Mem.nextblock m0)).
- apply inj_of_bc_valid; auto.
+ destruct (plt x (Mem.nextblock m0)).
+ apply inj_of_bc_valid; auto.
unfold inj_of_bc. erewrite bc_below_invalid; eauto.
- exact GE.
- exact NOSTACK.
@@ -1848,7 +1848,7 @@ Lemma avalue_sound:
/\ bc sp = BCstack.
Proof.
intros. inv H. exists bc; split; auto. rewrite AN. apply EM.
-Qed.
+Qed.
Definition aaddr (a: VA.t) (r: reg) : aptr :=
match a with
@@ -1867,7 +1867,7 @@ Lemma aaddr_sound:
Proof.
intros. inv H. exists bc; split; auto.
unfold aaddr; rewrite AN. apply match_aptr_of_aval. rewrite <- H0. apply EM.
-Qed.
+Qed.
Definition aaddressing (a: VA.t) (addr: addressing) (args: list reg) : aptr :=
match a with
@@ -1884,8 +1884,8 @@ Lemma aaddressing_sound:
/\ genv_match bc (Genv.globalenv prog)
/\ bc sp = BCstack.
Proof.
- intros. inv H. exists bc; split; auto.
- unfold aaddressing. rewrite AN. apply match_aptr_of_aval.
+ intros. inv H. exists bc; split; auto.
+ unfold aaddressing. rewrite AN. apply match_aptr_of_aval.
eapply eval_static_addressing_sound; eauto with va.
Qed.
@@ -1895,7 +1895,7 @@ Qed.
Definition aaddr_arg (a: VA.t) (ba: builtin_arg reg) : aptr :=
match a with
| VA.Bot => Pbot
- | VA.State ae am =>
+ | VA.State ae am =>
match ba with
| BA r => aptr_of_aval (AE.get r ae)
| BA_addrstack ofs => Stk ofs
@@ -1914,7 +1914,7 @@ Lemma aaddr_arg_sound_1:
eval_builtin_arg ge (fun r : positive => rs # r) (Vptr sp Int.zero) m a (Vptr b ofs) ->
pmatch bc b ofs (aaddr_arg (VA.State ae am) a).
Proof.
- intros.
+ intros.
apply pmatch_ge with (aptr_of_aval (abuiltin_arg ae am rm a)).
simpl. destruct a; try (apply pge_top); simpl; apply pge_refl.
apply match_aptr_of_aval. eapply abuiltin_arg_sound; eauto.
diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v
index 3d95bdd1..048ab816 100644
--- a/backend/ValueDomain.v
+++ b/backend/ValueDomain.v
@@ -47,7 +47,7 @@ Definition bc_below (bc: block_classification) (bound: block) : Prop :=
Lemma bc_below_invalid:
forall b bc bound, ~Plt b bound -> bc_below bc bound -> bc b = BCinvalid.
Proof.
- intros. destruct (block_class_eq (bc b) BCinvalid); auto.
+ intros. destruct (block_class_eq (bc b) BCinvalid); auto.
elim H. apply H0; auto.
Qed.
@@ -96,7 +96,7 @@ Lemma cmatch_lub_l:
Proof.
intros. unfold club; inv H; destruct y; try constructor;
destruct (eqb b b0) eqn:EQ; try constructor.
- replace b0 with b by (apply eqb_prop; auto). constructor.
+ replace b0 with b by (apply eqb_prop; auto). constructor.
Qed.
Lemma cmatch_lub_r:
@@ -136,7 +136,7 @@ Inductive aptr : Type :=
Definition eq_aptr: forall (p1 p2: aptr), {p1=p2} + {p1<>p2}.
Proof.
- intros. generalize ident_eq, Int.eq_dec; intros. decide equality.
+ intros. generalize ident_eq, Int.eq_dec; intros. decide equality.
Defined.
Inductive pmatch (b: block) (ofs: int): aptr -> Prop :=
@@ -192,7 +192,7 @@ Qed.
Lemma pmatch_top': forall b ofs p, pmatch b ofs p -> pmatch b ofs Ptop.
Proof.
- intros. apply pmatch_ge with p; auto with va.
+ intros. apply pmatch_ge with p; auto with va.
Qed.
Definition plub (p q: aptr) : aptr :=
@@ -215,7 +215,7 @@ Definition plub (p q: aptr) : aptr :=
Nonstack
| Stk ofs1, Stk ofs2 =>
if Int.eq_dec ofs1 ofs2 then p else Stack
- | (Stk _ | Stack), Stack =>
+ | (Stk _ | Stack), Stack =>
Stack
| Stack, Stk _ =>
Stack
@@ -226,22 +226,22 @@ Lemma plub_comm:
forall p q, plub p q = plub q p.
Proof.
intros; unfold plub; destruct p; destruct q; auto.
- destruct (ident_eq id id0). subst id0.
- rewrite dec_eq_true.
+ destruct (ident_eq id id0). subst id0.
+ rewrite dec_eq_true.
destruct (Int.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true. auto.
- rewrite dec_eq_false by auto. auto.
- rewrite dec_eq_false by auto. auto.
- destruct (ident_eq id id0). subst id0.
+ rewrite dec_eq_false by auto. auto.
+ rewrite dec_eq_false by auto. auto.
+ destruct (ident_eq id id0). subst id0.
rewrite dec_eq_true; auto.
rewrite dec_eq_false; auto.
- destruct (ident_eq id id0). subst id0.
+ destruct (ident_eq id id0). subst id0.
rewrite dec_eq_true; auto.
rewrite dec_eq_false; auto.
- destruct (ident_eq id id0). subst id0.
+ destruct (ident_eq id id0). subst id0.
rewrite dec_eq_true; auto.
rewrite dec_eq_false; auto.
destruct (Int.eq_dec ofs ofs0). subst ofs0. rewrite dec_eq_true; auto.
- rewrite dec_eq_false; auto.
+ rewrite dec_eq_false; auto.
Qed.
Lemma pge_lub_l:
@@ -283,7 +283,7 @@ Proof.
- unfold plub; destruct q; repeat rewrite dec_eq_true; constructor.
- rewrite dec_eq_true; constructor.
- rewrite dec_eq_true; constructor.
-- rewrite dec_eq_true. destruct (Int.eq_dec ofs ofs0); constructor.
+- rewrite dec_eq_true. destruct (Int.eq_dec ofs ofs0); constructor.
- destruct (ident_eq id id0). destruct (Int.eq_dec ofs ofs0); constructor. constructor.
- destruct (ident_eq id id0); constructor.
- destruct (ident_eq id id0); constructor.
@@ -328,7 +328,7 @@ Lemma pincl_sound:
forall b ofs p q,
pincl p q = true -> pmatch b ofs p -> pmatch b ofs q.
Proof.
- intros. eapply pmatch_ge; eauto. apply pincl_ge; auto.
+ intros. eapply pmatch_ge; eauto. apply pincl_ge; auto.
Qed.
Definition padd (p: aptr) (n: int) : aptr :=
@@ -343,7 +343,7 @@ Lemma padd_sound:
pmatch b ofs p ->
pmatch b (Int.add ofs delta) (padd p delta).
Proof.
- intros. inv H; simpl padd; eauto with va.
+ intros. inv H; simpl padd; eauto with va.
Qed.
Definition psub (p: aptr) (n: int) : aptr :=
@@ -358,7 +358,7 @@ Lemma psub_sound:
pmatch b ofs p ->
pmatch b (Int.sub ofs delta) (psub p delta).
Proof.
- intros. inv H; simpl psub; eauto with va.
+ intros. inv H; simpl psub; eauto with va.
Qed.
Definition poffset (p: aptr) : aptr :=
@@ -373,7 +373,7 @@ Lemma poffset_sound:
pmatch b ofs1 p ->
pmatch b ofs2 (poffset p).
Proof.
- intros. inv H; simpl poffset; eauto with va.
+ intros. inv H; simpl poffset; eauto with va.
Qed.
Definition psub2 (p q: aptr) : option int :=
@@ -442,7 +442,7 @@ Lemma pcmp_sound:
cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2)) (pcmp c p1 p2).
Proof.
intros.
- assert (DIFF: b1 <> b2 ->
+ assert (DIFF: b1 <> b2 ->
cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2))
(cmp_different_blocks c)).
{
@@ -455,19 +455,19 @@ Proof.
cmatch (Val.cmpu_bool valid c (Vptr b1 ofs1) (Vptr b2 ofs2))
(Maybe (Int.cmpu c ofs1 ofs2))).
{
- intros. subst b2. simpl. rewrite dec_eq_true.
+ intros. subst b2. simpl. rewrite dec_eq_true.
destruct ((valid b1 (Int.unsigned ofs1) || valid b1 (Int.unsigned ofs1 - 1)) &&
(valid b1 (Int.unsigned ofs2) || valid b1 (Int.unsigned ofs2 - 1))); simpl.
- constructor.
+ constructor.
constructor.
}
unfold pcmp; inv H; inv H0; (apply cmatch_top || (apply DIFF; congruence) || idtac).
- - destruct (peq id id0). subst id0. apply SAME. eapply bc_glob; eauto.
+ - destruct (peq id id0). subst id0. apply SAME. eapply bc_glob; eauto.
auto with va.
- destruct (peq id id0); auto with va.
- destruct (peq id id0); auto with va.
- destruct (peq id id0); auto with va.
- - apply SAME. eapply bc_stack; eauto.
+ - apply SAME. eapply bc_stack; eauto.
Qed.
Lemma pcmp_none:
@@ -539,7 +539,7 @@ Definition Vtop := Ifptr Ptop.
Definition eq_aval: forall (v1 v2: aval), {v1=v2} + {v1<>v2}.
Proof.
intros. generalize zeq Int.eq_dec Int64.eq_dec Float.eq_dec Float32.eq_dec eq_aptr; intros.
- decide equality.
+ decide equality.
Defined.
Definition is_uns (n: Z) (i: int) : Prop :=
@@ -570,7 +570,7 @@ Lemma vmatch_ifptr:
(forall b ofs, v = Vptr b ofs -> pmatch b ofs p) ->
vmatch v (Ifptr p).
Proof.
- intros. destruct v; constructor; auto.
+ intros. destruct v; constructor; auto.
Qed.
Lemma vmatch_top: forall v x, vmatch v x -> vmatch v Vtop.
@@ -604,8 +604,8 @@ Definition ssize (i: int) := Int.size (if Int.lt i Int.zero then Int.not i else
Lemma is_uns_usize:
forall i, is_uns (usize i) i.
Proof.
- unfold usize; intros; red; intros.
- apply Int.bits_size_2. omega.
+ unfold usize; intros; red; intros.
+ apply Int.bits_size_2. omega.
Qed.
Lemma is_sgn_ssize:
@@ -627,7 +627,7 @@ Qed.
Lemma is_uns_zero_ext:
forall n i, is_uns n i <-> Int.zero_ext n i = i.
Proof.
- intros; split; intros.
+ intros; split; intros.
Int.bit_solve. destruct (zlt i0 n); auto. symmetry; apply H; auto. omega.
rewrite <- H. red; intros. rewrite Int.bits_zero_ext by omega. rewrite zlt_false by omega. auto.
Qed.
@@ -635,12 +635,12 @@ Qed.
Lemma is_sgn_sign_ext:
forall n i, 0 < n -> (is_sgn n i <-> Int.sign_ext n i = i).
Proof.
- intros; split; intros.
+ intros; split; intros.
Int.bit_solve. destruct (zlt i0 n); auto.
transitivity (Int.testbit i (Int.zwordsize - 1)).
- apply H0; omega. symmetry; apply H0; omega.
+ apply H0; omega. symmetry; apply H0; omega.
rewrite <- H0. red; intros. rewrite ! Int.bits_sign_ext by omega.
- f_equal. transitivity (n-1). destruct (zlt m n); omega.
+ f_equal. transitivity (n-1). destruct (zlt m n); omega.
destruct (zlt (Int.zwordsize - 1) n); omega.
Qed.
@@ -649,7 +649,7 @@ Lemma is_zero_ext_uns:
is_uns m i \/ n <= m -> is_uns m (Int.zero_ext n i).
Proof.
intros. red; intros. rewrite Int.bits_zero_ext by omega.
- destruct (zlt m0 n); auto. destruct H. apply H; omega. omegaContradiction.
+ destruct (zlt m0 n); auto. destruct H. apply H; omega. omegaContradiction.
Qed.
Lemma is_zero_ext_sgn:
@@ -657,7 +657,7 @@ Lemma is_zero_ext_sgn:
n < m ->
is_sgn m (Int.zero_ext n i).
Proof.
- intros. red; intros. rewrite ! Int.bits_zero_ext by omega.
+ intros. red; intros. rewrite ! Int.bits_zero_ext by omega.
transitivity false. apply zlt_false; omega.
symmetry; apply zlt_false; omega.
Qed.
@@ -668,18 +668,18 @@ Lemma is_sign_ext_uns:
is_uns m i ->
is_uns m (Int.sign_ext n i).
Proof.
- intros; red; intros. rewrite Int.bits_sign_ext by omega.
+ intros; red; intros. rewrite Int.bits_sign_ext by omega.
apply H0. destruct (zlt m0 n); omega. destruct (zlt m0 n); omega.
Qed.
-
+
Lemma is_sign_ext_sgn:
forall i n m,
0 < n -> 0 < m ->
is_sgn m i \/ n <= m -> is_sgn m (Int.sign_ext n i).
Proof.
intros. apply is_sgn_sign_ext; auto.
- destruct (zlt m n). destruct H1. apply is_sgn_sign_ext in H1; auto.
- rewrite <- H1. rewrite (Int.sign_ext_widen i) by omega. apply Int.sign_ext_idem; auto.
+ destruct (zlt m n). destruct H1. apply is_sgn_sign_ext in H1; auto.
+ rewrite <- H1. rewrite (Int.sign_ext_widen i) by omega. apply Int.sign_ext_idem; auto.
omegaContradiction.
apply Int.sign_ext_widen; omega.
Qed.
@@ -690,7 +690,7 @@ Lemma is_uns_1:
forall n, is_uns 1 n -> n = Int.zero \/ n = Int.one.
Proof.
intros. destruct (Int.testbit n 0) eqn:B0; [right|left]; apply Int.same_bits_eq; intros.
- rewrite Int.bits_one. destruct (zeq i 0). subst i; auto. apply H; omega.
+ rewrite Int.bits_one. destruct (zeq i 0). subst i; auto. apply H; omega.
rewrite Int.bits_zero. destruct (zeq i 0). subst i; auto. apply H; omega.
Qed.
@@ -744,8 +744,8 @@ Lemma vmatch_uns':
Proof.
intros.
assert (A: forall n', n' >= 0 -> n' >= n -> is_uns n' i) by (eauto with va).
- unfold uns.
- destruct (zle n 1). auto with va.
+ unfold uns.
+ destruct (zle n 1). auto with va.
destruct (zle n 7). auto with va.
destruct (zle n 8). auto with va.
destruct (zle n 15). auto with va.
@@ -761,8 +761,8 @@ Qed.
Lemma vmatch_uns_undef: forall p n, vmatch Vundef (uns p n).
Proof.
- intros. unfold uns.
- destruct (zle n 1). auto with va.
+ intros. unfold uns.
+ destruct (zle n 1). auto with va.
destruct (zle n 7). auto with va.
destruct (zle n 8). auto with va.
destruct (zle n 15). auto with va.
@@ -774,7 +774,7 @@ Lemma vmatch_sgn':
Proof.
intros.
assert (A: forall n', n' >= 1 -> n' >= n -> is_sgn n' i) by (eauto with va).
- unfold sgn.
+ unfold sgn.
destruct (zle n 8). auto with va.
destruct (zle n 16); auto with va.
Qed.
@@ -893,8 +893,8 @@ Proof.
intros. unfold vlub; destruct v; destruct w; auto.
- rewrite Int.eq_sym. predSpec Int.eq Int.eq_spec n0 n.
congruence.
- rewrite orb_comm.
- destruct (Int.lt n0 Int.zero || Int.lt n Int.zero); f_equal; apply Z.max_comm.
+ rewrite orb_comm.
+ destruct (Int.lt n0 Int.zero || Int.lt n Int.zero); f_equal; apply Z.max_comm.
- f_equal. apply plub_comm. apply Z.max_comm.
- f_equal. apply plub_comm. apply Z.max_comm.
- f_equal; apply plub_comm.
@@ -918,7 +918,7 @@ Qed.
Lemma vge_uns_uns': forall p n, vge (uns p n) (Uns p n).
Proof.
- unfold uns; intros.
+ unfold uns; intros.
destruct (zle n 1). auto with va.
destruct (zle n 7). auto with va.
destruct (zle n 8). auto with va.
@@ -928,19 +928,19 @@ Qed.
Lemma vge_uns_i': forall p n i, 0 <= n -> is_uns n i -> vge (uns p n) (I i).
Proof.
- intros. apply vge_trans with (Uns p n). apply vge_uns_uns'. auto with va.
+ intros. apply vge_trans with (Uns p n). apply vge_uns_uns'. auto with va.
Qed.
Lemma vge_sgn_sgn': forall p n, vge (sgn p n) (Sgn p n).
Proof.
- unfold sgn; intros.
+ unfold sgn; intros.
destruct (zle n 8). auto with va.
destruct (zle n 16); auto with va.
Qed.
Lemma vge_sgn_i': forall p n i, 0 < n -> is_sgn n i -> vge (sgn p n) (I i).
Proof.
- intros. apply vge_trans with (Sgn p n). apply vge_sgn_sgn'. auto with va.
+ intros. apply vge_trans with (Sgn p n). apply vge_sgn_sgn'. auto with va.
Qed.
Hint Resolve vge_uns_uns' vge_uns_i' vge_sgn_sgn' vge_sgn_i' : va.
@@ -952,7 +952,7 @@ Qed.
Lemma ssize_pos: forall n, 0 < ssize n.
Proof.
- unfold ssize; intros.
+ unfold ssize; intros.
generalize (Int.size_range (if Int.lt n Int.zero then Int.not n else n)); omega.
Qed.
@@ -964,16 +964,16 @@ Proof.
unfold vlub; destruct x, y; eauto using pge_lub_l with va.
- predSpec Int.eq Int.eq_spec n n0. auto with va.
destruct (Int.lt n Int.zero || Int.lt n0 Int.zero).
- apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va.
- apply vge_uns_i'. generalize (usize_pos n); xomega. eauto with va.
+ apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va.
+ apply vge_uns_i'. generalize (usize_pos n); xomega. eauto with va.
- destruct (Int.lt n Int.zero).
- apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va.
- apply vge_uns_i'. generalize (usize_pos n); xomega. eauto with va.
-- apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va.
+ apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va.
+ apply vge_uns_i'. generalize (usize_pos n); xomega. eauto with va.
+- apply vge_sgn_i'. generalize (ssize_pos n); xomega. eauto with va.
- destruct (Int.lt n0 Int.zero).
eapply vge_trans. apply vge_sgn_sgn'.
apply vge_trans with (Sgn p (n + 1)); eauto with va.
- eapply vge_trans. apply vge_uns_uns'. eauto with va.
+ eapply vge_trans. apply vge_uns_uns'. eauto with va.
- eapply vge_trans. apply vge_sgn_sgn'.
apply vge_trans with (Sgn p (n + 1)); eauto using pge_lub_l with va.
- eapply vge_trans. apply vge_sgn_sgn'. eauto with va.
@@ -987,7 +987,7 @@ Qed.
Lemma vge_lub_r:
forall x y, vge (vlub x y) y.
Proof.
- intros. rewrite vlub_comm. apply vge_lub_l.
+ intros. rewrite vlub_comm. apply vge_lub_l.
Qed.
Lemma vmatch_lub_l:
@@ -1036,13 +1036,13 @@ Definition vplub (v: aval) (p: aptr) : aptr :=
Lemma vmatch_vplub_l:
forall v x p, vmatch v x -> vmatch v (Ifptr (vplub x p)).
Proof.
- intros. unfold vplub; inv H; auto with va; constructor; eapply pmatch_lub_l; eauto.
+ intros. unfold vplub; inv H; auto with va; constructor; eapply pmatch_lub_l; eauto.
Qed.
Lemma pmatch_vplub:
forall b ofs x p, pmatch b ofs p -> pmatch b ofs (vplub x p).
Proof.
- intros.
+ intros.
assert (DFL: pmatch b ofs (if va_strict tt then p else Ptop)).
{ destruct (va_strict tt); auto. eapply pmatch_top'; eauto. }
unfold vplub; destruct x; auto; apply pmatch_lub_r; auto.
@@ -1051,7 +1051,7 @@ Qed.
Lemma vmatch_vplub_r:
forall v x p, vmatch v (Ifptr p) -> vmatch v (Ifptr (vplub x p)).
Proof.
- intros. apply vmatch_ifptr; intros; subst v. inv H. apply pmatch_vplub; auto.
+ intros. apply vmatch_ifptr; intros; subst v. inv H. apply pmatch_vplub; auto.
Qed.
(** Inclusion *)
@@ -1065,13 +1065,13 @@ Definition vpincl (v: aval) (p: aptr) : bool :=
Lemma vpincl_ge:
forall x p, vpincl x p = true -> vge (Ifptr p) x.
Proof.
- unfold vpincl; intros. destruct x; constructor; apply pincl_ge; auto.
+ unfold vpincl; intros. destruct x; constructor; apply pincl_ge; auto.
Qed.
Lemma vpincl_sound:
forall v x p, vpincl x p = true -> vmatch v x -> vmatch v (Ifptr p).
Proof.
- intros. apply vmatch_ge with x; auto. apply vpincl_ge; auto.
+ intros. apply vmatch_ge with x; auto. apply vpincl_ge; auto.
Qed.
Definition vincl (v w: aval) : bool :=
@@ -1111,8 +1111,8 @@ Lemma symbol_address_sound:
genv_match ge ->
vmatch (Genv.symbol_address ge id ofs) (Ptr (Gl id ofs)).
Proof.
- intros. unfold Genv.symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:F.
- constructor. constructor. apply H; auto.
+ intros. unfold Genv.symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:F.
+ constructor. constructor. apply H; auto.
constructor.
Qed.
@@ -1122,8 +1122,8 @@ Lemma vmatch_ptr_gl:
vmatch v (Ptr (Gl id ofs)) ->
Val.lessdef v (Genv.symbol_address ge id ofs).
Proof.
- intros. unfold Genv.symbol_address. inv H0.
-- inv H3. replace (Genv.find_symbol ge id) with (Some b). constructor.
+ intros. unfold Genv.symbol_address. inv H0.
+- inv H3. replace (Genv.find_symbol ge id) with (Some b). constructor.
symmetry. apply H; auto.
- constructor.
Qed.
@@ -1134,7 +1134,7 @@ Lemma vmatch_ptr_stk:
bc sp = BCstack ->
Val.lessdef v (Vptr sp ofs).
Proof.
- intros. inv H.
+ intros. inv H.
- inv H3. replace b with sp by (eapply bc_stack; eauto). constructor.
- constructor.
Qed.
@@ -1223,19 +1223,19 @@ Definition shl (v w: aval) :=
| _ => ntop1 v
end.
-Lemma shl_sound:
+Lemma shl_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shl v w) (shl x y).
Proof.
intros.
- assert (DEFAULT: vmatch (Val.shl v w) (ntop1 x)).
+ assert (DEFAULT: vmatch (Val.shl v w) (ntop1 x)).
{
- destruct v; destruct w; simpl; try constructor.
+ destruct v; destruct w; simpl; try constructor.
destruct (Int.ltu i0 Int.iwordsize); constructor.
}
destruct y; auto. simpl. inv H0. unfold Val.shl.
destruct (Int.ltu n Int.iwordsize) eqn:LTU; auto.
exploit Int.ltu_inv; eauto. intros RANGE.
- inv H; auto with va.
+ inv H; auto with va.
- apply vmatch_uns'. red; intros. rewrite Int.bits_shl by omega.
destruct (zlt m (Int.unsigned n)). auto. apply H1; xomega.
- apply vmatch_sgn'. red; intros. zify.
@@ -1258,13 +1258,13 @@ Definition shru (v w: aval) :=
| _ => ntop1 v
end.
-Lemma shru_sound:
+Lemma shru_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shru v w) (shru x y).
Proof.
intros.
- assert (DEFAULT: vmatch (Val.shru v w) (ntop1 x)).
+ assert (DEFAULT: vmatch (Val.shru v w) (ntop1 x)).
{
- destruct v; destruct w; simpl; try constructor.
+ destruct v; destruct w; simpl; try constructor.
destruct (Int.ltu i0 Int.iwordsize); constructor.
}
destruct y; auto. inv H0. unfold shru, Val.shru.
@@ -1272,14 +1272,14 @@ Proof.
exploit Int.ltu_inv; eauto. intros RANGE. change (Int.unsigned Int.iwordsize) with Int.zwordsize in RANGE.
assert (DEFAULT2: forall i, vmatch (Vint (Int.shru i n)) (uns (provenance x) (Int.zwordsize - Int.unsigned n))).
{
- intros. apply vmatch_uns. red; intros.
- rewrite Int.bits_shru by omega. apply zlt_false. omega.
+ intros. apply vmatch_uns. red; intros.
+ rewrite Int.bits_shru by omega. apply zlt_false. omega.
}
inv H; auto with va.
- apply vmatch_uns'. red; intros. zify.
rewrite Int.bits_shru by omega.
destruct (zlt (m + Int.unsigned n) Int.zwordsize); auto.
- apply H1; omega.
+ apply H1; omega.
- destruct v; constructor.
Qed.
@@ -1297,13 +1297,13 @@ Definition shr (v w: aval) :=
| _ => ntop1 v
end.
-Lemma shr_sound:
+Lemma shr_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.shr v w) (shr x y).
Proof.
intros.
- assert (DEFAULT: vmatch (Val.shr v w) (ntop1 x)).
+ assert (DEFAULT: vmatch (Val.shr v w) (ntop1 x)).
{
- destruct v; destruct w; simpl; try constructor.
+ destruct v; destruct w; simpl; try constructor.
destruct (Int.ltu i0 Int.iwordsize); constructor.
}
destruct y; auto. inv H0. unfold shr, Val.shr.
@@ -1311,7 +1311,7 @@ Proof.
exploit Int.ltu_inv; eauto. intros RANGE. change (Int.unsigned Int.iwordsize) with Int.zwordsize in RANGE.
assert (DEFAULT2: forall i, vmatch (Vint (Int.shr i n)) (sgn (provenance x) (Int.zwordsize - Int.unsigned n))).
{
- intros. apply vmatch_sgn. red; intros.
+ intros. apply vmatch_sgn. red; intros.
rewrite ! Int.bits_shr by omega. f_equal.
destruct (zlt (m + Int.unsigned n) Int.zwordsize);
destruct (zlt (Int.zwordsize - 1 + Int.unsigned n) Int.zwordsize);
@@ -1346,12 +1346,12 @@ Definition and (v w: aval) :=
| _, _ => ntop2 v w
end.
-Lemma and_sound:
+Lemma and_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.and v w) (and x y).
Proof.
assert (UNS_l: forall i j n, is_uns n i -> is_uns n (Int.and i j)).
{
- intros; red; intros. rewrite Int.bits_and by auto. rewrite (H m) by auto.
+ intros; red; intros. rewrite Int.bits_and by auto. rewrite (H m) by auto.
apply andb_false_l.
}
assert (UNS_r: forall i j n, is_uns n i -> is_uns n (Int.and j i)).
@@ -1360,7 +1360,7 @@ Proof.
}
assert (UNS: forall i j n m, is_uns n i -> is_uns m j -> is_uns (Z.min n m) (Int.and i j)).
{
- intros. apply Z.min_case; auto.
+ intros. apply Z.min_case; auto.
}
assert (SGN: forall i j n m, is_sgn n i -> is_sgn m j -> is_sgn (Z.max n m) (Int.and i j)).
{
@@ -1379,12 +1379,12 @@ Definition or (v w: aval) :=
| _, _ => ntop2 v w
end.
-Lemma or_sound:
+Lemma or_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.or v w) (or x y).
Proof.
assert (UNS: forall i j n m, is_uns n i -> is_uns m j -> is_uns (Z.max n m) (Int.or i j)).
{
- intros; red; intros. rewrite Int.bits_or by auto.
+ intros; red; intros. rewrite Int.bits_or by auto.
rewrite H by xomega. rewrite H0 by xomega. auto.
}
assert (SGN: forall i j n m, is_sgn n i -> is_sgn m j -> is_sgn (Z.max n m) (Int.or i j)).
@@ -1404,12 +1404,12 @@ Definition xor (v w: aval) :=
| _, _ => ntop2 v w
end.
-Lemma xor_sound:
+Lemma xor_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.xor v w) (xor x y).
Proof.
assert (UNS: forall i j n m, is_uns n i -> is_uns m j -> is_uns (Z.max n m) (Int.xor i j)).
{
- intros; red; intros. rewrite Int.bits_xor by auto.
+ intros; red; intros. rewrite Int.bits_xor by auto.
rewrite H by xomega. rewrite H0 by xomega. auto.
}
assert (SGN: forall i j n m, is_sgn n i -> is_sgn m j -> is_sgn (Z.max n m) (Int.xor i j)).
@@ -1433,7 +1433,7 @@ Lemma notint_sound:
Proof.
assert (SGN: forall n i, is_sgn n i -> is_sgn n (Int.not i)).
{
- intros; red; intros. rewrite ! Int.bits_not by omega.
+ intros; red; intros. rewrite ! Int.bits_not by omega.
f_equal. apply H; auto.
}
intros. unfold Val.notint, notint; inv H; eauto with va.
@@ -1445,13 +1445,13 @@ Definition ror (x y: aval) :=
| _, _ => ntop1 x
end.
-Lemma ror_sound:
+Lemma ror_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.ror v w) (ror x y).
Proof.
- intros.
- assert (DEFAULT: forall p, vmatch (Val.ror v w) (Ifptr p)).
+ intros.
+ assert (DEFAULT: forall p, vmatch (Val.ror v w) (Ifptr p)).
{
- destruct v; destruct w; simpl; try constructor.
+ destruct v; destruct w; simpl; try constructor.
destruct (Int.ltu i0 Int.iwordsize); constructor.
}
unfold ror; destruct y; try apply DEFAULT; auto. inv H0. unfold Val.ror.
@@ -1473,7 +1473,7 @@ Proof.
intros.
assert (UNS_r: forall i j n, is_uns n j -> is_uns n (Int.and i j)).
{
- intros; red; intros. rewrite Int.bits_and by auto. rewrite (H0 m) by auto.
+ intros; red; intros. rewrite Int.bits_and by auto. rewrite (H0 m) by auto.
apply andb_false_r.
}
assert (UNS: forall i, vmatch (Vint (Int.rolm i amount mask))
@@ -1486,7 +1486,7 @@ Qed.
Definition neg := unop_int Int.neg.
-Lemma neg_sound:
+Lemma neg_sound:
forall v x, vmatch v x -> vmatch (Val.neg v) (neg x).
Proof (unop_int_sound Int.neg).
@@ -1506,8 +1506,8 @@ Lemma add_sound:
Proof.
intros. unfold Val.add, add; inv H; inv H0; constructor;
((apply padd_sound; assumption) || (eapply poffset_sound; eassumption) || idtac).
- apply pmatch_lub_r. eapply poffset_sound; eauto.
- apply pmatch_lub_l. eapply poffset_sound; eauto.
+ apply pmatch_lub_r. eapply poffset_sound; eauto.
+ apply pmatch_lub_l. eapply poffset_sound; eauto.
Qed.
Definition sub (v w: aval) :=
@@ -1524,7 +1524,7 @@ Definition sub (v w: aval) :=
| _, _ => ntop2 v w
end.
-Lemma sub_sound:
+Lemma sub_sound:
forall v w x y, vmatch v x -> vmatch w y -> vmatch (Val.sub v w) (sub x y).
Proof.
intros. inv H; subst; inv H0; subst; simpl;
@@ -1534,19 +1534,19 @@ Qed.
Definition mul := binop_int Int.mul.
-Lemma mul_sound:
+Lemma mul_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mul v w) (mul x y).
Proof (binop_int_sound Int.mul).
Definition mulhs := binop_int Int.mulhs.
-Lemma mulhs_sound:
+Lemma mulhs_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulhs v w) (mulhs x y).
Proof (binop_int_sound Int.mulhs).
Definition mulhu := binop_int Int.mulhu.
-Lemma mulhu_sound:
+Lemma mulhu_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulhu v w) (mulhu x y).
Proof (binop_int_sound Int.mulhu).
@@ -1563,7 +1563,7 @@ Definition divs (v w: aval) :=
Lemma divs_sound:
forall v w u x y, vmatch v x -> vmatch w y -> Val.divs v w = Some u -> vmatch u (divs x y).
Proof.
- intros. destruct v; destruct w; try discriminate; simpl in H1.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
destruct (Int.eq i0 Int.zero
|| Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone) eqn:E; inv H1.
inv H; inv H0; auto with va. simpl. rewrite E. constructor.
@@ -1581,7 +1581,7 @@ Definition divu (v w: aval) :=
Lemma divu_sound:
forall v w u x y, vmatch v x -> vmatch w y -> Val.divu v w = Some u -> vmatch u (divu x y).
Proof.
- intros. destruct v; destruct w; try discriminate; simpl in H1.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
destruct (Int.eq i0 Int.zero) eqn:E; inv H1.
inv H; inv H0; auto with va. simpl. rewrite E. constructor.
Qed.
@@ -1599,7 +1599,7 @@ Definition mods (v w: aval) :=
Lemma mods_sound:
forall v w u x y, vmatch v x -> vmatch w y -> Val.mods v w = Some u -> vmatch u (mods x y).
Proof.
- intros. destruct v; destruct w; try discriminate; simpl in H1.
+ intros. destruct v; destruct w; try discriminate; simpl in H1.
destruct (Int.eq i0 Int.zero
|| Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone) eqn:E; inv H1.
inv H; inv H0; auto with va. simpl. rewrite E. constructor.
@@ -1623,14 +1623,14 @@ Proof.
intros. apply is_uns_mon with (usize (Int.modu i j)); auto with va.
unfold usize, Int.size. apply Int.Zsize_monotone.
generalize (Int.unsigned_range_2 j); intros RANGE.
- assert (Int.unsigned j <> 0).
+ assert (Int.unsigned j <> 0).
{ red; intros; elim H. rewrite <- (Int.repr_unsigned j). rewrite H0. auto. }
exploit (Z_mod_lt (Int.unsigned i) (Int.unsigned j)). omega. intros MOD.
- unfold Int.modu. rewrite Int.unsigned_repr. omega. omega.
+ unfold Int.modu. rewrite Int.unsigned_repr. omega. omega.
}
intros. destruct v; destruct w; try discriminate; simpl in H1.
destruct (Int.eq i0 Int.zero) eqn:Z; inv H1.
- assert (i0 <> Int.zero) by (generalize (Int.eq_spec i0 Int.zero); rewrite Z; auto).
+ assert (i0 <> Int.zero) by (generalize (Int.eq_spec i0 Int.zero); rewrite Z; auto).
unfold modu. inv H; inv H0; auto with va. rewrite Z. constructor.
Qed.
@@ -1640,13 +1640,13 @@ Definition shrx (v w: aval) :=
| _, _ => ntop1 v
end.
-Lemma shrx_sound:
+Lemma shrx_sound:
forall v w u x y, vmatch v x -> vmatch w y -> Val.shrx v w = Some u -> vmatch u (shrx x y).
Proof.
intros.
- destruct v; destruct w; try discriminate; simpl in H1.
+ destruct v; destruct w; try discriminate; simpl in H1.
destruct (Int.ltu i0 (Int.repr 31)) eqn:LTU; inv H1.
- unfold shrx; inv H; auto with va; inv H0; auto with va.
+ unfold shrx; inv H; auto with va; inv H0; auto with va.
rewrite LTU; auto with va.
Qed.
@@ -1654,73 +1654,73 @@ Qed.
Definition negf := unop_float Float.neg.
-Lemma negf_sound:
+Lemma negf_sound:
forall v x, vmatch v x -> vmatch (Val.negf v) (negf x).
Proof (unop_float_sound Float.neg).
Definition absf := unop_float Float.abs.
-Lemma absf_sound:
+Lemma absf_sound:
forall v x, vmatch v x -> vmatch (Val.absf v) (absf x).
Proof (unop_float_sound Float.abs).
Definition addf := binop_float Float.add.
-Lemma addf_sound:
+Lemma addf_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.addf v w) (addf x y).
Proof (binop_float_sound Float.add).
Definition subf := binop_float Float.sub.
-Lemma subf_sound:
+Lemma subf_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.subf v w) (subf x y).
Proof (binop_float_sound Float.sub).
Definition mulf := binop_float Float.mul.
-Lemma mulf_sound:
+Lemma mulf_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulf v w) (mulf x y).
Proof (binop_float_sound Float.mul).
Definition divf := binop_float Float.div.
-Lemma divf_sound:
+Lemma divf_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.divf v w) (divf x y).
Proof (binop_float_sound Float.div).
Definition negfs := unop_single Float32.neg.
-Lemma negfs_sound:
+Lemma negfs_sound:
forall v x, vmatch v x -> vmatch (Val.negfs v) (negfs x).
Proof (unop_single_sound Float32.neg).
Definition absfs := unop_single Float32.abs.
-Lemma absfs_sound:
+Lemma absfs_sound:
forall v x, vmatch v x -> vmatch (Val.absfs v) (absfs x).
Proof (unop_single_sound Float32.abs).
Definition addfs := binop_single Float32.add.
-Lemma addfs_sound:
+Lemma addfs_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.addfs v w) (addfs x y).
Proof (binop_single_sound Float32.add).
Definition subfs := binop_single Float32.sub.
-Lemma subfs_sound:
+Lemma subfs_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.subfs v w) (subfs x y).
Proof (binop_single_sound Float32.sub).
Definition mulfs := binop_single Float32.mul.
-Lemma mulfs_sound:
+Lemma mulfs_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.mulfs v w) (mulfs x y).
Proof (binop_single_sound Float32.mul).
Definition divfs := binop_single Float32.div.
-Lemma divfs_sound:
+Lemma divfs_sound:
forall v x w y, vmatch v x -> vmatch w y -> vmatch (Val.divfs v w) (divfs x y).
Proof (binop_single_sound Float32.div).
@@ -1738,12 +1738,12 @@ Lemma zero_ext_sound:
Proof.
assert (DFL: forall nbits i, is_uns nbits (Int.zero_ext nbits i)).
{
- intros; red; intros. rewrite Int.bits_zero_ext by omega. apply zlt_false; auto.
+ intros; red; intros. rewrite Int.bits_zero_ext by omega. apply zlt_false; auto.
}
intros. inv H; simpl; auto with va. apply vmatch_uns.
red; intros. zify.
rewrite Int.bits_zero_ext by omega.
- destruct (zlt m nbits); auto. apply H1; omega.
+ destruct (zlt m nbits); auto. apply H1; omega.
Qed.
Definition sign_ext (nbits: Z) (v: aval) :=
@@ -1761,9 +1761,9 @@ Proof.
{
intros. apply vmatch_sgn. apply is_sign_ext_sgn; auto with va.
}
- intros. inv H0; simpl; auto with va.
+ intros. inv H0; simpl; auto with va.
- destruct (zlt n nbits); eauto with va.
- constructor; auto. eapply is_sign_ext_uns; eauto with va.
+ constructor; auto. eapply is_sign_ext_uns; eauto with va.
- destruct (zlt n nbits); auto with va.
- apply vmatch_sgn. apply is_sign_ext_sgn; auto with va.
apply Z.min_case; auto with va.
@@ -1778,10 +1778,10 @@ Definition singleoffloat (v: aval) :=
Lemma singleoffloat_sound:
forall v x, vmatch v x -> vmatch (Val.singleoffloat v) (singleoffloat x).
Proof.
- intros.
+ intros.
assert (DEFAULT: vmatch (Val.singleoffloat v) (ntop1 x)).
{ destruct v; constructor. }
- destruct x; auto. inv H. constructor.
+ destruct x; auto. inv H. constructor.
Qed.
Definition floatofsingle (v: aval) :=
@@ -1793,10 +1793,10 @@ Definition floatofsingle (v: aval) :=
Lemma floatofsingle_sound:
forall v x, vmatch v x -> vmatch (Val.floatofsingle v) (floatofsingle x).
Proof.
- intros.
+ intros.
assert (DEFAULT: vmatch (Val.floatofsingle v) (ntop1 x)).
{ destruct v; constructor. }
- destruct x; auto. inv H. constructor.
+ destruct x; auto. inv H. constructor.
Qed.
Definition intoffloat (x: aval) :=
@@ -1812,9 +1812,9 @@ Definition intoffloat (x: aval) :=
Lemma intoffloat_sound:
forall v x w, vmatch v x -> Val.intoffloat v = Some w -> vmatch w (intoffloat x).
Proof.
- unfold Val.intoffloat; intros. destruct v; try discriminate.
+ unfold Val.intoffloat; intros. destruct v; try discriminate.
destruct (Float.to_int f) as [i|] eqn:E; simpl in H0; inv H0.
- inv H; simpl; auto with va. rewrite E; constructor.
+ inv H; simpl; auto with va. rewrite E; constructor.
Qed.
Definition intuoffloat (x: aval) :=
@@ -1830,9 +1830,9 @@ Definition intuoffloat (x: aval) :=
Lemma intuoffloat_sound:
forall v x w, vmatch v x -> Val.intuoffloat v = Some w -> vmatch w (intuoffloat x).
Proof.
- unfold Val.intuoffloat; intros. destruct v; try discriminate.
+ unfold Val.intuoffloat; intros. destruct v; try discriminate.
destruct (Float.to_intu f) as [i|] eqn:E; simpl in H0; inv H0.
- inv H; simpl; auto with va. rewrite E; constructor.
+ inv H; simpl; auto with va. rewrite E; constructor.
Qed.
Definition floatofint (x: aval) :=
@@ -1874,9 +1874,9 @@ Definition intofsingle (x: aval) :=
Lemma intofsingle_sound:
forall v x w, vmatch v x -> Val.intofsingle v = Some w -> vmatch w (intofsingle x).
Proof.
- unfold Val.intofsingle; intros. destruct v; try discriminate.
+ unfold Val.intofsingle; intros. destruct v; try discriminate.
destruct (Float32.to_int f) as [i|] eqn:E; simpl in H0; inv H0.
- inv H; simpl; auto with va. rewrite E; constructor.
+ inv H; simpl; auto with va. rewrite E; constructor.
Qed.
Definition intuofsingle (x: aval) :=
@@ -1892,9 +1892,9 @@ Definition intuofsingle (x: aval) :=
Lemma intuofsingle_sound:
forall v x w, vmatch v x -> Val.intuofsingle v = Some w -> vmatch w (intuofsingle x).
Proof.
- unfold Val.intuofsingle; intros. destruct v; try discriminate.
+ unfold Val.intuofsingle; intros. destruct v; try discriminate.
destruct (Float32.to_intu f) as [i|] eqn:E; simpl in H0; inv H0.
- inv H; simpl; auto with va. rewrite E; constructor.
+ inv H; simpl; auto with va. rewrite E; constructor.
Qed.
Definition singleofint (x: aval) :=
@@ -2000,7 +2000,7 @@ Proof.
intros c [lo hi] x n; simpl; intros R.
destruct c; unfold zcmp, proj_sumbool.
- (* eq *)
- destruct (zlt n lo). rewrite zeq_false by omega. constructor.
+ destruct (zlt n lo). rewrite zeq_false by omega. constructor.
destruct (zlt hi n). rewrite zeq_false by omega. constructor.
constructor.
- (* ne *)
@@ -2055,7 +2055,7 @@ Proof.
intros. inv H; simpl; try (apply Int.unsigned_range_2).
- omega.
- destruct (zlt n0 Int.zwordsize); simpl.
-+ rewrite is_uns_zero_ext in H2. rewrite <- H2. rewrite Int.zero_ext_mod by omega.
++ rewrite is_uns_zero_ext in H2. rewrite <- H2. rewrite Int.zero_ext_mod by omega.
exploit (Z_mod_lt (Int.unsigned n) (two_p n0)). apply two_p_gt_ZERO; auto. omega.
+ apply Int.unsigned_range_2.
Qed.
@@ -2077,7 +2077,7 @@ Lemma cmpu_intv_sound_2:
vmatch (Vint n1) v1 ->
cmatch (Val.cmpu_bool valid c (Vint n2) (Vint n1)) (cmp_intv (swap_comparison c) (uintv v1) (Int.unsigned n2)).
Proof.
- intros. rewrite <- Val.swap_cmpu_bool. apply cmpu_intv_sound; auto.
+ intros. rewrite <- Val.swap_cmpu_bool. apply cmpu_intv_sound; auto.
Qed.
Definition sintv (v: aval) : Z * Z :=
@@ -2098,20 +2098,20 @@ Proof.
intros. inv H; simpl; try (apply Int.signed_range).
- omega.
- destruct (zlt n0 Int.zwordsize); simpl.
-+ rewrite is_uns_zero_ext in H2. rewrite <- H2.
++ rewrite is_uns_zero_ext in H2. rewrite <- H2.
assert (Int.unsigned (Int.zero_ext n0 n) = Int.unsigned n mod two_p n0) by (apply Int.zero_ext_mod; omega).
exploit (Z_mod_lt (Int.unsigned n) (two_p n0)). apply two_p_gt_ZERO; auto. intros.
replace (Int.signed (Int.zero_ext n0 n)) with (Int.unsigned (Int.zero_ext n0 n)).
- rewrite H. omega.
- unfold Int.signed. rewrite zlt_true. auto.
- assert (two_p n0 <= Int.half_modulus).
- { change Int.half_modulus with (two_p (Int.zwordsize - 1)).
+ rewrite H. omega.
+ unfold Int.signed. rewrite zlt_true. auto.
+ assert (two_p n0 <= Int.half_modulus).
+ { change Int.half_modulus with (two_p (Int.zwordsize - 1)).
apply two_p_monotone. omega. }
- omega.
+ omega.
+ apply Int.signed_range.
- destruct (zlt n0 (Int.zwordsize)); simpl.
-+ rewrite is_sgn_sign_ext in H2 by auto. rewrite <- H2.
- exploit (Int.sign_ext_range n0 n). omega. omega.
++ rewrite is_sgn_sign_ext in H2 by auto. rewrite <- H2.
+ exploit (Int.sign_ext_range n0 n). omega. omega.
+ apply Int.signed_range.
Qed.
@@ -2132,8 +2132,8 @@ Lemma cmp_intv_sound_2:
vmatch (Vint n1) v1 ->
cmatch (Val.cmp_bool c (Vint n2) (Vint n1)) (cmp_intv (swap_comparison c) (sintv v1) (Int.signed n2)).
Proof.
- intros. rewrite <- Val.swap_cmp_bool. apply cmp_intv_sound; auto.
-Qed.
+ intros. rewrite <- Val.swap_cmp_bool. apply cmp_intv_sound; auto.
+Qed.
(** Comparisons *)
@@ -2184,7 +2184,7 @@ Definition cmp_bool (c: comparison) (v w: aval) : abool :=
Lemma cmp_bool_sound:
forall c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmp_bool c v w) (cmp_bool c x y).
Proof.
- intros.
+ intros.
unfold cmp_bool; inversion H; subst; inversion H0; subst;
auto using cmatch_top, cmp_intv_sound, cmp_intv_sound_2, cmp_intv_None.
- constructor.
@@ -2199,7 +2199,7 @@ Definition cmpf_bool (c: comparison) (v w: aval) : abool :=
Lemma cmpf_bool_sound:
forall c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmpf_bool c v w) (cmpf_bool c x y).
Proof.
- intros. inv H; try constructor; inv H0; constructor.
+ intros. inv H; try constructor; inv H0; constructor.
Qed.
Definition cmpfs_bool (c: comparison) (v w: aval) : abool :=
@@ -2211,7 +2211,7 @@ Definition cmpfs_bool (c: comparison) (v w: aval) : abool :=
Lemma cmpfs_bool_sound:
forall c v w x y, vmatch v x -> vmatch w y -> cmatch (Val.cmpfs_bool c v w) (cmpfs_bool c x y).
Proof.
- intros. inv H; try constructor; inv H0; constructor.
+ intros. inv H; try constructor; inv H0; constructor.
Qed.
Definition maskzero (x: aval) (mask: int) : abool :=
@@ -2226,12 +2226,12 @@ Lemma maskzero_sound:
vmatch v x ->
cmatch (Val.maskzero_bool v mask) (maskzero x mask).
Proof.
- intros. inv H; simpl; auto with va.
+ intros. inv H; simpl; auto with va.
predSpec Int.eq Int.eq_spec (Int.zero_ext n mask) Int.zero; auto with va.
replace (Int.and i mask) with Int.zero.
rewrite Int.eq_true. constructor.
rewrite is_uns_zero_ext in H1. rewrite Int.zero_ext_and in * by auto.
- rewrite <- H1. rewrite Int.and_assoc. rewrite Int.and_commut in H. rewrite H.
+ rewrite <- H1. rewrite Int.and_assoc. rewrite Int.and_commut in H. rewrite H.
rewrite Int.and_zero; auto.
destruct (Int.eq (Int.zero_ext n mask) Int.zero); constructor.
Qed.
@@ -2249,7 +2249,7 @@ Proof.
assert (DEFAULT: vmatch (Val.of_optbool ob) (Uns Pbot 1)).
{
destruct ob; simpl; auto with va.
- destruct b; constructor; try omega.
+ destruct b; constructor; try omega.
change 1 with (usize Int.one). apply is_uns_usize.
red; intros. apply Int.bits_zero.
}
@@ -2329,16 +2329,16 @@ Lemma vnormalize_cast:
vmatch v (Ifptr p) ->
vmatch v (vnormalize chunk (Ifptr p)).
Proof.
- intros. exploit Mem.load_cast; eauto. exploit Mem.load_type; eauto.
+ intros. exploit Mem.load_cast; eauto. exploit Mem.load_type; eauto.
destruct chunk; simpl; intros.
- (* int8signed *)
- rewrite H2. destruct v; simpl; constructor. omega. apply is_sign_ext_sgn; auto with va.
+ rewrite H2. destruct v; simpl; constructor. omega. apply is_sign_ext_sgn; auto with va.
- (* int8unsigned *)
- rewrite H2. destruct v; simpl; constructor. omega. apply is_zero_ext_uns; auto with va.
+ rewrite H2. destruct v; simpl; constructor. omega. apply is_zero_ext_uns; auto with va.
- (* int16signed *)
- rewrite H2. destruct v; simpl; constructor. omega. apply is_sign_ext_sgn; auto with va.
+ rewrite H2. destruct v; simpl; constructor. omega. apply is_sign_ext_sgn; auto with va.
- (* int16unsigned *)
- rewrite H2. destruct v; simpl; constructor. omega. apply is_zero_ext_uns; auto with va.
+ rewrite H2. destruct v; simpl; constructor. omega. apply is_zero_ext_uns; auto with va.
- (* int32 *)
auto.
- (* int64 *)
@@ -2391,7 +2391,7 @@ Proof with (auto using provenance_monotone with va).
- constructor... apply is_zero_ext_uns...
- unfold provenance; destruct (va_strict tt)...
- destruct (va_strict tt)...
-- destruct (zlt n2 8); constructor...
+- destruct (zlt n2 8); constructor...
- destruct (zlt n2 16); constructor...
- destruct (va_strict tt)...
- destruct (va_strict tt)...
@@ -2423,7 +2423,7 @@ Inductive acontent : Type :=
Definition eq_acontent : forall (c1 c2: acontent), {c1=c2} + {c1<>c2}.
Proof.
- intros. generalize chunk_eq eq_aval. decide equality.
+ intros. generalize chunk_eq eq_aval. decide equality.
Defined.
Record ablock : Type := ABlock {
@@ -2494,7 +2494,7 @@ Qed.
Remark fst_inval_before: forall hi lo c, fst (inval_before hi lo c) = fst c.
Proof.
intros. functional induction (inval_before hi lo c); auto.
- rewrite IHt. unfold inval_if. destruct c##lo; auto.
+ rewrite IHt. unfold inval_if. destruct c##lo; auto.
destruct (zle (lo + size_chunk chunk) hi); auto.
Qed.
@@ -2507,7 +2507,7 @@ Program Definition ablock_store (chunk: memory_chunk) (ab: ablock) (i: Z) (av: a
vplub av ab.(ab_summary);
ab_default := _ |}.
Next Obligation.
- rewrite fst_inval_before, fst_inval_after. apply ab_default.
+ rewrite fst_inval_before, fst_inval_after. apply ab_default.
Qed.
Definition ablock_store_anywhere (chunk: memory_chunk) (ab: ablock) (av: aval) : ablock :=
@@ -2539,7 +2539,7 @@ Remark loadbytes_load_ext:
forall chunk ofs v, Mem.load chunk m' b ofs = Some v -> Mem.load chunk m b ofs = Some v.
Proof.
intros. exploit Mem.load_loadbytes; eauto. intros [bytes [A B]].
- exploit Mem.load_valid_access; eauto. intros [C D].
+ exploit Mem.load_valid_access; eauto. intros [C D].
subst v. apply Mem.loadbytes_load; auto. apply H; auto. generalize (size_chunk_pos chunk); omega.
Qed.
@@ -2567,7 +2567,7 @@ Qed.
Lemma smatch_ge:
forall m b p q, smatch m b p -> pge q p -> smatch m b q.
Proof.
- intros. destruct H as [A B]. split; intros.
+ intros. destruct H as [A B]. split; intros.
apply vmatch_ge with (Ifptr p); eauto with va.
apply pmatch_ge with p; eauto with va.
Qed.
@@ -2578,26 +2578,26 @@ Lemma In_loadbytes:
In byte bytes ->
exists ofs', ofs <= ofs' < ofs + n /\ Mem.loadbytes m b ofs' 1 = Some(byte :: nil).
Proof.
- intros until n. pattern n.
+ intros until n. pattern n.
apply well_founded_ind with (R := Zwf 0).
- apply Zwf_well_founded.
- intros sz REC ofs bytes LOAD IN.
- destruct (zle sz 0).
+ destruct (zle sz 0).
+ rewrite (Mem.loadbytes_empty m b ofs sz) in LOAD by auto.
inv LOAD. contradiction.
+ exploit (Mem.loadbytes_split m b ofs 1 (sz - 1) bytes).
replace (1 + (sz - 1)) with sz by omega. auto.
omega.
omega.
- intros (bytes1 & bytes2 & LOAD1 & LOAD2 & CONCAT).
- subst bytes.
+ intros (bytes1 & bytes2 & LOAD1 & LOAD2 & CONCAT).
+ subst bytes.
exploit Mem.loadbytes_length. eexact LOAD1. change (nat_of_Z 1) with 1%nat. intros LENGTH1.
rewrite in_app_iff in IN. destruct IN.
- * destruct bytes1; try discriminate. destruct bytes1; try discriminate.
+ * destruct bytes1; try discriminate. destruct bytes1; try discriminate.
simpl in H. destruct H; try contradiction. subst m0.
exists ofs; split. omega. auto.
* exploit (REC (sz - 1)). red; omega. eexact LOAD2. auto.
- intros (ofs' & A & B).
+ intros (ofs' & A & B).
exists ofs'; split. omega. auto.
Qed.
@@ -2609,7 +2609,7 @@ Lemma smatch_loadbytes:
pmatch b' ofs' p.
Proof.
intros. exploit In_loadbytes; eauto. intros (ofs1 & A & B).
- eapply H0; eauto.
+ eapply H0; eauto.
Qed.
Lemma loadbytes_provenance:
@@ -2619,7 +2619,7 @@ Lemma loadbytes_provenance:
ofs <= ofs' < ofs + n ->
In byte bytes.
Proof.
- intros until n. pattern n.
+ intros until n. pattern n.
apply well_founded_ind with (R := Zwf 0).
- apply Zwf_well_founded.
- intros sz REC ofs bytes LOAD LOAD1 IN.
@@ -2640,15 +2640,15 @@ Lemma storebytes_provenance:
In (Fragment (Vptr b'' ofs'') q i) bytes
\/ Mem.loadbytes m b' ofs' 1 = Some (Fragment (Vptr b'' ofs'') q i :: nil).
Proof.
- intros.
+ intros.
assert (EITHER:
(b' <> b \/ ofs' + 1 <= ofs \/ ofs + Z.of_nat (length bytes) <= ofs')
\/ (b' = b /\ ofs <= ofs' < ofs + Z.of_nat (length bytes))).
{
destruct (eq_block b' b); auto.
- destruct (zle (ofs' + 1) ofs); auto.
- destruct (zle (ofs + Z.of_nat (length bytes)) ofs'); auto.
- right. split. auto. omega.
+ destruct (zle (ofs' + 1) ofs); auto.
+ destruct (zle (ofs + Z.of_nat (length bytes)) ofs'); auto.
+ right. split. auto. omega.
}
destruct EITHER as [A | (A & B)].
- right. rewrite <- H0. symmetry. eapply Mem.loadbytes_storebytes_other; eauto. omega.
@@ -2664,15 +2664,15 @@ Lemma store_provenance:
v = Vptr b'' ofs'' /\ (chunk = Mint32 \/ chunk = Many32 \/ chunk = Many64)
\/ Mem.loadbytes m b' ofs' 1 = Some (Fragment (Vptr b'' ofs'') q i :: nil).
Proof.
- intros. exploit storebytes_provenance; eauto. eapply Mem.store_storebytes; eauto.
+ intros. exploit storebytes_provenance; eauto. eapply Mem.store_storebytes; eauto.
intros [A|A]; auto. left.
generalize (encode_val_shape chunk v). intros ENC; inv ENC.
- split; auto. rewrite <- H1 in A; destruct A.
+ congruence.
+ exploit H5; eauto. intros (j & P & Q); congruence.
-- rewrite <- H1 in A; destruct A.
+- rewrite <- H1 in A; destruct A.
+ congruence.
- + exploit H3; eauto. intros [byte P]; congruence.
+ + exploit H3; eauto. intros [byte P]; congruence.
- rewrite <- H1 in A; destruct A.
+ congruence.
+ exploit H2; eauto. congruence.
@@ -2687,18 +2687,18 @@ Lemma smatch_store:
Proof.
intros. destruct H0 as [A B]. split.
- intros chunk' ofs' v' LOAD. destruct v'; auto with va.
- exploit Mem.load_pointer_store; eauto.
- intros [(P & Q & R & S) | DISJ].
+ exploit Mem.load_pointer_store; eauto.
+ intros [(P & Q & R & S) | DISJ].
+ subst. apply vmatch_vplub_l. auto.
+ apply vmatch_vplub_r. apply A with (chunk := chunk') (ofs := ofs').
rewrite <- LOAD. symmetry. eapply Mem.load_store_other; eauto.
-- intros. exploit store_provenance; eauto. intros [[P Q] | P].
-+ subst.
+- intros. exploit store_provenance; eauto. intros [[P Q] | P].
++ subst.
assert (V: vmatch (Vptr b'0 ofs') (Ifptr (vplub av p))).
{
- apply vmatch_vplub_l. auto.
+ apply vmatch_vplub_l. auto.
}
- inv V; auto.
+ inv V; auto.
+ apply pmatch_vplub. eapply B; eauto.
Qed.
@@ -2710,22 +2710,22 @@ Lemma smatch_storebytes:
smatch m' b' (plub p' p).
Proof.
intros. destruct H0 as [A B]. split.
-- intros. apply vmatch_ifptr. intros bx ofsx EQ; subst v.
+- intros. apply vmatch_ifptr. intros bx ofsx EQ; subst v.
exploit Mem.load_loadbytes; eauto. intros (bytes' & P & Q).
- destruct bytes' as [ | byte1' bytes'].
+ destruct bytes' as [ | byte1' bytes'].
exploit Mem.loadbytes_length; eauto. intros. destruct chunk; discriminate.
- generalize (decode_val_shape chunk byte1' bytes'). rewrite <- Q.
+ generalize (decode_val_shape chunk byte1' bytes'). rewrite <- Q.
intros DEC; inv DEC; try contradiction.
assert (v = Vptr bx ofsx).
{ destruct H5 as [E|[E|E]]; rewrite E in H4; destruct v; simpl in H4; congruence. }
- exploit In_loadbytes; eauto. eauto with coqlib.
- intros (ofs' & X & Y). subst v.
+ exploit In_loadbytes; eauto. eauto with coqlib.
+ intros (ofs' & X & Y). subst v.
exploit storebytes_provenance; eauto. intros [Z | Z].
- apply pmatch_lub_l. eauto.
- apply pmatch_lub_r. eauto.
+ apply pmatch_lub_l. eauto.
+ apply pmatch_lub_r. eauto.
- intros. exploit storebytes_provenance; eauto. intros [Z | Z].
- apply pmatch_lub_l. eauto.
- apply pmatch_lub_r. eauto.
+ apply pmatch_lub_l. eauto.
+ apply pmatch_lub_r. eauto.
Qed.
Definition bmatch (m: mem) (b: block) (ab: ablock) : Prop :=
@@ -2740,7 +2740,7 @@ Lemma bmatch_ext:
Proof.
intros. destruct H as [A B]. split; intros.
apply smatch_ext with m; auto.
- eapply B; eauto. eapply loadbytes_load_ext; eauto.
+ eapply B; eauto. eapply loadbytes_load_ext; eauto.
Qed.
Lemma bmatch_inv:
@@ -2759,7 +2759,7 @@ Lemma ablock_load_sound:
bmatch m b ab ->
vmatch v (ablock_load chunk ab ofs).
Proof.
- intros. destruct H0. eauto.
+ intros. destruct H0. eauto.
Qed.
Lemma ablock_load_anywhere_sound:
@@ -2768,16 +2768,16 @@ Lemma ablock_load_anywhere_sound:
bmatch m b ab ->
vmatch v (ablock_load_anywhere chunk ab).
Proof.
- intros. destruct H0. destruct H0. unfold ablock_load_anywhere.
- eapply vnormalize_cast; eauto.
+ intros. destruct H0. destruct H0. unfold ablock_load_anywhere.
+ eapply vnormalize_cast; eauto.
Qed.
Lemma ablock_init_sound:
forall m b p, smatch m b p -> bmatch m b (ablock_init p).
Proof.
- intros; split; auto; intros.
+ intros; split; auto; intros.
unfold ablock_load, ablock_init; simpl. rewrite ZMap.gi.
- eapply vnormalize_cast; eauto. eapply H; eauto.
+ eapply vnormalize_cast; eauto. eapply H; eauto.
Qed.
Lemma ablock_store_anywhere_sound:
@@ -2788,14 +2788,14 @@ Lemma ablock_store_anywhere_sound:
bmatch m' b' (ablock_store_anywhere chunk ab av).
Proof.
intros. destruct H0 as [A B]. unfold ablock_store_anywhere.
- apply ablock_init_sound. eapply smatch_store; eauto.
+ apply ablock_init_sound. eapply smatch_store; eauto.
Qed.
Remark inval_after_outside:
forall i lo hi c, i < lo \/ i > hi -> (inval_after lo hi c)##i = c##i.
Proof.
intros until c. functional induction (inval_after lo hi c); intros.
- rewrite IHt by omega. apply ZMap.gso. unfold ZIndexed.t; omega.
+ rewrite IHt by omega. apply ZMap.gso. unfold ZIndexed.t; omega.
auto.
Qed.
@@ -2805,7 +2805,7 @@ Remark inval_after_contents:
c##i = ACval chunk av /\ (i < lo \/ i > hi).
Proof.
intros until c. functional induction (inval_after lo hi c); intros.
- destruct (zeq i hi).
+ destruct (zeq i hi).
subst i. rewrite inval_after_outside in H by omega. rewrite ZMap.gss in H. discriminate.
exploit IHt; eauto. intros [A B]. rewrite ZMap.gso in A by auto. split. auto. omega.
split. auto. omega.
@@ -2815,9 +2815,9 @@ Remark inval_before_outside:
forall i hi lo c, i < lo \/ i >= hi -> (inval_before hi lo c)##i = c##i.
Proof.
intros until c. functional induction (inval_before hi lo c); intros.
- rewrite IHt by omega. unfold inval_if. destruct (c##lo); auto.
+ rewrite IHt by omega. unfold inval_if. destruct (c##lo); auto.
destruct (zle (lo + size_chunk chunk) hi); auto.
- apply ZMap.gso. unfold ZIndexed.t; omega.
+ apply ZMap.gso. unfold ZIndexed.t; omega.
auto.
Qed.
@@ -2828,15 +2828,15 @@ Remark inval_before_contents_1:
Proof.
intros until c. functional induction (inval_before hi lo c); intros.
- destruct (zeq lo i).
-+ subst i. rewrite inval_before_outside in H0 by omega.
- unfold inval_if in H0. destruct (c##lo) eqn:C. congruence.
++ subst i. rewrite inval_before_outside in H0 by omega.
+ unfold inval_if in H0. destruct (c##lo) eqn:C. congruence.
destruct (zle (lo + size_chunk chunk0) hi).
- rewrite C in H0; inv H0. auto.
+ rewrite C in H0; inv H0. auto.
rewrite ZMap.gss in H0. congruence.
-+ exploit IHt. omega. auto. intros [A B]; split; auto.
++ exploit IHt. omega. auto. intros [A B]; split; auto.
unfold inval_if in A. destruct (c##lo) eqn:C. auto.
destruct (zle (lo + size_chunk chunk0) hi); auto.
- rewrite ZMap.gso in A; auto.
+ rewrite ZMap.gso in A; auto.
- omegaContradiction.
Qed.
@@ -2850,12 +2850,12 @@ Remark inval_before_contents:
(inval_before i (i - 7) c)##j = ACval chunk' av' ->
c##j = ACval chunk' av' /\ (j + size_chunk chunk' <= i \/ i <= j).
Proof.
- intros. destruct (zlt j (i - 7)).
- rewrite inval_before_outside in H by omega.
+ intros. destruct (zlt j (i - 7)).
+ rewrite inval_before_outside in H by omega.
split. auto. left. generalize (max_size_chunk chunk'); omega.
- destruct (zlt j i).
+ destruct (zlt j i).
exploit inval_before_contents_1; eauto. omega. tauto.
- rewrite inval_before_outside in H by omega.
+ rewrite inval_before_outside in H by omega.
split. auto. omega.
Qed.
@@ -2863,12 +2863,12 @@ Lemma ablock_store_contents:
forall chunk ab i av j chunk' av',
(ablock_store chunk ab i av).(ab_contents)##j = ACval chunk' av' ->
(i = j /\ chunk' = chunk /\ av' = av)
- \/ (ab.(ab_contents)##j = ACval chunk' av'
+ \/ (ab.(ab_contents)##j = ACval chunk' av'
/\ (j + size_chunk chunk' <= i \/ i + size_chunk chunk <= j)).
Proof.
unfold ablock_store; simpl; intros.
- destruct (zeq i j).
- subst j. rewrite ZMap.gss in H. inv H; auto.
+ destruct (zeq i j).
+ subst j. rewrite ZMap.gss in H. inv H; auto.
right. rewrite ZMap.gso in H by auto.
exploit inval_before_contents; eauto. intros [A B].
exploit inval_after_contents; eauto. intros [C D].
@@ -2891,20 +2891,20 @@ Lemma ablock_store_sound:
bmatch m' b (ablock_store chunk ab ofs av).
Proof.
intros until av; intros STORE BIN VIN. destruct BIN as [BIN1 BIN2]. split.
- eapply smatch_store; eauto.
+ eapply smatch_store; eauto.
intros chunk' ofs' v' LOAD.
assert (SUMMARY: vmatch v' (vnormalize chunk' (Ifptr (vplub av ab.(ab_summary))))).
{ exploit smatch_store; eauto. intros [A B]. eapply vnormalize_cast; eauto. }
- unfold ablock_load.
+ unfold ablock_load.
destruct ((ab_contents (ablock_store chunk ab ofs av)) ## ofs') as [ | chunk1 av1] eqn:C.
apply SUMMARY.
destruct (chunk_compat chunk' chunk1) eqn:COMPAT; auto.
exploit chunk_compat_true; eauto. intros (U & V & W).
exploit ablock_store_contents; eauto. intros [(P & Q & R) | (P & Q)].
- (* same offset and compatible chunks *)
- subst.
- assert (v' = Val.load_result chunk' v).
- { exploit Mem.load_store_similar_2; eauto. congruence. }
+ subst.
+ assert (v' = Val.load_result chunk' v).
+ { exploit Mem.load_store_similar_2; eauto. congruence. }
subst v'. apply vnormalize_sound; auto.
- (* disjoint load/store *)
assert (Mem.load chunk' m b ofs' = Some v').
@@ -2920,7 +2920,7 @@ Lemma ablock_loadbytes_sound:
In (Fragment (Vptr b' ofs') q i) bytes ->
pmatch b' ofs' (ablock_loadbytes ab).
Proof.
- intros. destruct H0. eapply smatch_loadbytes; eauto.
+ intros. destruct H0. eapply smatch_loadbytes; eauto.
Qed.
Lemma ablock_storebytes_anywhere_sound:
@@ -2937,7 +2937,7 @@ Qed.
Lemma ablock_storebytes_contents:
forall ab p i sz j chunk' av',
(ablock_storebytes ab p i sz).(ab_contents)##j = ACval chunk' av' ->
- ab.(ab_contents)##j = ACval chunk' av'
+ ab.(ab_contents)##j = ACval chunk' av'
/\ (j + size_chunk chunk' <= i \/ i + Zmax sz 0 <= j).
Proof.
unfold ablock_storebytes; simpl; intros.
@@ -2954,15 +2954,15 @@ Lemma ablock_storebytes_sound:
bmatch m b ab ->
bmatch m' b (ablock_storebytes ab p ofs sz).
Proof.
- intros until sz; intros STORE LENGTH CONTENTS BM. destruct BM as [BM1 BM2]. split.
- eapply smatch_storebytes; eauto.
+ intros until sz; intros STORE LENGTH CONTENTS BM. destruct BM as [BM1 BM2]. split.
+ eapply smatch_storebytes; eauto.
intros chunk' ofs' v' LOAD'.
assert (SUMMARY: vmatch v' (vnormalize chunk' (Ifptr (plub p ab.(ab_summary))))).
{ exploit smatch_storebytes; eauto. intros [A B]. eapply vnormalize_cast; eauto. }
- unfold ablock_load.
+ unfold ablock_load.
destruct (ab_contents (ablock_storebytes ab p ofs sz))##ofs' eqn:C.
exact SUMMARY.
- destruct (chunk_compat chunk' chunk) eqn:COMPAT; auto.
+ destruct (chunk_compat chunk' chunk) eqn:COMPAT; auto.
exploit chunk_compat_true; eauto. intros (U & V & W).
exploit ablock_storebytes_contents; eauto. intros [A B].
assert (Mem.load chunk' m b ofs' = Some v').
@@ -2975,7 +2975,7 @@ Qed.
Definition bbeq (ab1 ab2: ablock) : bool :=
eq_aptr ab1.(ab_summary) ab2.(ab_summary) &&
- PTree.beq (fun c1 c2 => proj_sumbool (eq_acontent c1 c2))
+ PTree.beq (fun c1 c2 => proj_sumbool (eq_acontent c1 c2))
(snd ab1.(ab_contents)) (snd ab2.(ab_contents)).
Lemma bbeq_load:
@@ -2989,7 +2989,7 @@ Proof.
- rewrite PTree.beq_correct in H1.
assert (A: forall i, ZMap.get i (ab_contents ab1) = ZMap.get i (ab_contents ab2)).
{
- intros. unfold ZMap.get, PMap.get. set (j := ZIndexed.index i).
+ intros. unfold ZMap.get, PMap.get. set (j := ZIndexed.index i).
specialize (H1 j).
destruct (snd (ab_contents ab1))!j; destruct (snd (ab_contents ab2))!j; try contradiction.
InvBooleans; auto.
@@ -3004,9 +3004,9 @@ Lemma bbeq_sound:
bbeq ab1 ab2 = true ->
forall m b, bmatch m b ab1 <-> bmatch m b ab2.
Proof.
- intros. exploit bbeq_load; eauto. intros [A B].
- unfold bmatch. rewrite A. intuition. rewrite <- B; eauto. rewrite B; eauto.
-Qed.
+ intros. exploit bbeq_load; eauto. intros [A B].
+ unfold bmatch. rewrite A. intuition. rewrite <- B; eauto. rewrite B; eauto.
+Qed.
(** Least upper bound *)
@@ -3039,11 +3039,11 @@ Lemma get_combine_contentmaps:
ZMap.get i (combine_contentmaps m1 m2) = combine_acontents (ZMap.get i m1) (ZMap.get i m2).
Proof.
intros. destruct m1 as [dfl1 pt1]. destruct m2 as [dfl2 pt2]; simpl in *.
- subst dfl1 dfl2. unfold combine_contentmaps, ZMap.get, PMap.get, fst, snd.
- set (j := ZIndexed.index i).
+ subst dfl1 dfl2. unfold combine_contentmaps, ZMap.get, PMap.get, fst, snd.
+ set (j := ZIndexed.index i).
rewrite PTree.gcombine by auto.
destruct (pt1!j) as [[]|]; destruct (pt2!j) as [[]|]; simpl; auto.
- destruct (chunk_eq chunk chunk0); auto.
+ destruct (chunk_eq chunk chunk0); auto.
Qed.
Lemma smatch_lub_l:
@@ -3051,7 +3051,7 @@ Lemma smatch_lub_l:
Proof.
intros. destruct H as [A B]. split; intros.
change (vmatch v (vlub (Ifptr p) (Ifptr q))). apply vmatch_lub_l. eapply A; eauto.
- apply pmatch_lub_l. eapply B; eauto.
+ apply pmatch_lub_l. eapply B; eauto.
Qed.
Lemma smatch_lub_r:
@@ -3059,14 +3059,14 @@ Lemma smatch_lub_r:
Proof.
intros. destruct H as [A B]. split; intros.
change (vmatch v (vlub (Ifptr p) (Ifptr q))). apply vmatch_lub_r. eapply A; eauto.
- apply pmatch_lub_r. eapply B; eauto.
+ apply pmatch_lub_r. eapply B; eauto.
Qed.
Lemma bmatch_lub_l:
forall m b x y, bmatch m b x -> bmatch m b (blub x y).
Proof.
intros. destruct H as [BM1 BM2]. split; unfold blub; simpl.
-- apply smatch_lub_l; auto.
+- apply smatch_lub_l; auto.
- intros.
assert (SUMMARY: vmatch v (vnormalize chunk (Ifptr (plub (ab_summary x) (ab_summary y))))
).
@@ -3077,14 +3077,14 @@ Proof.
unfold combine_acontents; destruct (ab_contents x)##ofs, (ab_contents y)##ofs; auto.
destruct (chunk_eq chunk0 chunk1); auto. subst chunk0.
destruct (chunk_compat chunk chunk1); auto.
- intros. eapply vmatch_ge; eauto. apply vnormalize_monotone. apply vge_lub_l.
+ intros. eapply vmatch_ge; eauto. apply vnormalize_monotone. apply vge_lub_l.
Qed.
Lemma bmatch_lub_r:
forall m b x y, bmatch m b y -> bmatch m b (blub x y).
Proof.
intros. destruct H as [BM1 BM2]. split; unfold blub; simpl.
-- apply smatch_lub_r; auto.
+- apply smatch_lub_r; auto.
- intros.
assert (SUMMARY: vmatch v (vnormalize chunk (Ifptr (plub (ab_summary x) (ab_summary y))))
).
@@ -3095,7 +3095,7 @@ Proof.
unfold combine_acontents; destruct (ab_contents x)##ofs, (ab_contents y)##ofs; auto.
destruct (chunk_eq chunk0 chunk1); auto. subst chunk0.
destruct (chunk_compat chunk chunk1); auto.
- intros. eapply vmatch_ge; eauto. apply vnormalize_monotone. apply vge_lub_r.
+ intros. eapply vmatch_ge; eauto. apply vnormalize_monotone. apply vge_lub_r.
Qed.
(** * Abstracting read-only global variables *)
@@ -3119,7 +3119,7 @@ Proof.
intros; red; intros. exploit H0; eauto. intros (A & B & C). split; auto. split.
- exploit Mem.store_valid_access_3; eauto. intros [P _].
apply bmatch_inv with m; auto.
-+ intros. eapply Mem.loadbytes_store_other; eauto.
++ intros. eapply Mem.loadbytes_store_other; eauto.
left. red; intros; subst b0. elim (C ofs). apply Mem.perm_cur_max.
apply P. generalize (size_chunk_pos chunk); omega.
- intros; red; intros; elim (C ofs0). eauto with mem.
@@ -3133,7 +3133,7 @@ Lemma romatch_storebytes:
Proof.
intros; red; intros. exploit H0; eauto. intros (A & B & C). split; auto. split.
- apply bmatch_inv with m; auto.
- intros. eapply Mem.loadbytes_storebytes_disjoint; eauto.
+ intros. eapply Mem.loadbytes_storebytes_disjoint; eauto.
destruct (eq_block b0 b); auto. subst b0. right; red; unfold Intv.In; simpl; red; intros.
elim (C x). apply Mem.perm_cur_max. eapply Mem.storebytes_range_perm; eauto.
- intros; red; intros; elim (C ofs0). eauto with mem.
@@ -3149,7 +3149,7 @@ Proof.
intros; red; intros. exploit H; eauto. intros (A & B & C).
split. auto.
split. apply bmatch_ext with m; auto. intros. eapply H0; eauto.
- intros; red; intros. elim (C ofs). eapply H1; eauto.
+ intros; red; intros. elim (C ofs). eapply H1; eauto.
Qed.
Lemma romatch_free:
@@ -3158,8 +3158,8 @@ Lemma romatch_free:
romatch m rm ->
romatch m' rm.
Proof.
- intros. apply romatch_ext with m; auto.
- intros. eapply Mem.loadbytes_free_2; eauto.
+ intros. apply romatch_ext with m; auto.
+ intros. eapply Mem.loadbytes_free_2; eauto.
intros. eauto with mem.
Qed.
@@ -3170,7 +3170,7 @@ Lemma romatch_alloc:
romatch m rm ->
romatch m' rm.
Proof.
- intros. apply romatch_ext with m; auto.
+ intros. apply romatch_ext with m; auto.
intros. rewrite <- H3; symmetry. eapply Mem.loadbytes_alloc_unchanged; eauto.
apply H0. congruence.
intros. eapply Mem.perm_alloc_4; eauto. apply Mem.valid_not_valid_diff with m; eauto with mem.
@@ -3191,7 +3191,7 @@ Record mmatch (m: mem) (am: amem) : Prop := mk_mem_match {
bc b = BCstack ->
bmatch m b am.(am_stack);
mmatch_glob: forall id ab b,
- bc b = BCglob id ->
+ bc b = BCglob id ->
am.(am_glob)!id = Some ab ->
bmatch m b ab;
mmatch_nonstack: forall b,
@@ -3323,18 +3323,18 @@ Theorem load_sound:
pmatch b ofs p ->
vmatch v (load chunk rm am p).
Proof.
- intros. unfold load. inv H2.
+ intros. unfold load. inv H2.
- (* Gl id ofs *)
- destruct (rm!id) as [ab|] eqn:RM.
- eapply ablock_load_sound; eauto. eapply H0; eauto.
+ destruct (rm!id) as [ab|] eqn:RM.
+ eapply ablock_load_sound; eauto. eapply H0; eauto.
destruct (am_glob am)!id as [ab|] eqn:AM.
- eapply ablock_load_sound; eauto. eapply mmatch_glob; eauto.
+ eapply ablock_load_sound; eauto. eapply mmatch_glob; eauto.
eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto; congruence.
- (* Glo id *)
- destruct (rm!id) as [ab|] eqn:RM.
- eapply ablock_load_anywhere_sound; eauto. eapply H0; eauto.
+ destruct (rm!id) as [ab|] eqn:RM.
+ eapply ablock_load_anywhere_sound; eauto. eapply H0; eauto.
destruct (am_glob am)!id as [ab|] eqn:AM.
- eapply ablock_load_anywhere_sound; eauto. eapply mmatch_glob; eauto.
+ eapply ablock_load_anywhere_sound; eauto. eapply mmatch_glob; eauto.
eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto; congruence.
- (* Glob *)
eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto. congruence. congruence.
@@ -3343,7 +3343,7 @@ Proof.
- (* Stack *)
eapply ablock_load_anywhere_sound; eauto. eapply mmatch_stack; eauto.
- (* Nonstack *)
- eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto.
+ eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto.
- (* Top *)
eapply vnormalize_cast; eauto. eapply mmatch_top; eauto.
Qed.
@@ -3357,7 +3357,7 @@ Theorem loadv_sound:
vmatch v (loadv chunk rm am aaddr).
Proof.
intros. destruct addr; simpl in H; try discriminate.
- eapply load_sound; eauto. apply match_aptr_of_aval; auto.
+ eapply load_sound; eauto. apply match_aptr_of_aval; auto.
Qed.
Theorem store_sound:
@@ -3372,27 +3372,27 @@ Proof.
unfold store; constructor; simpl; intros.
- (* Stack *)
assert (DFL: bc b <> BCstack -> bmatch m' b0 (am_stack am)).
- { intros. apply bmatch_inv with m. eapply mmatch_stack; eauto.
+ { intros. apply bmatch_inv with m. eapply mmatch_stack; eauto.
intros. eapply Mem.loadbytes_store_other; eauto. left; congruence. }
inv PM; try (apply DFL; congruence).
- + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
- eapply ablock_store_sound; eauto. eapply mmatch_stack; eauto.
- + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
- eapply ablock_store_anywhere_sound; eauto. eapply mmatch_stack; eauto.
+ + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
+ eapply ablock_store_sound; eauto. eapply mmatch_stack; eauto.
+ + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
+ eapply ablock_store_anywhere_sound; eauto. eapply mmatch_stack; eauto.
+ eapply ablock_store_anywhere_sound; eauto. eapply mmatch_stack; eauto.
- (* Globals *)
rename b0 into b'.
- assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab ->
+ assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab ->
bmatch m' b' ab).
{ intros. apply bmatch_inv with m. eapply mmatch_glob; eauto.
intros. eapply Mem.loadbytes_store_other; eauto. left; congruence. }
- inv PM.
+ inv PM.
+ rewrite PTree.gsspec in H0. destruct (peq id id0).
subst id0; inv H0.
assert (b' = b) by (eapply bc_glob; eauto). subst b'.
eapply ablock_store_sound; eauto.
- destruct (am_glob am)!id as [ab0|] eqn:GL.
+ destruct (am_glob am)!id as [ab0|] eqn:GL.
eapply mmatch_glob; eauto.
apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence.
eapply DFL; eauto. congruence.
@@ -3400,13 +3400,13 @@ Proof.
subst id0; inv H0.
assert (b' = b) by (eapply bc_glob; eauto). subst b'.
eapply ablock_store_anywhere_sound; eauto.
- destruct (am_glob am)!id as [ab0|] eqn:GL.
+ destruct (am_glob am)!id as [ab0|] eqn:GL.
eapply mmatch_glob; eauto.
apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence.
eapply DFL; eauto. congruence.
+ rewrite PTree.gempty in H0; congruence.
+ eapply DFL; eauto. congruence.
- + eapply DFL; eauto. congruence.
+ + eapply DFL; eauto. congruence.
+ rewrite PTree.gempty in H0; congruence.
+ rewrite PTree.gempty in H0; congruence.
@@ -3434,7 +3434,7 @@ Theorem storev_sound:
vmatch v av ->
mmatch m' (storev chunk am aaddr av).
Proof.
- intros. destruct addr; simpl in H; try discriminate.
+ intros. destruct addr; simpl in H; try discriminate.
eapply store_sound; eauto. apply match_aptr_of_aval; auto.
Qed.
@@ -3451,21 +3451,21 @@ Proof.
destruct (rm!id) as [ab|] eqn:RM.
exploit H0; eauto. intros (A & B & C). eapply ablock_loadbytes_sound; eauto.
destruct (am_glob am)!id as [ab|] eqn:GL.
- eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto.
+ eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto.
eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va.
- (* Glo id *)
destruct (rm!id) as [ab|] eqn:RM.
exploit H0; eauto. intros (A & B & C). eapply ablock_loadbytes_sound; eauto.
destruct (am_glob am)!id as [ab|] eqn:GL.
- eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto.
+ eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto.
eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va.
- (* Glob *)
eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va.
- (* Stk ofs *)
- eapply ablock_loadbytes_sound; eauto. eapply mmatch_stack; eauto.
+ eapply ablock_loadbytes_sound; eauto. eapply mmatch_stack; eauto.
- (* Stack *)
eapply ablock_loadbytes_sound; eauto. eapply mmatch_stack; eauto.
-- (* Nonstack *)
+- (* Nonstack *)
eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va.
- (* Top *)
eapply smatch_loadbytes; eauto. eapply mmatch_top; eauto with va.
@@ -3484,27 +3484,27 @@ Proof.
unfold storebytes; constructor; simpl; intros.
- (* Stack *)
assert (DFL: bc b <> BCstack -> bmatch m' b0 (am_stack am)).
- { intros. apply bmatch_inv with m. eapply mmatch_stack; eauto.
+ { intros. apply bmatch_inv with m. eapply mmatch_stack; eauto.
intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence. }
inv PM; try (apply DFL; congruence).
- + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
- eapply ablock_storebytes_sound; eauto. eapply mmatch_stack; eauto.
- + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
- eapply ablock_storebytes_anywhere_sound; eauto. eapply mmatch_stack; eauto.
+ + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
+ eapply ablock_storebytes_sound; eauto. eapply mmatch_stack; eauto.
+ + assert (b0 = b) by (eapply bc_stack; eauto). subst b0.
+ eapply ablock_storebytes_anywhere_sound; eauto. eapply mmatch_stack; eauto.
+ eapply ablock_storebytes_anywhere_sound; eauto. eapply mmatch_stack; eauto.
- (* Globals *)
rename b0 into b'.
- assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab ->
+ assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab ->
bmatch m' b' ab).
{ intros. apply bmatch_inv with m. eapply mmatch_glob; eauto.
intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence. }
- inv PM.
+ inv PM.
+ rewrite PTree.gsspec in H0. destruct (peq id id0).
subst id0; inv H0.
assert (b' = b) by (eapply bc_glob; eauto). subst b'.
eapply ablock_storebytes_sound; eauto.
- destruct (am_glob am)!id as [ab0|] eqn:GL.
+ destruct (am_glob am)!id as [ab0|] eqn:GL.
eapply mmatch_glob; eauto.
apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence.
eapply DFL; eauto. congruence.
@@ -3512,13 +3512,13 @@ Proof.
subst id0; inv H0.
assert (b' = b) by (eapply bc_glob; eauto). subst b'.
eapply ablock_storebytes_anywhere_sound; eauto.
- destruct (am_glob am)!id as [ab0|] eqn:GL.
+ destruct (am_glob am)!id as [ab0|] eqn:GL.
eapply mmatch_glob; eauto.
apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence.
eapply DFL; eauto. congruence.
+ rewrite PTree.gempty in H0; congruence.
+ eapply DFL; eauto. congruence.
- + eapply DFL; eauto. congruence.
+ + eapply DFL; eauto. congruence.
+ rewrite PTree.gempty in H0; congruence.
+ rewrite PTree.gempty in H0; congruence.
@@ -3550,7 +3550,7 @@ Proof.
- apply bmatch_ext with m; eauto with va.
- apply smatch_ext with m; auto with va.
- apply smatch_ext with m; auto with va.
-- red; intros. exploit mmatch_below0; eauto. xomega.
+- red; intros. exploit mmatch_below0; eauto. xomega.
Qed.
Lemma mmatch_free:
@@ -3559,16 +3559,16 @@ Lemma mmatch_free:
mmatch m am ->
mmatch m' am.
Proof.
- intros. apply mmatch_ext with m; auto.
- intros. eapply Mem.loadbytes_free_2; eauto.
- erewrite <- Mem.nextblock_free by eauto. xomega.
+ intros. apply mmatch_ext with m; auto.
+ intros. eapply Mem.loadbytes_free_2; eauto.
+ erewrite <- Mem.nextblock_free by eauto. xomega.
Qed.
Lemma mmatch_top':
forall m am, mmatch m am -> mmatch m mtop.
Proof.
intros. constructor; simpl; intros.
-- apply ablock_init_sound. apply smatch_ge with (ab_summary (am_stack am)).
+- apply ablock_init_sound. apply smatch_ge with (ab_summary (am_stack am)).
eapply mmatch_stack; eauto. constructor.
- rewrite PTree.gempty in H1; discriminate.
- eapply smatch_ge. eapply mmatch_nonstack; eauto. constructor.
@@ -3589,16 +3589,16 @@ Lemma mbeq_sound:
Proof.
unfold mbeq; intros. InvBooleans. rewrite PTree.beq_correct in H1.
split; intros M; inv M; constructor; intros.
-- erewrite <- bbeq_sound; eauto.
-- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m1)!id eqn:G; try contradiction.
- erewrite <- bbeq_sound; eauto.
-- rewrite <- H; eauto.
+- erewrite <- bbeq_sound; eauto.
+- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m1)!id eqn:G; try contradiction.
+ erewrite <- bbeq_sound; eauto.
+- rewrite <- H; eauto.
- rewrite <- H0; eauto.
- auto.
-- erewrite bbeq_sound; eauto.
-- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m2)!id eqn:G; try contradiction.
- erewrite bbeq_sound; eauto.
-- rewrite H; eauto.
+- erewrite bbeq_sound; eauto.
+- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m2)!id eqn:G; try contradiction.
+ erewrite bbeq_sound; eauto.
+- rewrite H; eauto.
- rewrite H0; eauto.
- auto.
Qed.
@@ -3620,14 +3620,14 @@ Definition mlub (m1 m2: amem) : amem :=
Lemma mmatch_lub_l:
forall m x y, mmatch m x -> mmatch m (mlub x y).
Proof.
- intros. inv H. constructor; simpl; intros.
-- apply bmatch_lub_l; auto.
-- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0.
+ intros. inv H. constructor; simpl; intros.
+- apply bmatch_lub_l; auto.
+- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0.
destruct (am_glob x)!id as [b1|] eqn:G1;
destruct (am_glob y)!id as [b2|] eqn:G2;
inv H0.
- apply bmatch_lub_l; eauto.
-- apply smatch_lub_l; auto.
+ apply bmatch_lub_l; eauto.
+- apply smatch_lub_l; auto.
- apply smatch_lub_l; auto.
- auto.
Qed.
@@ -3635,14 +3635,14 @@ Qed.
Lemma mmatch_lub_r:
forall m x y, mmatch m y -> mmatch m (mlub x y).
Proof.
- intros. inv H. constructor; simpl; intros.
-- apply bmatch_lub_r; auto.
-- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0.
+ intros. inv H. constructor; simpl; intros.
+- apply bmatch_lub_r; auto.
+- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0.
destruct (am_glob x)!id as [b1|] eqn:G1;
destruct (am_glob y)!id as [b2|] eqn:G2;
inv H0.
- apply bmatch_lub_r; eauto.
-- apply smatch_lub_r; auto.
+ apply bmatch_lub_r; eauto.
+- apply smatch_lub_r; auto.
- apply smatch_lub_r; auto.
- auto.
Qed.
@@ -3658,9 +3658,9 @@ Lemma genv_match_exten:
(forall b, bc1 b = BCother -> bc2 b = BCother) ->
genv_match bc2 ge.
Proof.
- intros. destruct H as [A B]. split; intros.
-- rewrite <- H0. eauto.
-- exploit B; eauto. destruct (bc1 b) eqn:BC1.
+ intros. destruct H as [A B]. split; intros.
+- rewrite <- H0. eauto.
+- exploit B; eauto. destruct (bc1 b) eqn:BC1.
+ intuition congruence.
+ rewrite H0 in BC1. intuition congruence.
+ intuition congruence.
@@ -3678,19 +3678,19 @@ Proof.
assert (PM: forall b ofs p, pmatch bc1 b ofs p -> pmatch bc1 b ofs (ab_summary ab) -> pmatch bc2 b ofs p).
{
intros.
- assert (pmatch bc1 b0 ofs Glob) by (eapply pmatch_ge; eauto).
+ assert (pmatch bc1 b0 ofs Glob) by (eapply pmatch_ge; eauto).
inv H5.
- assert (bc2 b0 = BCglob id0) by (rewrite H0; auto).
+ assert (bc2 b0 = BCglob id0) by (rewrite H0; auto).
inv H3; econstructor; eauto with va.
}
assert (VM: forall v x, vmatch bc1 v x -> vmatch bc1 v (Ifptr (ab_summary ab)) -> vmatch bc2 v x).
{
- intros. inv H3; constructor; auto; inv H4; eapply PM; eauto.
+ intros. inv H3; constructor; auto; inv H4; eapply PM; eauto.
}
destruct B as [[B1 B2] B3]. split. split.
-- intros. apply VM; eauto.
-- intros. apply PM; eauto.
-- intros. apply VM; eauto.
+- intros. apply VM; eauto.
+- intros. apply PM; eauto.
+- intros. apply VM; eauto.
Qed.
Definition bc_incr (bc1 bc2: block_classification) : Prop :=
@@ -3703,28 +3703,28 @@ Hypothesis INCR: bc_incr bc1 bc2.
Lemma pmatch_incr: forall b ofs p, pmatch bc1 b ofs p -> pmatch bc2 b ofs p.
Proof.
- induction 1;
+ induction 1;
assert (bc2 b = bc1 b) by (apply INCR; congruence);
- econstructor; eauto with va. rewrite H0; eauto.
+ econstructor; eauto with va. rewrite H0; eauto.
Qed.
Lemma vmatch_incr: forall v x, vmatch bc1 v x -> vmatch bc2 v x.
Proof.
- induction 1; constructor; auto; apply pmatch_incr; auto.
+ induction 1; constructor; auto; apply pmatch_incr; auto.
Qed.
Lemma smatch_incr: forall m b p, smatch bc1 m b p -> smatch bc2 m b p.
Proof.
- intros. destruct H as [A B]. split; intros.
- apply vmatch_incr; eauto.
+ intros. destruct H as [A B]. split; intros.
+ apply vmatch_incr; eauto.
apply pmatch_incr; eauto.
Qed.
Lemma bmatch_incr: forall m b ab, bmatch bc1 m b ab -> bmatch bc2 m b ab.
Proof.
- intros. destruct H as [B1 B2]. split.
- apply smatch_incr; auto.
- intros. apply vmatch_incr; eauto.
+ intros. destruct H as [B1 B2]. split.
+ apply smatch_incr; auto.
+ intros. apply vmatch_incr; eauto.
Qed.
End MATCH_INCR.
@@ -3737,7 +3737,7 @@ Definition inj_of_bc (bc: block_classification) : meminj :=
Lemma inj_of_bc_valid:
forall (bc: block_classification) b, bc b <> BCinvalid -> inj_of_bc bc b = Some(b, 0).
Proof.
- intros. unfold inj_of_bc. destruct (bc b); congruence.
+ intros. unfold inj_of_bc. destruct (bc b); congruence.
Qed.
Lemma inj_of_bc_inv:
@@ -3750,44 +3750,44 @@ Qed.
Lemma pmatch_inj:
forall bc b ofs p, pmatch bc b ofs p -> inj_of_bc bc b = Some(b, 0).
Proof.
- intros. apply inj_of_bc_valid. inv H; congruence.
+ intros. apply inj_of_bc_valid. inv H; congruence.
Qed.
Lemma vmatch_inj:
forall bc v x, vmatch bc v x -> Val.inject (inj_of_bc bc) v v.
Proof.
- induction 1; econstructor.
- eapply pmatch_inj; eauto. rewrite Int.add_zero; auto.
- eapply pmatch_inj; eauto. rewrite Int.add_zero; auto.
+ induction 1; econstructor.
+ eapply pmatch_inj; eauto. rewrite Int.add_zero; auto.
+ eapply pmatch_inj; eauto. rewrite Int.add_zero; auto.
Qed.
Lemma vmatch_list_inj:
forall bc vl xl, list_forall2 (vmatch bc) vl xl -> Val.inject_list (inj_of_bc bc) vl vl.
Proof.
- induction 1; constructor. eapply vmatch_inj; eauto. auto.
-Qed.
+ induction 1; constructor. eapply vmatch_inj; eauto. auto.
+Qed.
Lemma mmatch_inj:
forall bc m am, mmatch bc m am -> bc_below bc (Mem.nextblock m) -> Mem.inject (inj_of_bc bc) m m.
Proof.
intros. constructor. constructor.
- (* perms *)
- intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
- rewrite Zplus_0_r. auto.
+ intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
+ rewrite Zplus_0_r. auto.
- (* alignment *)
- intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
- apply Zdivide_0.
+ intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
+ apply Zdivide_0.
- (* contents *)
- intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
- rewrite Zplus_0_r.
+ intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
+ rewrite Zplus_0_r.
set (mv := ZMap.get ofs (Mem.mem_contents m)#b1).
assert (Mem.loadbytes m b1 ofs 1 = Some (mv :: nil)).
{
Local Transparent Mem.loadbytes.
- unfold Mem.loadbytes. rewrite pred_dec_true. reflexivity.
+ unfold Mem.loadbytes. rewrite pred_dec_true. reflexivity.
red; intros. replace ofs0 with ofs by omega. auto.
}
- destruct mv; econstructor. destruct v; econstructor.
+ destruct mv; econstructor. destruct v; econstructor.
apply inj_of_bc_valid.
assert (PM: pmatch bc b i Ptop).
{ exploit mmatch_top; eauto. intros [P Q].
@@ -3795,17 +3795,17 @@ Proof.
inv PM; auto.
rewrite Int.add_zero; auto.
- (* free blocks *)
- intros. unfold inj_of_bc. erewrite bc_below_invalid; eauto.
+ intros. unfold inj_of_bc. erewrite bc_below_invalid; eauto.
- (* mapped blocks *)
intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
apply H0; auto.
- (* overlap *)
- red; intros.
+ red; intros.
exploit inj_of_bc_inv. eexact H2. intros (A1 & B & C); subst.
exploit inj_of_bc_inv. eexact H3. intros (A2 & B & C); subst.
auto.
- (* overflow *)
- intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
+ intros. exploit inj_of_bc_inv; eauto. intros (A & B & C); subst.
rewrite Zplus_0_r. split. omega. apply Int.unsigned_range_2.
Qed.
@@ -3814,8 +3814,8 @@ Lemma inj_of_bc_preserves_globals:
Proof.
intros. destruct H as [A B].
split. intros. apply inj_of_bc_valid. rewrite A in H. congruence.
- split. intros. apply inj_of_bc_valid. apply B. eapply Genv.genv_vars_range; eauto.
- intros. exploit inj_of_bc_inv; eauto. intros (P & Q & R). auto.
+ split. intros. apply inj_of_bc_valid. apply B. eapply Genv.genv_vars_range; eauto.
+ intros. exploit inj_of_bc_inv; eauto. intros (P & Q & R). auto.
Qed.
Lemma pmatch_inj_top:
@@ -3827,27 +3827,27 @@ Qed.
Lemma vmatch_inj_top:
forall bc v v', Val.inject (inj_of_bc bc) v v' -> vmatch bc v Vtop.
Proof.
- intros. inv H; constructor. eapply pmatch_inj_top; eauto.
+ intros. inv H; constructor. eapply pmatch_inj_top; eauto.
Qed.
Lemma mmatch_inj_top:
forall bc m m', Mem.inject (inj_of_bc bc) m m' -> mmatch bc m mtop.
Proof.
- intros.
+ intros.
assert (SM: forall b, bc b <> BCinvalid -> smatch bc m b Ptop).
{
- intros; split; intros.
- - exploit Mem.load_inject. eauto. eauto. apply inj_of_bc_valid; auto.
- intros (v' & A & B). eapply vmatch_inj_top; eauto.
- - exploit Mem.loadbytes_inject. eauto. eauto. apply inj_of_bc_valid; auto.
- intros (bytes' & A & B). inv B. inv H4. inv H8. eapply pmatch_inj_top; eauto.
+ intros; split; intros.
+ - exploit Mem.load_inject. eauto. eauto. apply inj_of_bc_valid; auto.
+ intros (v' & A & B). eapply vmatch_inj_top; eauto.
+ - exploit Mem.loadbytes_inject. eauto. eauto. apply inj_of_bc_valid; auto.
+ intros (bytes' & A & B). inv B. inv H4. inv H8. eapply pmatch_inj_top; eauto.
}
- constructor; simpl; intros.
+ constructor; simpl; intros.
- apply ablock_init_sound. apply SM. congruence.
- rewrite PTree.gempty in H1; discriminate.
- apply SM; auto.
- apply SM; auto.
- - red; intros. eapply Mem.valid_block_inject_1. eapply inj_of_bc_valid; eauto. eauto.
+ - red; intros. eapply Mem.valid_block_inject_1. eapply inj_of_bc_valid; eauto. eauto.
Qed.
(** * Abstracting RTL register environments *)
@@ -3882,30 +3882,30 @@ End AVal.
Module AE := LPMap(AVal).
-Definition aenv := AE.t.
+Definition aenv := AE.t.
Section MATCHENV.
Variable bc: block_classification.
Definition ematch (e: regset) (ae: aenv) : Prop :=
- forall r, vmatch bc e#r (AE.get r ae).
+ forall r, vmatch bc e#r (AE.get r ae).
Lemma ematch_ge:
forall e ae1 ae2,
ematch e ae1 -> AE.ge ae2 ae1 -> ematch e ae2.
Proof.
- intros; red; intros. apply vmatch_ge with (AE.get r ae1); auto. apply H0.
+ intros; red; intros. apply vmatch_ge with (AE.get r ae1); auto. apply H0.
Qed.
Lemma ematch_update:
forall e ae v av r,
ematch e ae -> vmatch bc v av -> ematch (e#r <- v) (AE.set r av ae).
Proof.
- intros; red; intros. rewrite AE.gsspec. rewrite PMap.gsspec.
- destruct (peq r0 r); auto.
- red; intros. specialize (H xH). subst ae. simpl in H. inv H.
- unfold AVal.eq; red; intros. subst av. inv H0.
+ intros; red; intros. rewrite AE.gsspec. rewrite PMap.gsspec.
+ destruct (peq r0 r); auto.
+ red; intros. specialize (H xH). subst ae. simpl in H. inv H.
+ unfold AVal.eq; red; intros. subst av. inv H0.
Qed.
Fixpoint einit_regs (rl: list reg) : aenv :=
@@ -3919,23 +3919,23 @@ Lemma ematch_init:
(forall v, In v vl -> vmatch bc v (Ifptr Nonstack)) ->
ematch (init_regs vl rl) (einit_regs rl).
Proof.
- induction rl; simpl; intros.
-- red; intros. rewrite Regmap.gi. simpl AE.get. rewrite PTree.gempty.
- constructor.
-- destruct vl as [ | v1 vs ].
+ induction rl; simpl; intros.
+- red; intros. rewrite Regmap.gi. simpl AE.get. rewrite PTree.gempty.
+ constructor.
+- destruct vl as [ | v1 vs ].
+ assert (ematch (init_regs nil rl) (einit_regs rl)).
{ apply IHrl. simpl; tauto. }
- replace (init_regs nil rl) with (Regmap.init Vundef) in H0 by (destruct rl; auto).
- red; intros. rewrite AE.gsspec. destruct (peq r a).
- rewrite Regmap.gi. constructor.
- apply H0.
+ replace (init_regs nil rl) with (Regmap.init Vundef) in H0 by (destruct rl; auto).
+ red; intros. rewrite AE.gsspec. destruct (peq r a).
+ rewrite Regmap.gi. constructor.
+ apply H0.
red; intros EQ; rewrite EQ in H0. specialize (H0 xH). simpl in H0. inv H0.
unfold AVal.eq, AVal.bot. congruence.
+ assert (ematch (init_regs vs rl) (einit_regs rl)).
{ apply IHrl. eauto with coqlib. }
- red; intros. rewrite Regmap.gsspec. rewrite AE.gsspec. destruct (peq r a).
+ red; intros. rewrite Regmap.gsspec. rewrite AE.gsspec. destruct (peq r a).
auto with coqlib.
- apply H0.
+ apply H0.
red; intros EQ; rewrite EQ in H0. specialize (H0 xH). simpl in H0. inv H0.
unfold AVal.eq, AVal.bot. congruence.
Qed.
@@ -3954,14 +3954,14 @@ Proof.
destruct ae. unfold AE.get at 2. apply AVal.ge_bot.
eapply AVal.ge_trans. apply IHrl. rewrite AE.gsspec.
destruct (peq p a). apply AVal.ge_top. apply AVal.ge_refl. apply AVal.eq_refl.
- congruence.
+ congruence.
unfold AVal.eq, Vtop, AVal.bot. congruence.
Qed.
Lemma ematch_forget:
forall e rl ae, ematch e ae -> ematch e (eforget rl ae).
Proof.
- intros. eapply ematch_ge; eauto. apply eforget_ge.
+ intros. eapply ematch_ge; eauto. apply eforget_ge.
Qed.
End MATCHENV.
@@ -3969,7 +3969,7 @@ End MATCHENV.
Lemma ematch_incr:
forall bc bc' e ae, ematch bc e ae -> bc_incr bc bc' -> ematch bc' e ae.
Proof.
- intros; red; intros. apply vmatch_incr with bc; auto.
+ intros; red; intros. apply vmatch_incr with bc; auto.
Qed.
(** * Lattice for dataflow analysis *)
@@ -3989,12 +3989,12 @@ Module VA <: SEMILATTICE.
Lemma eq_refl: forall x, eq x x.
Proof.
- destruct x; simpl. auto. split. apply AE.eq_refl. tauto.
+ destruct x; simpl. auto. split. apply AE.eq_refl. tauto.
Qed.
Lemma eq_sym: forall x y, eq x y -> eq y x.
Proof.
- destruct x, y; simpl; auto. intros [A B].
- split. apply AE.eq_sym; auto. intros. rewrite B. tauto.
+ destruct x, y; simpl; auto. intros [A B].
+ split. apply AE.eq_sym; auto. intros. rewrite B. tauto.
Qed.
Lemma eq_trans: forall x y z, eq x y -> eq y z -> eq x z.
Proof.
@@ -4009,16 +4009,16 @@ Module VA <: SEMILATTICE.
| State ae1 am1, State ae2 am2 => AE.beq ae1 ae2 && mbeq am1 am2
| _, _ => false
end.
-
+
Lemma beq_correct: forall x y, beq x y = true -> eq x y.
Proof.
- destruct x, y; simpl; intros.
+ destruct x, y; simpl; intros.
auto.
congruence.
congruence.
InvBooleans; split.
apply AE.beq_correct; auto.
- intros. apply mbeq_sound; auto.
+ intros. apply mbeq_sound; auto.
Qed.
Definition ge (x y: t) : Prop :=
@@ -4030,21 +4030,21 @@ Module VA <: SEMILATTICE.
Lemma ge_refl: forall x y, eq x y -> ge x y.
Proof.
- destruct x, y; simpl; try tauto. intros [A B]; split.
+ destruct x, y; simpl; try tauto. intros [A B]; split.
apply AE.ge_refl; auto.
- intros. rewrite B; auto.
+ intros. rewrite B; auto.
Qed.
Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z.
Proof.
destruct x, y, z; simpl; try tauto. intros [A B] [C D]; split.
- eapply AE.ge_trans; eauto.
- eauto.
+ eapply AE.ge_trans; eauto.
+ eauto.
Qed.
Definition bot : t := Bot.
Lemma ge_bot: forall x, ge x bot.
Proof.
- destruct x; simpl; auto.
+ destruct x; simpl; auto.
Qed.
Definition lub (x y: t) : t :=
@@ -4078,7 +4078,7 @@ Hint Constructors pmatch: va.
Hint Constructors vmatch: va.
Hint Resolve cnot_sound symbol_address_sound
shl_sound shru_sound shr_sound
- and_sound or_sound xor_sound notint_sound
+ and_sound or_sound xor_sound notint_sound
ror_sound rolm_sound
neg_sound add_sound sub_sound
mul_sound mulhs_sound mulhu_sound
@@ -4090,6 +4090,6 @@ Hint Resolve cnot_sound symbol_address_sound
zero_ext_sound sign_ext_sound singleoffloat_sound floatofsingle_sound
intoffloat_sound intuoffloat_sound floatofint_sound floatofintu_sound
intofsingle_sound intuofsingle_sound singleofint_sound singleofintu_sound
- longofwords_sound loword_sound hiword_sound
+ longofwords_sound loword_sound hiword_sound
cmpu_bool_sound cmp_bool_sound cmpf_bool_sound cmpfs_bool_sound
maskzero_sound : va.
diff --git a/backend/XTL.ml b/backend/XTL.ml
index dde9bdb0..2ddbc50a 100644
--- a/backend/XTL.ml
+++ b/backend/XTL.ml
@@ -186,7 +186,7 @@ let type_block blk =
let type_function f =
PTree.fold
(fun () pc blk ->
- try
+ try
type_block blk
with Type_error ->
raise (Type_error_at pc))