aboutsummaryrefslogtreecommitdiffstats
path: root/backend
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2006-09-04 15:08:29 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2006-09-04 15:08:29 +0000
commit73729d23ac13275c0d28d23bc1b1f6056104e5d9 (patch)
treee3044ce75edb30377bd8c87356b7617eba5ed223 /backend
parentc79434827bf2bd71f857f4471f7bbf381fddd189 (diff)
downloadcompcert-kvx-73729d23ac13275c0d28d23bc1b1f6056104e5d9.tar.gz
compcert-kvx-73729d23ac13275c0d28d23bc1b1f6056104e5d9.zip
Fusion de la branche "traces":
- Ajout de traces d'evenements d'E/S dans les semantiques - Ajout constructions switch et allocation dynamique - Initialisation des variables globales - Portage Coq 8.1 beta Debut d'integration du front-end C: - Traduction Clight -> Csharpminor dans cfrontend/ - Modifications de Csharpminor et Globalenvs en consequence. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@72 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend')
-rw-r--r--backend/AST.v96
-rw-r--r--backend/Allocation.v41
-rw-r--r--backend/Allocproof.v336
-rw-r--r--backend/Allocproof_aux.v850
-rw-r--r--backend/Alloctyping.v74
-rw-r--r--backend/Alloctyping_aux.v895
-rw-r--r--backend/CSE.v15
-rw-r--r--backend/CSEproof.v134
-rw-r--r--backend/Cmconstr.v51
-rw-r--r--backend/Cmconstrproof.v727
-rw-r--r--backend/Cminor.v207
-rw-r--r--backend/Cminorgen.v74
-rw-r--r--backend/Cminorgenproof.v1209
-rw-r--r--backend/Coloring.v6
-rw-r--r--backend/Coloringproof.v49
-rw-r--r--backend/Constprop.v24
-rw-r--r--backend/Constpropproof.v112
-rw-r--r--backend/Conventions.v115
-rw-r--r--backend/Csharpminor.v260
-rw-r--r--backend/Events.v103
-rw-r--r--backend/Globalenvs.v94
-rw-r--r--backend/InterfGraph.v137
-rw-r--r--backend/Kildall.v280
-rw-r--r--backend/LTL.v152
-rw-r--r--backend/LTLtyping.v14
-rw-r--r--backend/Linear.v97
-rw-r--r--backend/Linearize.v7
-rw-r--r--backend/Linearizeproof.v138
-rw-r--r--backend/Linearizetyping.v16
-rw-r--r--backend/Lineartyping.v13
-rw-r--r--backend/Mach.v101
-rw-r--r--backend/Machabstr.v274
-rw-r--r--backend/Machabstr2mach.v100
-rw-r--r--backend/Machtyping.v107
-rw-r--r--backend/Main.v58
-rw-r--r--backend/Mem.v132
-rw-r--r--backend/Op.v21
-rw-r--r--backend/PPC.v99
-rw-r--r--backend/PPCgen.v13
-rw-r--r--backend/PPCgenproof.v286
-rw-r--r--backend/PPCgenproof1.v112
-rw-r--r--backend/Parallelmove.v2759
-rw-r--r--backend/RTL.v123
-rw-r--r--backend/RTLgen.v23
-rw-r--r--backend/RTLgenproof.v503
-rw-r--r--backend/RTLgenproof1.v51
-rw-r--r--backend/RTLtyping.v1230
-rw-r--r--backend/Stacking.v7
-rw-r--r--backend/Stackingproof.v185
-rw-r--r--backend/Stackingtyping.v20
-rw-r--r--backend/Tunneling.v7
-rw-r--r--backend/Tunnelingproof.v136
-rw-r--r--backend/Tunnelingtyping.v11
53 files changed, 4666 insertions, 8018 deletions
diff --git a/backend/AST.v b/backend/AST.v
index aae9e860..1342bef1 100644
--- a/backend/AST.v
+++ b/backend/AST.v
@@ -2,6 +2,8 @@
the abstract syntax trees of many of the intermediate languages. *)
Require Import Coqlib.
+Require Import Integers.
+Require Import Floats.
Set Implicit Arguments.
@@ -33,6 +35,12 @@ Record signature : Set := mksignature {
sig_res: option typ
}.
+Definition proj_sig_res (s: signature) : typ :=
+ match s.(sig_res) with
+ | None => Tint
+ | Some t => t
+ end.
+
(** Memory accesses (load and store instructions) are annotated by
a ``memory chunk'' indicating the type, size and signedness of the
chunk of memory being accessed. *)
@@ -46,41 +54,20 @@ Inductive memory_chunk : Set :=
| Mfloat32 : memory_chunk (**r 32-bit single-precision float *)
| Mfloat64 : memory_chunk. (**r 64-bit double-precision float *)
-(** Comparison instructions can perform one of the six following comparisons
- between their two operands. *)
-
-Inductive comparison : Set :=
- | Ceq : comparison (**r same *)
- | Cne : comparison (**r different *)
- | Clt : comparison (**r less than *)
- | Cle : comparison (**r less than or equal *)
- | Cgt : comparison (**r greater than *)
- | Cge : comparison. (**r greater than or equal *)
-
-Definition negate_comparison (c: comparison): comparison :=
- match c with
- | Ceq => Cne
- | Cne => Ceq
- | Clt => Cge
- | Cle => Cgt
- | Cgt => Cle
- | Cge => Clt
- end.
+(** Initialization data for global variables. *)
-Definition swap_comparison (c: comparison): comparison :=
- match c with
- | Ceq => Ceq
- | Cne => Cne
- | Clt => Cgt
- | Cle => Cge
- | Cgt => Clt
- | Cge => Cle
- end.
+Inductive init_data: Set :=
+ | Init_int8: int -> init_data
+ | Init_int16: int -> init_data
+ | Init_int32: int -> init_data
+ | Init_float32: float -> init_data
+ | Init_float64: float -> init_data
+ | Init_space: Z -> init_data.
(** Whole programs consist of:
- a collection of function definitions (name and description);
- the name of the ``main'' function that serves as entry point in the program;
-- a collection of global variable declarations (name and size in bytes).
+- a collection of global variable declarations (name and initializer).
The type of function descriptions varies among the various intermediate
languages and is taken as parameter to the [program] type. The other parts
@@ -89,7 +76,7 @@ of whole programs are common to all languages. *)
Record program (funct: Set) : Set := mkprogram {
prog_funct: list (ident * funct);
prog_main: ident;
- prog_vars: list (ident * Z)
+ prog_vars: list (ident * list init_data)
}.
(** We now define a general iterator over programs that applies a given
@@ -214,3 +201,50 @@ Proof.
Qed.
End TRANSF_PARTIAL_PROGRAM.
+
+(** For most languages, the functions composing the program are either
+ internal functions, defined within the language, or external functions
+ (a.k.a. system calls) that emit an event when applied. We define
+ a type for such functions and some generic transformation functions. *)
+
+Record external_function : Set := mkextfun {
+ ef_id: ident;
+ ef_sig: signature
+}.
+
+Inductive fundef (F: Set): Set :=
+ | Internal: F -> fundef F
+ | External: external_function -> fundef F.
+
+Implicit Arguments External [F].
+
+Section TRANSF_FUNDEF.
+
+Variable A B: Set.
+Variable transf: A -> B.
+
+Definition transf_fundef (fd: fundef A): fundef B :=
+ match fd with
+ | Internal f => Internal (transf f)
+ | External ef => External ef
+ end.
+
+End TRANSF_FUNDEF.
+
+Section TRANSF_PARTIAL_FUNDEF.
+
+Variable A B: Set.
+Variable transf_partial: A -> option B.
+
+Definition transf_partial_fundef (fd: fundef A): option (fundef B) :=
+ match fd with
+ | Internal f =>
+ match transf_partial f with
+ | None => None
+ | Some f' => Some (Internal f')
+ end
+ | External ef => Some (External ef)
+ end.
+
+End TRANSF_PARTIAL_FUNDEF.
+
diff --git a/backend/Allocation.v b/backend/Allocation.v
index 30f9dcc6..d919d1eb 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -87,6 +87,8 @@ Definition transfer
| Icall sig ros args res s =>
reg_list_live args
(reg_sum_live ros (reg_dead res after))
+ | Ialloc arg res s =>
+ reg_live arg (reg_dead res after)
| Icond cond args ifso ifnot =>
reg_list_live args after
| Ireturn optarg =>
@@ -214,20 +216,10 @@ Definition add_move (src dst: loc) (k: block) :=
length. This is a parallel move, meaning that arbitrary overlap
between the sources and destinations is permitted. *)
-Fixpoint listsLoc2Moves (src dst : list loc) {struct src} : Moves :=
- match src, dst with
- | nil, _ => nil
- | s :: srcs, nil => nil
- | s :: srcs, d :: dsts => (s, d) :: listsLoc2Moves srcs dsts
- end.
-
-Definition parallel_move_order (src dst : list loc) :=
- Parallelmove.P_move (listsLoc2Moves src dst).
-
Definition parallel_move (srcs dsts: list loc) (k: block) : block :=
- List.fold_left
- (fun k p => add_move (fst p) (snd p) k)
- (parallel_move_order srcs dsts) k.
+ List.fold_right
+ (fun p k => add_move (fst p) (snd p) k)
+ k (parmove srcs dsts).
(** ** Constructors for LTL instructions *)
@@ -261,6 +253,10 @@ Definition add_store (chunk: memory_chunk) (addr: addressing)
(Bstore chunk addr rargs rsrc (Bgoto s))
end.
+Definition add_alloc (arg: loc) (res: loc) (s: node) :=
+ add_reload arg Conventions.loc_alloc_argument
+ (Balloc (add_spill Conventions.loc_alloc_result res (Bgoto s))).
+
(** For function calls, we also add appropriate moves to and from
the canonical locations for function arguments and function results,
as dictated by the calling conventions. *)
@@ -344,10 +340,12 @@ Definition transf_instr
| Icall sig ros args res s =>
add_call sig (sum_left_map assign ros) (List.map assign args)
(assign res) s
+ | Ialloc arg res s =>
+ add_alloc (assign arg) (assign res) s
| Icond cond args ifso ifnot =>
add_cond cond (List.map assign args) ifso ifnot
| Ireturn optarg =>
- add_return (RTL.fn_sig f) (option_map assign optarg)
+ add_return f.(RTL.fn_sig) (option_map assign optarg)
end.
Definition transf_entrypoint
@@ -391,7 +389,7 @@ Qed.
transformation as described above. *)
Definition transf_function (f: RTL.function) : option LTL.function :=
- match type_rtl_function f with
+ match type_function f with
| None => None
| Some env =>
match analyze f with
@@ -413,6 +411,17 @@ Definition transf_function (f: RTL.function) : option LTL.function :=
end
end.
+Definition transf_fundef (fd: RTL.fundef) : option LTL.fundef :=
+ match fd with
+ | External ef =>
+ if type_external_function ef then Some (External ef) else None
+ | Internal f =>
+ match transf_function f with
+ | None => None
+ | Some tf => Some (Internal tf)
+ end
+ end.
+
Definition transf_program (p: RTL.program) : option LTL.program :=
- transform_partial_program transf_function p.
+ transform_partial_program transf_fundef p.
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index 138e6d79..07c0f58b 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -8,6 +8,7 @@ Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Registers.
@@ -17,8 +18,8 @@ Require Import Locations.
Require Import Conventions.
Require Import Coloring.
Require Import Coloringproof.
+Require Import Parallelmove.
Require Import Allocation.
-Require Import Allocproof_aux.
(** * Semantic properties of calling conventions *)
@@ -473,7 +474,7 @@ Qed.
Lemma add_reload_correct:
forall src dst k rs m,
exists rs',
- exec_instrs ge sp (add_reload src dst k) rs m k rs' m /\
+ exec_instrs ge sp (add_reload src dst k) rs m E0 k rs' m /\
rs' (R dst) = rs src /\
forall l, Loc.diff (R dst) l -> rs' l = rs l.
Proof.
@@ -493,7 +494,7 @@ Qed.
Lemma add_spill_correct:
forall src dst k rs m,
exists rs',
- exec_instrs ge sp (add_spill src dst k) rs m k rs' m /\
+ exec_instrs ge sp (add_spill src dst k) rs m E0 k rs' m /\
rs' dst = rs (R src) /\
forall l, Loc.diff dst l -> rs' l = rs l.
Proof.
@@ -520,7 +521,7 @@ Lemma add_reloads_correct_rec:
list_norepet itmps ->
list_norepet ftmps ->
exists rs',
- exec_instrs ge sp (add_reloads srcs (regs_for_rec srcs itmps ftmps) k) rs m k rs' m /\
+ exec_instrs ge sp (add_reloads srcs (regs_for_rec srcs itmps ftmps) k) rs m E0 k rs' m /\
reglist (regs_for_rec srcs itmps ftmps) rs' = map rs srcs /\
(forall r, ~(In r itmps) -> ~(In r ftmps) -> rs' (R r) = rs (R r)) /\
(forall s, rs' (S s) = rs (S s)).
@@ -578,7 +579,7 @@ Proof.
generalize (IHsrcs itmps ftmps k rs1 m R1 R2 R3 R4 R5 R6 R7).
intros [rs' [EX [RES [OTH1 OTH2]]]].
exists rs'.
- split. eapply exec_trans; eauto.
+ split. eapply exec_trans; eauto. traceEq.
split. simpl. apply (f_equal2 (@cons val)).
rewrite OTH1; auto.
rewrite RES. apply list_map_exten. intros.
@@ -599,7 +600,7 @@ Lemma add_reloads_correct:
(List.length srcs <= 3)%nat ->
Loc.disjoint srcs temporaries ->
exists rs',
- exec_instrs ge sp (add_reloads srcs (regs_for srcs) k) rs m k rs' m /\
+ exec_instrs ge sp (add_reloads srcs (regs_for srcs) k) rs m E0 k rs' m /\
reglist (regs_for srcs) rs' = List.map rs srcs /\
forall l, Loc.notin l temporaries -> rs' l = rs l.
Proof.
@@ -638,7 +639,7 @@ Qed.
Lemma add_move_correct:
forall src dst k rs m,
exists rs',
- exec_instrs ge sp (add_move src dst k) rs m k rs' m /\
+ exec_instrs ge sp (add_move src dst k) rs m E0 k rs' m /\
rs' dst = rs src /\
forall l, Loc.diff l dst -> Loc.diff l (R IT1) -> Loc.diff l (R FT1) -> rs' l = rs l.
Proof.
@@ -660,13 +661,33 @@ Proof.
generalize (add_spill_correct tmp (S s0) k rs1 m);
intros [rs2 [EX2 [RES2 OTH2]]].
exists rs2. split.
- eapply exec_trans; eauto.
+ eapply exec_trans; eauto. traceEq.
split. congruence.
intros. rewrite OTH2. apply OTH1.
apply Loc.diff_sym. unfold tmp; case (slot_type s); auto.
apply Loc.diff_sym; auto.
Qed.
+Lemma effect_move_sequence:
+ forall k moves rs m,
+ let k' := List.fold_right (fun p k => add_move (fst p) (snd p) k) k moves in
+ exists rs',
+ exec_instrs ge sp k' rs m E0 k rs' m /\
+ effect_seqmove moves rs rs'.
+Proof.
+ induction moves; intros until m; simpl.
+ exists rs; split. constructor. constructor.
+ destruct a as [src dst]; simpl.
+ set (k1 := fold_right
+ (fun (p : loc * loc) (k : block) => add_move (fst p) (snd p) k)
+ k moves) in *.
+ destruct (add_move_correct src dst k1 rs m) as [rs1 [A [B C]]].
+ destruct (IHmoves rs1 m) as [rs' [D E]].
+ exists rs'; split.
+ eapply exec_trans; eauto. traceEq.
+ econstructor; eauto. red. tauto.
+Qed.
+
Theorem parallel_move_correct:
forall srcs dsts k rs m,
List.length srcs = List.length dsts ->
@@ -675,13 +696,16 @@ Theorem parallel_move_correct:
Loc.disjoint srcs temporaries ->
Loc.disjoint dsts temporaries ->
exists rs',
- exec_instrs ge sp (parallel_move srcs dsts k) rs m k rs' m /\
+ exec_instrs ge sp (parallel_move srcs dsts k) rs m E0 k rs' m /\
List.map rs' dsts = List.map rs srcs /\
rs' (R IT3) = rs (R IT3) /\
forall l, Loc.notin l dsts -> Loc.notin l temporaries -> rs' l = rs l.
Proof.
- apply (parallel_move_correctX ge sp).
- apply add_move_correct.
+ intros.
+ generalize (effect_move_sequence k (parmove srcs dsts) rs m).
+ intros [rs' [EXEC EFFECT]].
+ exists rs'. split. exact EXEC.
+ apply effect_parmove; auto.
Qed.
Lemma add_op_correct:
@@ -690,7 +714,7 @@ Lemma add_op_correct:
Loc.disjoint args temporaries ->
eval_operation ge sp op (List.map rs args) = Some v ->
exists rs',
- exec_block ge sp (add_op op args res s) rs m (Cont s) rs' m /\
+ exec_block ge sp (add_op op args res s) rs m E0 (Cont s) rs' m /\
rs' res = v /\
forall l, Loc.diff l res -> Loc.notin l temporaries -> rs' l = rs l.
Proof.
@@ -721,7 +745,7 @@ Proof.
split. apply exec_Bgoto. eapply exec_trans. eexact EX1.
eapply exec_trans; eauto.
apply exec_one. unfold rs2. apply exec_Bop.
- unfold rargs. rewrite RES1. auto.
+ unfold rargs. rewrite RES1. auto. traceEq.
split. rewrite RES3. unfold rs2; apply Locmap.gss.
intros. rewrite OTHER3. unfold rs2. rewrite Locmap.gso.
apply OTHER1. assumption.
@@ -737,7 +761,7 @@ Lemma add_load_correct:
eval_addressing ge sp addr (List.map rs args) = Some a ->
loadv chunk m a = Some v ->
exists rs',
- exec_block ge sp (add_load chunk addr args res s) rs m (Cont s) rs' m /\
+ exec_block ge sp (add_load chunk addr args res s) rs m E0 (Cont s) rs' m /\
rs' res = v /\
forall l, Loc.diff l res -> Loc.notin l temporaries -> rs' l = rs l.
Proof.
@@ -755,7 +779,7 @@ Proof.
split. apply exec_Bgoto. eapply exec_trans; eauto.
eapply exec_trans; eauto.
apply exec_one. unfold rs2. apply exec_Bload with a.
- unfold rargs; rewrite RES1. assumption. assumption.
+ unfold rargs; rewrite RES1. assumption. assumption. traceEq.
split. rewrite RES3. unfold rs2; apply Locmap.gss.
intros. rewrite OTHER3. unfold rs2. rewrite Locmap.gso.
apply OTHER1. assumption.
@@ -772,7 +796,7 @@ Lemma add_store_correct:
eval_addressing ge sp addr (List.map rs args) = Some a ->
storev chunk m a (rs src) = Some m' ->
exists rs',
- exec_block ge sp (add_store chunk addr args src s) rs m (Cont s) rs' m' /\
+ exec_block ge sp (add_store chunk addr args src s) rs m E0 (Cont s) rs' m' /\
forall l, Loc.notin l temporaries -> rs' l = rs l.
Proof.
intros.
@@ -795,10 +819,42 @@ Proof.
split. apply exec_Bgoto.
eapply exec_trans. rewrite <- EQ. eexact EX1.
apply exec_one. apply exec_Bstore with a.
- rewrite RES2. assumption. rewrite RES3. assumption.
+ rewrite RES2. assumption. rewrite RES3. assumption. traceEq.
exact OTHER1.
Qed.
+Lemma add_alloc_correct:
+ forall arg res s rs m m' sz b,
+ rs arg = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', b) ->
+ exists rs',
+ exec_block ge sp (add_alloc arg res s) rs m E0 (Cont s) rs' m' /\
+ rs' res = Vptr b Int.zero /\
+ forall l,
+ Loc.diff l (R Conventions.loc_alloc_argument) ->
+ Loc.diff l (R Conventions.loc_alloc_result) ->
+ Loc.diff l res ->
+ rs' l = rs l.
+Proof.
+ intros; unfold add_alloc.
+ generalize (add_reload_correct arg loc_alloc_argument
+ (Balloc (add_spill loc_alloc_result res (Bgoto s)))
+ rs m).
+ intros [rs1 [EX1 [RES1 OTHER1]]].
+ pose (rs2 := Locmap.set (R loc_alloc_result) (Vptr b Int.zero) rs1).
+ generalize (add_spill_correct loc_alloc_result res (Bgoto s) rs2 m').
+ intros [rs3 [EX3 [RES3 OTHER3]]].
+ exists rs3.
+ split. apply exec_Bgoto. eapply exec_trans. eexact EX1.
+ eapply exec_trans. apply exec_one. eapply exec_Balloc; eauto. congruence.
+ fold rs2. eexact EX3. reflexivity. traceEq.
+ split. rewrite RES3; unfold rs2. apply Locmap.gss.
+ intros. rewrite OTHER3. unfold rs2. rewrite Locmap.gso.
+ apply OTHER1. apply Loc.diff_sym; auto.
+ apply Loc.diff_sym; auto.
+ apply Loc.diff_sym; auto.
+Qed.
+
Lemma add_cond_correct:
forall cond args ifso ifnot rs m b s,
(List.length args <= 3)%nat ->
@@ -806,7 +862,7 @@ Lemma add_cond_correct:
eval_condition cond (List.map rs args) = Some b ->
s = (if b then ifso else ifnot) ->
exists rs',
- exec_block ge sp (add_cond cond args ifso ifnot) rs m (Cont s) rs' m /\
+ exec_block ge sp (add_cond cond args ifso ifnot) rs m E0 (Cont s) rs' m /\
forall l, Loc.notin l temporaries -> rs' l = rs l.
Proof.
intros. unfold add_cond.
@@ -825,7 +881,7 @@ Proof.
exact OTHER1.
Qed.
-Definition find_function2 (los: loc + ident) (ls: locset) : option function :=
+Definition find_function2 (los: loc + ident) (ls: locset) : option fundef :=
match los with
| inl l => Genv.find_funct ge (ls l)
| inr symb =>
@@ -836,21 +892,21 @@ Definition find_function2 (los: loc + ident) (ls: locset) : option function :=
end.
Lemma add_call_correct:
- forall f vargs m vres m' sig los args res s ls
+ forall f vargs m t vres m' sig los args res s ls
(EXECF:
forall lsi,
- List.map lsi (loc_arguments f.(fn_sig)) = vargs ->
+ List.map lsi (loc_arguments (funsig f)) = vargs ->
exists lso,
- exec_function ge f lsi m lso m'
- /\ lso (R (loc_result f.(fn_sig))) = vres)
+ exec_function ge f lsi m t lso m'
+ /\ lso (R (loc_result (funsig f))) = vres)
(FIND: find_function2 los ls = Some f)
- (SIG: sig = f.(fn_sig))
+ (SIG: sig = funsig f)
(VARGS: List.map ls args = vargs)
(LARGS: List.length args = List.length sig.(sig_args))
(AARGS: locs_acceptable args)
(RES: loc_acceptable res),
exists ls',
- exec_block ge sp (add_call sig los args res s) ls m (Cont s) ls' m' /\
+ exec_block ge sp (add_call sig los args res s) ls m t (Cont s) ls' m' /\
ls' res = vres /\
forall l,
Loc.notin l destroyed_at_call -> loc_acceptable l -> Loc.diff l res ->
@@ -896,7 +952,7 @@ Proof.
eapply exec_trans. apply exec_one. apply exec_Bcall with f.
unfold find_function. rewrite TMP2. rewrite RES1.
assumption. assumption. eexact EX3.
- exact EX5.
+ eexact EX5. reflexivity. reflexivity. traceEq.
(* Result *)
split. rewrite RES5. unfold ls4. rewrite return_regs_result.
assumption.
@@ -936,7 +992,7 @@ Proof.
eapply exec_trans. eexact EX2.
eapply exec_trans. apply exec_one. apply exec_Bcall with f.
unfold find_function. assumption. assumption. eexact EX3.
- exact EX5.
+ eexact EX5. reflexivity. traceEq.
(* Result *)
split. rewrite RES5.
unfold ls4. rewrite return_regs_result.
@@ -955,7 +1011,7 @@ Lemma add_undefs_correct:
(forall l, In l res -> loc_acceptable l) ->
(forall ofs ty, In (S (Local ofs ty)) res -> ls (S (Local ofs ty)) = Vundef) ->
exists ls',
- exec_instrs ge sp (add_undefs res b) ls m b ls' m /\
+ exec_instrs ge sp (add_undefs res b) ls m E0 b ls' m /\
(forall l, In l res -> ls' l = Vundef) /\
(forall l, Loc.notin l res -> ls' l = ls l).
Proof.
@@ -972,7 +1028,7 @@ Proof.
intros [ls2 [EX2 [RES2 OTHER2]]].
exists ls2. split.
eapply exec_trans. apply exec_one. apply exec_Bop.
- simpl; reflexivity. exact EX2.
+ simpl; reflexivity. eexact EX2. traceEq.
split. intros. case (In_dec Loc.eq l res); intro.
apply RES2; auto.
rewrite OTHER2. elim H1; intro.
@@ -1006,7 +1062,7 @@ Lemma add_entry_correct:
locs_acceptable undefs ->
(forall ofs ty, ls (S (Local ofs ty)) = Vundef) ->
exists ls',
- exec_block ge sp (add_entry sig params undefs s) ls m (Cont s) ls' m /\
+ exec_block ge sp (add_entry sig params undefs s) ls m E0 (Cont s) ls' m /\
List.map ls' params = List.map ls (loc_parameters sig) /\
(forall l, In l undefs -> ls' l = Vundef).
Proof.
@@ -1029,7 +1085,7 @@ Proof.
intros [ls2 [EX2 [RES2 OTHER2]]].
exists ls2.
split. apply exec_Bgoto. unfold add_entry.
- eapply exec_trans. eexact EX1. eexact EX2.
+ eapply exec_trans. eexact EX1. eexact EX2. traceEq.
split. rewrite <- RES1. apply list_map_exten.
intros. symmetry. apply OTHER2. eapply Loc.disjoint_notin; eauto.
exact RES2.
@@ -1038,7 +1094,7 @@ Qed.
Lemma add_return_correct:
forall sig optarg ls m,
exists ls',
- exec_block ge sp (add_return sig optarg) ls m Return ls' m /\
+ exec_block ge sp (add_return sig optarg) ls m E0 Return ls' m /\
match optarg with
| Some arg => ls' (R (loc_result sig)) = ls arg
| None => ls' (R (loc_result sig)) = Vundef
@@ -1131,51 +1187,53 @@ Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
intro. unfold ge, tge.
- apply Genv.find_symbol_transf_partial with transf_function.
+ apply Genv.find_symbol_transf_partial with transf_fundef.
exact TRANSF.
Qed.
Lemma functions_translated:
- forall (v: val) (f: RTL.function),
+ forall (v: val) (f: RTL.fundef),
Genv.find_funct ge v = Some f ->
exists tf,
- Genv.find_funct tge v = Some tf /\ transf_function f = Some tf.
+ Genv.find_funct tge v = Some tf /\ transf_fundef f = Some tf.
Proof.
intros.
generalize
- (Genv.find_funct_transf_partial transf_function TRANSF H).
- case (transf_function f).
+ (Genv.find_funct_transf_partial transf_fundef TRANSF H).
+ case (transf_fundef f).
intros tf [A B]. exists tf. tauto.
intros [A B]. elim B. reflexivity.
Qed.
Lemma function_ptr_translated:
- forall (b: Values.block) (f: RTL.function),
+ forall (b: Values.block) (f: RTL.fundef),
Genv.find_funct_ptr ge b = Some f ->
exists tf,
- Genv.find_funct_ptr tge b = Some tf /\ transf_function f = Some tf.
+ Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = Some tf.
Proof.
intros.
generalize
- (Genv.find_funct_ptr_transf_partial transf_function TRANSF H).
- case (transf_function f).
+ (Genv.find_funct_ptr_transf_partial transf_fundef TRANSF H).
+ case (transf_fundef f).
intros tf [A B]. exists tf. tauto.
intros [A B]. elim B. reflexivity.
Qed.
Lemma sig_function_translated:
forall f tf,
- transf_function f = Some tf ->
- tf.(LTL.fn_sig) = f.(RTL.fn_sig).
+ transf_fundef f = Some tf ->
+ LTL.funsig tf = RTL.funsig f.
Proof.
- intros f tf. unfold transf_function.
- destruct (type_rtl_function f).
+ intros f tf. destruct f; simpl.
+ unfold transf_function.
+ destruct (type_function f).
destruct (analyze f).
- destruct (regalloc f t0).
+ destruct (regalloc f t).
intro EQ; injection EQ; intro EQ1; rewrite <- EQ1; simpl; auto.
- intros; discriminate.
- intros; discriminate.
- intros; discriminate.
+ congruence. congruence. congruence.
+ destruct (type_external_function e).
+ intro EQ; inversion EQ; subst tf. reflexivity.
+ congruence.
Qed.
Lemma entrypoint_function_translated:
@@ -1184,9 +1242,9 @@ Lemma entrypoint_function_translated:
tf.(LTL.fn_entrypoint) = f.(RTL.fn_nextpc).
Proof.
intros f tf. unfold transf_function.
- destruct (type_rtl_function f).
+ destruct (type_function f).
destruct (analyze f).
- destruct (regalloc f t0).
+ destruct (regalloc f t).
intro EQ; injection EQ; intro EQ1; rewrite <- EQ1; simpl; auto.
intros; discriminate.
intros; discriminate.
@@ -1220,43 +1278,43 @@ Qed.
Definition exec_instr_prop
(c: RTL.code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
forall f env live assign ls
(CF: c = f.(RTL.fn_code))
- (WT: wt_function env f)
+ (WT: wt_function f env)
(ASG: regalloc f live (live0 f live) env = Some assign)
(AG: agree assign (transfer f pc live!!pc) rs ls),
let tc := PTree.map (transf_instr f live assign) c in
exists ls',
- exec_blocks tge tc sp pc ls m (Cont pc') ls' m' /\
+ exec_blocks tge tc sp pc ls m t (Cont pc') ls' m' /\
agree assign live!!pc rs' ls'.
Definition exec_instrs_prop
(c: RTL.code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
forall f env live assign ls,
forall (CF: c = f.(RTL.fn_code))
- (WT: wt_function env f)
+ (WT: wt_function f env)
(ANL: analyze f = Some live)
(ASG: regalloc f live (live0 f live) env = Some assign)
(AG: agree assign (transfer f pc live!!pc) rs ls)
(VALIDPC': c!pc' <> None),
let tc := PTree.map (transf_instr f live assign) c in
exists ls',
- exec_blocks tge tc sp pc ls m (Cont pc') ls' m' /\
+ exec_blocks tge tc sp pc ls m t (Cont pc') ls' m' /\
agree assign (transfer f pc' live!!pc') rs' ls'.
Definition exec_function_prop
- (f: RTL.function) (args: list val) (m: mem)
- (res: val) (m': mem) : Prop :=
+ (f: RTL.fundef) (args: list val) (m: mem)
+ (t: trace) (res: val) (m': mem) : Prop :=
forall ls tf,
- transf_function f = Some tf ->
- List.map ls (Conventions.loc_arguments tf.(fn_sig)) = args ->
+ transf_fundef f = Some tf ->
+ List.map ls (Conventions.loc_arguments (funsig tf)) = args ->
exists ls',
- LTL.exec_function tge tf ls m ls' m' /\
- ls' (R (Conventions.loc_result tf.(fn_sig))) = res.
+ LTL.exec_function tge tf ls m t ls' m' /\
+ ls' (R (Conventions.loc_result (funsig tf))) = res.
(** The simulation proof is by structural induction over the RTL evaluation
derivation. We prove each case of the proof as a separate lemma.
@@ -1298,7 +1356,7 @@ Lemma transl_Inop_correct:
forall (c : PTree.t instruction) (sp: val) (pc : positive)
(rs : regset) (m : mem) (pc' : RTL.node),
c ! pc = Some (Inop pc') ->
- exec_instr_prop c sp pc rs m pc' rs m.
+ exec_instr_prop c sp pc rs m E0 pc' rs m.
Proof.
intros; red; intros; CleanupHyps.
exists ls. split.
@@ -1312,7 +1370,7 @@ Lemma transl_Iop_correct:
(res : reg) (pc' : RTL.node) (v: val),
c ! pc = Some (Iop op args res pc') ->
eval_operation ge sp op (rs ## args) = Some v ->
- exec_instr_prop c sp pc rs m pc' (rs # res <- v) m.
+ exec_instr_prop c sp pc rs m E0 pc' (rs # res <- v) m.
Proof.
intros; red; intros; CleanupHyps.
caseEq (Regset.mem res live!!pc); intro LV;
@@ -1364,7 +1422,7 @@ Lemma transl_Iload_correct:
c ! pc = Some (Iload chunk addr args dst pc') ->
eval_addressing ge sp addr rs ## args = Some a ->
loadv chunk m a = Some v ->
- exec_instr_prop c sp pc rs m pc' rs # dst <- v m.
+ exec_instr_prop c sp pc rs m E0 pc' rs # dst <- v m.
Proof.
intros; red; intros; CleanupHyps.
caseEq (Regset.mem dst live!!pc); intro LV;
@@ -1407,7 +1465,7 @@ Lemma transl_Istore_correct:
c ! pc = Some (Istore chunk addr args src pc') ->
eval_addressing ge sp addr rs ## args = Some a ->
storev chunk m a rs # src = Some m' ->
- exec_instr_prop c sp pc rs m pc' rs m'.
+ exec_instr_prop c sp pc rs m E0 pc' rs m'.
Proof.
intros; red; intros; CleanupHyps.
assert (LL: (List.length (List.map assign args) <= 2)%nat).
@@ -1444,19 +1502,19 @@ Lemma transl_Icall_correct:
forall (c : PTree.t instruction) (sp: val) (pc : positive)
(rs : regset) (m : mem) (sig : signature) (ros : reg + ident)
(args : list reg) (res : reg) (pc' : RTL.node)
- (f : RTL.function) (vres : val) (m' : mem),
+ (f : RTL.fundef) (vres : val) (m' : mem) (t: trace),
c ! pc = Some (Icall sig ros args res pc') ->
RTL.find_function ge ros rs = Some f ->
- sig = RTL.fn_sig f ->
- RTL.exec_function ge f (rs##args) m vres m' ->
- exec_function_prop f (rs##args) m vres m' ->
- exec_instr_prop c sp pc rs m pc' (rs#res <- vres) m'.
+ RTL.funsig f = sig ->
+ RTL.exec_function ge f (rs##args) m t vres m' ->
+ exec_function_prop f (rs##args) m t vres m' ->
+ exec_instr_prop c sp pc rs m t pc' (rs#res <- vres) m'.
Proof.
intros; red; intros; CleanupHyps.
set (los := sum_left_map assign ros).
assert (FIND: exists tf,
find_function2 tge los ls = Some tf /\
- transf_function f = Some tf).
+ transf_fundef f = Some tf).
unfold los. destruct ros; simpl; simpl in H0.
apply functions_translated.
replace (ls (assign r)) with rs#r. assumption.
@@ -1466,7 +1524,7 @@ Proof.
apply function_ptr_translated. auto.
discriminate.
elim FIND; intros tf [AFIND TRF]; clear FIND.
- assert (ASIG: sig = fn_sig tf).
+ assert (ASIG: sig = funsig tf).
rewrite (sig_function_translated _ _ TRF). auto.
generalize (fun ls => H3 ls tf TRF); intro AEXECF.
assert (AVARGS: List.map ls (List.map assign args) = rs##args).
@@ -1482,7 +1540,7 @@ Proof.
unfold correct_alloc_instr. intros [CORR1 CORR2].
assert (ARES: loc_acceptable (assign res)).
eapply regalloc_acceptable; eauto.
- generalize (add_call_correct tge sp tf _ _ _ _ _ _ _ _ pc' _
+ generalize (add_call_correct tge sp tf _ _ _ _ _ _ _ _ _ pc' _
AEXECF AFIND ASIG AVARGS ALARGS
AACCEPT ARES).
intros [ls' [EX [RES OTHER]]].
@@ -1491,13 +1549,42 @@ Proof.
simpl. eapply agree_call; eauto.
Qed.
+Lemma transl_Ialloc_correct:
+ forall (c : PTree.t instruction) (sp: val) (pc : positive)
+ (rs : Regmap.t val) (m : mem) (pc': RTL.node) (arg res: reg)
+ (sz: int) (m': mem) (b: Values.block),
+ c ! pc = Some (Ialloc arg res pc') ->
+ rs#arg = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', b) ->
+ exec_instr_prop c sp pc rs m E0 pc' (rs # res <- (Vptr b Int.zero)) m'.
+Proof.
+ intros; red; intros; CleanupHyps.
+ assert (SZ: ls (assign arg) = Vint sz).
+ rewrite <- H0. eapply agree_eval_reg. eauto.
+ generalize (add_alloc_correct tge sp (assign arg) (assign res)
+ pc' ls m m' sz b SZ H1).
+ intros [ls' [EX [RES OTHER]]].
+ exists ls'.
+ split. CleanupGoal. exact EX.
+ rewrite CF in H.
+ generalize (regalloc_correct_1 f env live _ _ _ _ ASG H).
+ unfold correct_alloc_instr. intros [CORR1 CORR2].
+ eapply agree_call with (args := arg :: nil) (ros := inr reg 1%positive) (ls := ls) (ls' := ls'); eauto.
+ intros. apply OTHER.
+ eapply Loc.in_notin_diff; eauto.
+ unfold loc_alloc_argument, destroyed_at_call; simpl; tauto.
+ eapply Loc.in_notin_diff; eauto.
+ unfold loc_alloc_argument, destroyed_at_call; simpl; tauto.
+ auto.
+Qed.
+
Lemma transl_Icond_true_correct:
forall (c : PTree.t instruction) (sp: val) (pc : positive)
(rs : Regmap.t val) (m : mem) (cond : condition) (args : list reg)
(ifso ifnot : RTL.node),
c ! pc = Some (Icond cond args ifso ifnot) ->
eval_condition cond rs ## args = Some true ->
- exec_instr_prop c sp pc rs m ifso rs m.
+ exec_instr_prop c sp pc rs m E0 ifso rs m.
Proof.
intros; red; intros; CleanupHyps.
assert (LL: (List.length (map assign args) <= 3)%nat).
@@ -1523,7 +1610,7 @@ Lemma transl_Icond_false_correct:
(ifso ifnot : RTL.node),
c ! pc = Some (Icond cond args ifso ifnot) ->
eval_condition cond rs ## args = Some false ->
- exec_instr_prop c sp pc rs m ifnot rs m.
+ exec_instr_prop c sp pc rs m E0 ifnot rs m.
Proof.
intros; red; intros; CleanupHyps.
assert (LL: (List.length (map assign args) <= 3)%nat).
@@ -1545,7 +1632,7 @@ Qed.
Lemma transl_refl_correct:
forall (c : RTL.code) (sp: val) (pc : RTL.node) (rs : regset)
- (m : mem), exec_instrs_prop c sp pc rs m pc rs m.
+ (m : mem), exec_instrs_prop c sp pc rs m E0 pc rs m.
Proof.
intros; red; intros.
exists ls. split. apply exec_blocks_refl. assumption.
@@ -1553,10 +1640,10 @@ Qed.
Lemma transl_one_correct:
forall (c : RTL.code) (sp: val) (pc : RTL.node) (rs : regset)
- (m : mem) (pc' : RTL.node) (rs' : regset) (m' : mem),
- RTL.exec_instr ge c sp pc rs m pc' rs' m' ->
- exec_instr_prop c sp pc rs m pc' rs' m' ->
- exec_instrs_prop c sp pc rs m pc' rs' m'.
+ (m : mem) (t: trace) (pc' : RTL.node) (rs' : regset) (m' : mem),
+ RTL.exec_instr ge c sp pc rs m t pc' rs' m' ->
+ exec_instr_prop c sp pc rs m t pc' rs' m' ->
+ exec_instrs_prop c sp pc rs m t pc' rs' m'.
Proof.
intros; red; intros.
generalize (H0 f env live assign ls CF WT ASG AG).
@@ -1573,13 +1660,14 @@ Qed.
Lemma transl_trans_correct:
forall (c : RTL.code) (sp: val) (pc1 : RTL.node) (rs1 : regset)
- (m1 : mem) (pc2 : RTL.node) (rs2 : regset) (m2 : mem)
- (pc3 : RTL.node) (rs3 : regset) (m3 : mem),
- RTL.exec_instrs ge c sp pc1 rs1 m1 pc2 rs2 m2 ->
- exec_instrs_prop c sp pc1 rs1 m1 pc2 rs2 m2 ->
- RTL.exec_instrs ge c sp pc2 rs2 m2 pc3 rs3 m3 ->
- exec_instrs_prop c sp pc2 rs2 m2 pc3 rs3 m3 ->
- exec_instrs_prop c sp pc1 rs1 m1 pc3 rs3 m3.
+ (m1 : mem) (t1: trace) (pc2 : RTL.node) (rs2 : regset) (m2 : mem)
+ (t2: trace) (pc3 : RTL.node) (rs3 : regset) (m3 : mem) (t3: trace),
+ RTL.exec_instrs ge c sp pc1 rs1 m1 t1 pc2 rs2 m2 ->
+ exec_instrs_prop c sp pc1 rs1 m1 t1 pc2 rs2 m2 ->
+ RTL.exec_instrs ge c sp pc2 rs2 m2 t2 pc3 rs3 m3 ->
+ exec_instrs_prop c sp pc2 rs2 m2 t2 pc3 rs3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_instrs_prop c sp pc1 rs1 m1 t3 pc3 rs3 m3.
Proof.
intros; red; intros.
assert (VALIDPC2: c!pc2 <> None).
@@ -1589,7 +1677,7 @@ Proof.
generalize (H2 f env live assign ls1 CF WT ANL ASG AG1 VALIDPC').
intros [ls2 [EX2 AG2]].
exists ls2. split.
- eapply exec_blocks_trans. eexact EX1. exact EX2.
+ eapply exec_blocks_trans. eexact EX1. eexact EX2. auto.
exact AG2.
Qed.
@@ -1609,14 +1697,14 @@ Qed.
Lemma transf_entrypoint_correct:
forall f env live assign c ls args sp m,
- wt_function env f ->
+ wt_function f env ->
regalloc f live (live0 f live) env = Some assign ->
c!(RTL.fn_nextpc f) = None ->
List.map ls (loc_parameters (RTL.fn_sig f)) = args ->
(forall ofs ty, ls (S (Local ofs ty)) = Vundef) ->
let tc := transf_entrypoint f live assign c in
exists ls',
- exec_blocks tge tc sp (RTL.fn_nextpc f) ls m
+ exec_blocks tge tc sp (RTL.fn_nextpc f) ls m E0
(Cont (RTL.fn_entrypoint f)) ls' m /\
agree assign (transfer f (RTL.fn_entrypoint f) live!!(RTL.fn_entrypoint f))
(init_regs args (RTL.fn_params f)) ls'.
@@ -1680,20 +1768,20 @@ Qed.
Lemma transl_function_correct:
forall (f : RTL.function) (m m1 : mem) (stk : Values.block)
- (args : list val) (pc : RTL.node) (rs : regset) (m2 : mem)
+ (args : list val) (t: trace) (pc : RTL.node) (rs : regset) (m2 : mem)
(or : option reg) (vres : val),
alloc m 0 (RTL.fn_stacksize f) = (m1, stk) ->
RTL.exec_instrs ge (RTL.fn_code f) (Vptr stk Int.zero)
- (RTL.fn_entrypoint f) (init_regs args (fn_params f)) m1 pc rs m2 ->
+ (RTL.fn_entrypoint f) (init_regs args (fn_params f)) m1 t pc rs m2 ->
exec_instrs_prop (RTL.fn_code f) (Vptr stk Int.zero)
- (RTL.fn_entrypoint f) (init_regs args (fn_params f)) m1 pc rs m2 ->
+ (RTL.fn_entrypoint f) (init_regs args (fn_params f)) m1 t pc rs m2 ->
(RTL.fn_code f) ! pc = Some (Ireturn or) ->
vres = regmap_optget or Vundef rs ->
- exec_function_prop f args m vres (free m2 stk).
+ exec_function_prop (Internal f) args m t vres (free m2 stk).
Proof.
intros; red; intros until tf.
- unfold transf_function.
- caseEq (type_rtl_function f).
+ unfold transf_fundef, transf_function.
+ caseEq (type_function f).
intros env TRF.
caseEq (analyze f).
intros live ANL.
@@ -1704,7 +1792,7 @@ Proof.
set (tc1 := PTree.map (transf_instr f live alloc) (RTL.fn_code f)).
set (tc2 := transf_entrypoint f live alloc tc1).
intro EQ; injection EQ; intro TF; clear EQ. intro VARGS.
- generalize (type_rtl_function_correct _ _ TRF); intro WTF.
+ generalize (type_function_correct _ _ TRF); intro WTF.
assert (NEWINSTR: tc1!(RTL.fn_nextpc f) = None).
unfold tc1; rewrite PTree.gmap. unfold option_map.
elim (RTL.fn_code_wf f (fn_nextpc f)); intro.
@@ -1712,7 +1800,7 @@ Proof.
rewrite H4. auto.
pose (ls1 := call_regs ls).
assert (VARGS1: List.map ls1 (loc_parameters (RTL.fn_sig f)) = args).
- replace (RTL.fn_sig f) with (fn_sig tf).
+ replace (RTL.fn_sig f) with (funsig tf).
rewrite <- VARGS. unfold loc_parameters.
rewrite list_map_compose. apply list_map_exten.
intros. symmetry. unfold ls1. eapply call_regs_param_of_arg; eauto.
@@ -1732,9 +1820,9 @@ Proof.
intros [ls4 [EX4 RES4]].
exists ls4.
(* Execution *)
- split. apply exec_funct with m1.
- rewrite <- TF; simpl; assumption.
- rewrite <- TF; simpl. fold ls1.
+ split. rewrite <- TF; apply exec_funct_internal with m1; simpl.
+ assumption.
+ fold ls1.
eapply exec_blocks_trans. eexact EX2.
apply exec_blocks_extends with tc1.
intro p. unfold tc2. unfold transf_entrypoint.
@@ -1746,7 +1834,7 @@ Proof.
eapply exec_blocks_trans. eexact EX3.
eapply exec_blocks_one.
unfold tc1. rewrite PTree.gmap. rewrite H2. simpl. reflexivity.
- exact EX4.
+ eexact EX4. reflexivity. traceEq.
destruct or.
simpl in RES4. simpl in H3.
rewrite H3. rewrite <- TF; simpl. rewrite RES4.
@@ -1760,6 +1848,20 @@ Proof.
intros; discriminate.
Qed.
+Lemma transl_external_function_correct:
+ forall (ef : external_function) (args : list val) (m : mem)
+ (t: trace) (res: val),
+ event_match ef args t res ->
+ exec_function_prop (External ef) args m t res m.
+Proof.
+ intros; red; intros.
+ simpl in H0.
+ destruct (type_external_function ef); simplify_eq H0; intro.
+ exists (Locmap.set (R (loc_result (funsig tf))) res ls); split.
+ subst tf. eapply exec_funct_external; eauto.
+ apply Locmap.gss.
+Qed.
+
(** The correctness of the code transformation is obtained by
structural induction (and case analysis on the last rule used)
on the RTL evaluation derivation.
@@ -1767,14 +1869,10 @@ Qed.
[exec_instrs_prop] and [exec_function_prop] as the induction
hypotheses, and the lemmas above as cases for the induction. *)
-Scheme exec_instr_ind_3 := Minimality for RTL.exec_instr Sort Prop
- with exec_instrs_ind_3 := Minimality for RTL.exec_instrs Sort Prop
- with exec_function_ind_3 := Minimality for RTL.exec_function Sort Prop.
-
Theorem transl_function_correctness:
- forall f args m res m',
- RTL.exec_function ge f args m res m' ->
- exec_function_prop f args m res m'.
+ forall f args m t res m',
+ RTL.exec_function ge f args m t res m' ->
+ exec_function_prop f args m t res m'.
Proof
(exec_function_ind_3 ge
exec_instr_prop
@@ -1786,6 +1884,7 @@ Proof
transl_Iload_correct
transl_Istore_correct
transl_Icall_correct
+ transl_Ialloc_correct
transl_Icond_true_correct
transl_Icond_false_correct
@@ -1793,23 +1892,24 @@ Proof
transl_one_correct
transl_trans_correct
- transl_function_correct).
+ transl_function_correct
+ transl_external_function_correct).
(** The semantic equivalence between the original and transformed programs
follows easily. *)
Theorem transl_program_correct:
- forall (r: val),
- RTL.exec_program prog r -> LTL.exec_program tprog r.
+ forall (t: trace) (r: val),
+ RTL.exec_program prog t r -> LTL.exec_program tprog t r.
Proof.
- intros r [b [f [m [FIND1 [FIND2 [SIG EXEC]]]]]].
+ intros t r [b [f [m [FIND1 [FIND2 [SIG EXEC]]]]]].
generalize (function_ptr_translated _ _ FIND2).
intros [tf [TFIND TRF]].
- assert (SIG2: tf.(fn_sig) = mksignature nil (Some Tint)).
+ assert (SIG2: funsig tf = mksignature nil (Some Tint)).
rewrite <- SIG. apply sig_function_translated; auto.
- assert (VPARAMS: map (Locmap.init Vundef) (loc_arguments (fn_sig tf)) = nil).
+ assert (VPARAMS: map (Locmap.init Vundef) (loc_arguments (funsig tf)) = nil).
rewrite SIG2. reflexivity.
- generalize (transl_function_correctness _ _ _ _ _ EXEC
+ generalize (transl_function_correctness _ _ _ _ _ _ EXEC
(Locmap.init Vundef) tf TRF VPARAMS).
intros [ls' [TEXEC RES]].
red. exists b; exists tf; exists ls'; exists m.
diff --git a/backend/Allocproof_aux.v b/backend/Allocproof_aux.v
deleted file mode 100644
index d1b213e2..00000000
--- a/backend/Allocproof_aux.v
+++ /dev/null
@@ -1,850 +0,0 @@
-(** Correctness results for the [parallel_move] function defined in
- file [Allocation].
-
- The present file, contributed by Laurence Rideau, glues the general
- results over the parallel move algorithm (see file [Parallelmove])
- with the correctness proof for register allocation (file [Allocproof]).
-*)
-
-Require Import Coqlib.
-Require Import Values.
-Require Import Parallelmove.
-Require Import Allocation.
-Require Import LTL.
-Require Import Locations.
-Require Import Conventions.
-
-Section parallel_move_correction.
-Variable ge : LTL.genv.
-Variable sp : val.
-Hypothesis
- add_move_correct :
- forall src dst k rs m,
- (exists rs' ,
- exec_instrs ge sp (add_move src dst k) rs m k rs' m /\
- (rs' dst = rs src /\
- (forall l,
- Loc.diff l dst ->
- Loc.diff l (R IT1) -> Loc.diff l (R FT1) -> rs' l = rs l)) ).
-
-Lemma exec_instr_update:
- forall a1 a2 k e m,
- (exists rs' ,
- exec_instrs ge sp (add_move a1 a2 k) e m k rs' m /\
- (rs' a2 = update e a2 (e a1) a2 /\
- (forall l,
- Loc.diff l a2 ->
- Loc.diff l (R IT1) -> Loc.diff l (R FT1) -> rs' l = (update e a2 (e a1)) l))
- ).
-Proof.
-intros; destruct (add_move_correct a1 a2 k e m) as [rs' [Hf [R1 R2]]].
-exists rs'; (repeat split); auto.
-generalize (get_update_id e a2); unfold get, Locmap.get; intros H; rewrite H;
- auto.
-intros l H0; generalize (get_update_diff e a2); unfold get, Locmap.get;
- intros H; rewrite H; auto.
-apply Loc.diff_sym; assumption.
-Qed.
-
-Lemma map_inv:
- forall (A B : Set) (f f' : A -> B) l,
- map f l = map f' l -> forall x, In x l -> f x = f' x.
-Proof.
-induction l; simpl; intros Hmap x Hin.
-elim Hin.
-inversion Hmap.
-elim Hin; intros H; [subst a | apply IHl]; auto.
-Qed.
-
-Fixpoint reswellFormed (p : Moves) : Prop :=
- match p with
- nil => True
- | (s, d) :: l => Loc.notin s (getdst l) /\ reswellFormed l
- end.
-
-Definition temporaries1 := R IT1 :: (R FT1 :: nil).
-
-Lemma no_overlap_list_pop:
- forall p m, no_overlap_list (m :: p) -> no_overlap_list p.
-Proof.
-induction p; unfold no_overlap_list, no_overlap; destruct m as [m1 m2]; simpl;
- auto.
-Qed.
-
-Lemma exec_instrs_pmov:
- forall p k rs m,
- no_overlap_list p ->
- Loc.disjoint (getdst p) temporaries1 ->
- Loc.disjoint (getsrc p) temporaries1 ->
- (exists rs' ,
- exec_instrs
- ge sp
- (fold_left
- (fun (k0 : block) => fun (p0 : loc * loc) => add_move (fst p0) (snd p0) k0)
- p k) rs m k rs' m /\
- (forall l,
- (forall r, In r (getdst p) -> r = l \/ Loc.diff r l) ->
- Loc.diff l (Locations.R IT1) ->
- Loc.diff l (Locations.R FT1) -> rs' l = (sexec p rs) l) ).
-Proof.
-induction p; intros k rs m.
-simpl; exists rs; (repeat split); auto; apply exec_refl.
-destruct a as [a1 a2]; simpl; intros Hno_O Hdisd Hdiss;
- elim (IHp (add_move a1 a2 k) rs m);
- [idtac | apply no_overlap_list_pop with (a1, a2) |
- apply (Loc.disjoint_cons_left a2) | apply (Loc.disjoint_cons_left a1)];
- (try assumption).
-intros rs' Hexec;
- destruct (add_move_correct a1 a2 k rs' m) as [rs'' [Hexec2 [R1 R2]]].
-exists rs''; elim Hexec; intros; (repeat split).
-apply exec_trans with ( b2 := add_move a1 a2 k ) ( rs2 := rs' ) ( m2 := m );
- auto.
-intros l Heqd Hdi Hdf; case (Loc.eq a2 l); intro.
-subst l; generalize get_update_id; unfold get, Locmap.get; intros Hgui;
- rewrite Hgui; rewrite R1.
-apply H0; auto.
-unfold no_overlap_list, no_overlap in Hno_O |-; simpl in Hno_O |-; intros;
- generalize (Hno_O a1).
-intros H8; elim H8 with ( s := r );
- [intros H9; left | intros H9; right; apply Loc.diff_sym | left | right]; auto.
-unfold Loc.disjoint in Hdiss |-; apply Hdiss; simpl; left; trivial.
-apply Hdiss; simpl; [left | right; left]; auto.
-elim (Heqd a2);
- [intros H7; absurd (a2 = l); auto | intros H7; auto | left; trivial].
-generalize get_update_diff; unfold get, Locmap.get; intros H6; rewrite H6; auto.
-rewrite R2; auto.
-apply Loc.diff_sym; auto.
-Qed.
-
-Definition p_move :=
- fun (l : list (loc * loc)) =>
- fun (k : block) =>
- fold_left
- (fun (k0 : block) => fun (p : loc * loc) => add_move (fst p) (snd p) k0)
- (P_move l) k.
-
-Lemma norepet_SD: forall p, Loc.norepet (getdst p) -> simpleDest p.
-Proof.
-induction p; (simpl; auto).
-destruct a as [a1 a2].
-intro; inversion H.
-split.
-apply notindst_nW; auto.
-apply IHp; auto.
-Qed.
-
-Theorem SDone_stepf:
- forall S1, StateDone (stepf S1) = nil -> StateDone S1 = nil.
-Proof.
-destruct S1 as [[t b] d]; destruct t.
-destruct b; auto.
-destruct m as [m1 m2]; simpl.
-destruct b.
-simpl; intro; discriminate.
-case (Loc.eq m2 (fst (last (m :: b)))); simpl; intros; discriminate.
-destruct m as [a1 a2]; simpl.
-destruct b.
-case (Loc.eq a1 a2); simpl; intros; auto.
-destruct m as [m1 m2]; case (Loc.eq a1 m2); intro; (try (simpl; auto; fail)).
-case (split_move t m2).
-(repeat destruct p); simpl; intros; auto.
-destruct b; (try (simpl; intro; discriminate)); auto.
-case (Loc.eq m2 (fst (last (m :: b)))); intro; simpl; intro; discriminate.
-Qed.
-
-Theorem SDone_Pmov: forall S1, StateDone (Pmov S1) = nil -> StateDone S1 = nil.
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1; intros Hrec.
-destruct S1 as [[t b] d].
-rewrite Pmov_equation.
-destruct t.
-destruct b; auto.
-intro; assert (StateDone (stepf (nil, m :: b, d)) = nil);
- [apply Hrec; auto; apply stepf1_dec | apply SDone_stepf]; auto.
-intro; assert (StateDone (stepf (m :: t, b, d)) = nil);
- [apply Hrec; auto; apply stepf1_dec | apply SDone_stepf]; auto.
-Qed.
-
-Lemma no_overlap_temp: forall r s, In s temporaries -> r = s \/ Loc.diff r s.
-Proof.
-intros r s H; case (Loc.eq r s); intros e; [left | right]; (try assumption).
-unfold Loc.diff; destruct r.
-destruct s; trivial.
-red; intro; elim e; rewrite H0; auto.
-destruct s0; destruct s; trivial;
- (elim H; [intros H1 | intros [H1|[H1|[H1|[H1|[H1|H1]]]]]]; (try discriminate);
- inversion H1).
-Qed.
-
-Lemma getsrcdst_app:
- forall l1 l2,
- getdst l1 ++ getdst l2 = getsrc l1 ++ getsrc l2 ->
- getdst l1 = getsrc l1 /\ getdst l2 = getsrc l2.
-Proof.
-induction l1; simpl; auto.
-destruct a as [a1 a2]; intros; inversion H.
-subst a1; inversion H;
- (elim IHl1 with ( l2 := l2 );
- [intros H0 H3; (try clear IHl1); (try exact H0) | idtac]; auto).
-rewrite H0; auto.
-Qed.
-
-Lemma In_norepet:
- forall r x l, Loc.norepet l -> In x l -> In r l -> r = x \/ Loc.diff r x.
-Proof.
-induction l; simpl; intros.
-elim H1.
-inversion H.
-subst hd.
-elim H1; intros H2; clear H1; elim H0; intros H1.
-rewrite <- H1; rewrite <- H2; auto.
-subst a; right; apply Loc.in_notin_diff with l; auto.
-subst a; right; apply Loc.diff_sym; apply Loc.in_notin_diff with l; auto.
-apply IHl; auto.
-Qed.
-
-Definition no_overlap_stateD (S : State) :=
- no_overlap
- (getsrc (StateToMove S ++ (StateBeing S ++ StateDone S)))
- (getdst (StateToMove S ++ (StateBeing S ++ StateDone S))).
-
-Ltac
-incl_tac_rec :=
-(auto with datatypes; fail)
- ||
- (apply in_cons; incl_tac_rec)
- ||
- (apply in_or_app; left; incl_tac_rec)
- ||
- (apply in_or_app; right; incl_tac_rec)
- ||
- (apply incl_appl; incl_tac_rec) ||
- (apply incl_appr; incl_tac_rec) || (apply incl_tl; incl_tac_rec).
-
-Ltac incl_tac := (repeat (apply incl_cons || apply incl_app)); incl_tac_rec.
-
-Ltac
-in_tac :=
-match goal with
-| |- In ?x ?L1 =>
-match goal with
-| H:In x ?L2 |- _ =>
-let H1 := fresh "H" in
-(assert (H1: incl L2 L1); [incl_tac | apply (H1 x)]); auto; fail
-| _ => fail end end.
-
-Lemma in_cons_noteq:
- forall (A : Set) (a b : A) (l : list A), In a (b :: l) -> a <> b -> In a l.
-Proof.
-intros A a b l; simpl; intros.
-elim H; intros H1; (try assumption).
-absurd (a = b); auto.
-Qed.
-
-Lemma no_overlapD_inv:
- forall S1 S2, step S1 S2 -> no_overlap_stateD S1 -> no_overlap_stateD S2.
-Proof.
-intros S1 S2 STEP; inversion STEP; unfold no_overlap_stateD, no_overlap; simpl;
- auto; (repeat (rewrite getsrc_app; rewrite getdst_app; simpl)); intros.
-apply H1; in_tac.
-destruct m as [m1 m2]; apply H1; in_tac.
-apply H1; in_tac.
-case (Loc.eq r (T r0)); intros e.
-elim (no_overlap_temp s0 r);
- [intro; left; auto | intro; right; apply Loc.diff_sym; auto | unfold T in e |-].
-destruct (Loc.type r0); simpl; [right; left | right; right; right; right; left];
- auto.
-case (Loc.eq s0 (T r0)); intros es.
-apply (no_overlap_temp r s0); unfold T in es |-; destruct (Loc.type r0); simpl;
- [right; left | right; right; right; right; left]; auto.
-apply H1; apply in_cons_noteq with ( b := T r0 ); auto; in_tac.
-apply H3; in_tac.
-Qed.
-
-Lemma no_overlapD_invpp:
- forall S1 S2, stepp S1 S2 -> no_overlap_stateD S1 -> no_overlap_stateD S2.
-Proof.
-intros; induction H as [r|r1 r2 r3 H H1 HrecH]; auto.
-apply HrecH; auto.
-apply no_overlapD_inv with r1; auto.
-Qed.
-
-Lemma no_overlapD_invf:
- forall S1, stepInv S1 -> no_overlap_stateD S1 -> no_overlap_stateD (stepf S1).
-Proof.
-intros; destruct S1 as [[t1 b1] d1].
-destruct t1; destruct b1; auto.
-set (S1:=(nil (A:=Move), m :: b1, d1)).
-apply (no_overlapD_invpp S1); [apply dstep_step; auto | assumption].
-apply f2ind; [unfold S1 | idtac | idtac]; auto.
-generalize H0; clear H0; unfold no_overlap_stateD; destruct m as [m1 m2].
-intros; apply no_overlap_noOverlap.
-unfold no_overlap_state; simpl.
-generalize H0; clear H0; unfold no_overlap; (repeat rewrite getdst_app);
- (repeat rewrite getsrc_app); simpl; intros.
-apply H0.
-elim H1; intros H4; [left; assumption | right; in_tac].
-elim H2; intros H4; [left; assumption | right; in_tac].
-set (S1:=(m :: t1, nil (A:=Move), d1)).
-apply (no_overlapD_invpp S1); [apply dstep_step; auto | assumption].
-apply f2ind; [unfold S1 | idtac | idtac]; auto.
-generalize H0; clear H0; unfold no_overlap_stateD; destruct m as [m1 m2].
-intros; apply no_overlap_noOverlap.
-unfold no_overlap_state; simpl.
-generalize H0; clear H0; unfold no_overlap; (repeat rewrite getdst_app);
- (repeat rewrite getsrc_app); simpl; (repeat rewrite app_nil); intros.
-apply H0.
-elim H1; intros H4; [left; assumption | right; (try in_tac)].
-elim H2; intros H4; [left; assumption | right; in_tac].
-set (S1:=(m :: t1, m0 :: b1, d1)).
-apply (no_overlapD_invpp S1); [apply dstep_step; auto | assumption].
-apply f2ind; [unfold S1 | idtac | idtac]; auto.
-generalize H0; clear H0; unfold no_overlap_stateD; destruct m as [m1 m2].
-intros; apply no_overlap_noOverlap.
-unfold no_overlap_state; simpl.
-generalize H0; clear H0; unfold no_overlap; (repeat rewrite getdst_app);
- (repeat rewrite getsrc_app); destruct m0; simpl; intros.
-apply H0.
-elim H1; intros H4; [left; assumption | right; in_tac].
-elim H2; intros H4; [left; assumption | right; in_tac].
-Qed.
-
-Lemma no_overlapD_res:
- forall S1, stepInv S1 -> no_overlap_stateD S1 -> no_overlap_stateD (Pmov S1).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d].
-rewrite Pmov_equation.
-destruct t; auto.
-destruct b; auto.
-intros; apply Hrec.
-apply stepf1_dec; auto.
-apply (dstep_inv (nil, m :: b, d)); auto.
-apply f2ind'; auto.
-apply no_overlap_noOverlap.
-generalize H0; unfold no_overlap_stateD, no_overlap_state, no_overlap; simpl.
-destruct m; (repeat rewrite getdst_app); (repeat rewrite getsrc_app).
-intros H1 r1 H2 s H3; (try assumption).
-apply H1; in_tac.
-apply no_overlapD_invf; auto.
-intros; apply Hrec.
-apply stepf1_dec; auto.
-apply (dstep_inv (m :: t, b, d)); auto.
-apply f2ind'; auto.
-apply no_overlap_noOverlap.
-generalize H0; destruct m;
- unfold no_overlap_stateD, no_overlap_state, no_overlap; simpl;
- (repeat (rewrite getdst_app; simpl)); (repeat (rewrite getsrc_app; simpl)).
-simpl; intros H1 r1 H2 s H3; (try assumption).
-apply H1.
-elim H2; intros H4; [left; (try assumption) | right; in_tac].
-elim H3; intros H4; [left; (try assumption) | right; in_tac].
-apply no_overlapD_invf; auto.
-Qed.
-
-Definition temporaries1_3 := R IT1 :: (R FT1 :: (R IT3 :: nil)).
-
-Definition temporaries2 := R IT2 :: (R FT2 :: nil).
-
-Definition no_tmp13_state (S1 : State) :=
- Loc.disjoint (getsrc (StateDone S1)) temporaries1_3 /\
- Loc.disjoint (getdst (StateDone S1)) temporaries1_3.
-
-Definition Done_well_formed (S1 S2 : State) :=
- forall x,
- (In x (getsrc (StateDone S2)) ->
- In x temporaries2 \/ In x (getsrc (StateToMove S1 ++ StateBeing S1))) /\
- (In x (getdst (StateDone S2)) ->
- In x temporaries2 \/ In x (getdst (StateToMove S1 ++ StateBeing S1))).
-
-Lemma Done_notmp3_inv:
- forall S1 S2,
- step S1 S2 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT3)) ->
- forall x,
- In x (getdst (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- Loc.diff x (R IT3).
-Proof.
-intros S1 S2 STEP; inversion STEP; (try (simpl; trivial; fail));
- simpl StateDone; simpl StateToMove; simpl StateBeing; simpl getdst;
- (repeat (rewrite getdst_app; simpl)); intros.
-apply H1; in_tac.
-destruct m; apply H1; in_tac.
-apply H1; in_tac.
-case (Loc.eq x (T r0)); intros e.
-unfold T in e |-; destruct (Loc.type r0); rewrite e; simpl; red; intro;
- discriminate.
-apply H1; apply in_cons_noteq with ( b := T r0 ); auto; in_tac.
-apply H3; in_tac.
-Qed.
-
-Lemma Done_notmp3_invpp:
- forall S1 S2,
- stepp S1 S2 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT3)) ->
- forall x,
- In x (getdst (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- Loc.diff x (R IT3).
-Proof.
-intros S1 S2 H H0; (try assumption).
-induction H as [r|r1 r2 r3 H1 H2 Hrec]; auto.
-apply Hrec; auto.
-apply Done_notmp3_inv with r1; auto.
-Qed.
-
-Lemma Done_notmp3_invf:
- forall S1,
- stepInv S1 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT3)) ->
- forall x,
- In
- x
- (getdst
- (StateToMove (stepf S1) ++ (StateBeing (stepf S1) ++ StateDone (stepf S1)))) ->
- Loc.diff x (R IT3).
-Proof.
-intros S1 H H0; (try assumption).
-generalize H; unfold stepInv; intros [Hpath [HSD [HnoO [Hnotmp HnotmpL]]]].
-destruct S1 as [[t1 b1] d1]; set (S1:=(t1, b1, d1)); destruct t1; destruct b1;
- auto; apply (Done_notmp3_invpp S1); auto; apply dstep_step; auto; apply f2ind;
- unfold S1; auto.
-Qed.
-
-Lemma Done_notmp3_res:
- forall S1,
- stepInv S1 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT3)) ->
- forall x,
- In
- x
- (getdst
- (StateToMove (Pmov S1) ++ (StateBeing (Pmov S1) ++ StateDone (Pmov S1)))) ->
- Loc.diff x (R IT3).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d]; set (S1:=(t, b, d)).
-unfold S1; rewrite Pmov_equation.
-intros H; generalize H; intros [Hpath [HSD [HnoO [Htmp HtmpL]]]].
-destruct t; [destruct b; auto | idtac];
- (intro; apply Hrec;
- [apply stepf1_dec | apply (dstep_inv S1); auto; apply f2ind'
- | apply Done_notmp3_invf]; auto).
-Qed.
-
-Lemma Done_notmp1_inv:
- forall S1 S2,
- step S1 S2 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In x (getdst (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1 S2 STEP; inversion STEP; (try (simpl; trivial; fail));
- (repeat (rewrite getdst_app; simpl)); intros.
-apply H1; in_tac.
-destruct m; apply H1; in_tac.
-apply H1; in_tac.
-case (Loc.eq x (T r0)); intro.
-rewrite e; unfold T; case (Loc.type r0); simpl; split; red; intro; discriminate.
-apply H1; apply in_cons_noteq with ( b := T r0 ); (try assumption); in_tac.
-apply H3; in_tac.
-Qed.
-
-Lemma Done_notmp1_invpp:
- forall S1 S2,
- stepp S1 S2 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In x (getdst (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1 S2 H H0; (try assumption).
-induction H as [r|r1 r2 r3 H1 H2 Hrec]; auto.
-apply Hrec; auto.
-apply Done_notmp1_inv with r1; auto.
-Qed.
-
-Lemma Done_notmp1_invf:
- forall S1,
- stepInv S1 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In
- x
- (getdst
- (StateToMove (stepf S1) ++ (StateBeing (stepf S1) ++ StateDone (stepf S1)))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1 H H0; (try assumption).
-generalize H; unfold stepInv; intros [Hpath [HSD [HnoO [Hnotmp HnotmpL]]]].
-destruct S1 as [[t1 b1] d1]; set (S1:=(t1, b1, d1)); destruct t1; destruct b1;
- auto; apply (Done_notmp1_invpp S1); auto; apply dstep_step; auto; apply f2ind;
- unfold S1; auto.
-Qed.
-
-Lemma Done_notmp1_res:
- forall S1,
- stepInv S1 ->
- (forall x,
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In
- x
- (getdst
- (StateToMove (Pmov S1) ++ (StateBeing (Pmov S1) ++ StateDone (Pmov S1)))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d]; set (S1:=(t, b, d)).
-intros H; generalize H; intros [Hpath [HSD [HnoO [Htmp HtmpL]]]].
-unfold S1; rewrite Pmov_equation.
-destruct t; [destruct b; auto | idtac];
- (intro; apply Hrec;
- [apply stepf1_dec | apply (dstep_inv S1); auto; apply f2ind'
- | apply Done_notmp1_invf]; auto).
-Qed.
-
-Lemma Done_notmp1src_inv:
- forall S1 S2,
- step S1 S2 ->
- (forall x,
- In x (getsrc (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In x (getsrc (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1 S2 STEP; inversion STEP; (try (simpl; trivial; fail));
- (repeat (rewrite getsrc_app; simpl)); intros.
-apply H1; in_tac.
-destruct m; apply H1; in_tac.
-apply H1; in_tac.
-case (Loc.eq x (T r0)); intro.
-rewrite e; unfold T; case (Loc.type r0); simpl; split; red; intro; discriminate.
-apply H1; apply in_cons_noteq with ( b := T r0 ); (try assumption); in_tac.
-apply H3; in_tac.
-Qed.
-
-Lemma Done_notmp1src_invpp:
- forall S1 S2,
- stepp S1 S2 ->
- (forall x,
- In x (getsrc (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In x (getsrc (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1 S2 H H0; (try assumption).
-induction H as [r|r1 r2 r3 H1 H2 Hrec]; auto.
-apply Hrec; auto.
-apply Done_notmp1src_inv with r1; auto.
-Qed.
-
-Lemma Done_notmp1src_invf:
- forall S1,
- stepInv S1 ->
- (forall x,
- In x (getsrc (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In
- x
- (getsrc
- (StateToMove (stepf S1) ++ (StateBeing (stepf S1) ++ StateDone (stepf S1)))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1 H H0.
-generalize H; unfold stepInv; intros [Hpath [HSD [HnoO [Hnotmp HnotmpL]]]].
-destruct S1 as [[t1 b1] d1]; set (S1:=(t1, b1, d1)); destruct t1; destruct b1;
- auto; apply (Done_notmp1src_invpp S1); auto; apply dstep_step; auto;
- apply f2ind; unfold S1; auto.
-Qed.
-
-Lemma Done_notmp1src_res:
- forall S1,
- stepInv S1 ->
- (forall x,
- In x (getsrc (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1)) ->
- forall x,
- In
- x
- (getsrc
- (StateToMove (Pmov S1) ++ (StateBeing (Pmov S1) ++ StateDone (Pmov S1)))) ->
- Loc.diff x (R IT1) /\ Loc.diff x (R FT1).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d]; set (S1:=(t, b, d)).
-intros H; generalize H; intros [Hpath [HSD [HnoO [Htmp HtmpL]]]].
-unfold S1; rewrite Pmov_equation.
-destruct t; [destruct b; auto | idtac];
- (intro; apply Hrec;
- [apply stepf1_dec | apply (dstep_inv S1); auto; apply f2ind'
- | apply Done_notmp1src_invf]; auto).
-Qed.
-
-Lemma dst_tmp2_step:
- forall S1 S2,
- step S1 S2 ->
- forall x,
- In x (getdst (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- In x temporaries2 \/
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))).
-Proof.
-intros S1 S2 STEP; inversion STEP; (repeat (rewrite getdst_app; simpl)); intros;
- (try (right; in_tac)).
-destruct m; right; in_tac.
-case (Loc.eq x (T r0)); intro.
-rewrite e; unfold T; case (Loc.type r0); left; [left | right; left]; trivial.
-right; apply in_cons_noteq with ( b := T r0 ); auto; in_tac.
-Qed.
-
-Lemma dst_tmp2_stepp:
- forall S1 S2,
- stepp S1 S2 ->
- forall x,
- In x (getdst (StateToMove S2 ++ (StateBeing S2 ++ StateDone S2))) ->
- In x temporaries2 \/
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))).
-Proof.
-intros S1 S2 H.
-induction H as [r|r1 r2 r3 H1 H2 Hrec]; auto.
-intros.
-elim Hrec with ( x := x );
- [intros H0; (try clear Hrec); (try exact H0) | intros H0; (try clear Hrec)
- | try clear Hrec]; auto.
-generalize (dst_tmp2_step r1 r2); auto.
-Qed.
-
-Lemma dst_tmp2_stepf:
- forall S1,
- stepInv S1 ->
- forall x,
- In
- x
- (getdst
- (StateToMove (stepf S1) ++ (StateBeing (stepf S1) ++ StateDone (stepf S1)))) ->
- In x temporaries2 \/
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))).
-Proof.
-intros S1 H H0.
-generalize H; unfold stepInv; intros [Hpath [HSD [HnoO [Hnotmp HnotmpL]]]].
-destruct S1 as [[t1 b1] d1]; set (S1:=(t1, b1, d1)); destruct t1; destruct b1;
- auto; apply (dst_tmp2_stepp S1); auto; apply dstep_step; auto; apply f2ind;
- unfold S1; auto.
-Qed.
-
-Lemma dst_tmp2_res:
- forall S1,
- stepInv S1 ->
- forall x,
- In
- x
- (getdst
- (StateToMove (Pmov S1) ++ (StateBeing (Pmov S1) ++ StateDone (Pmov S1)))) ->
- In x temporaries2 \/
- In x (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1))).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d]; set (S1:=(t, b, d)).
-intros H; generalize H; intros [Hpath [HSD [HnoO [Htmp HtmpL]]]].
-unfold S1; rewrite Pmov_equation; intros.
-destruct t; auto.
-destruct b; auto.
-elim Hrec with ( y := stepf S1 ) ( x := x );
- [intros H1 | intros H1 | try clear Hrec | try clear Hrec | try assumption].
-left; (try assumption).
-apply dst_tmp2_stepf; auto.
-apply stepf1_dec; auto.
-apply (dstep_inv S1); auto; unfold S1; apply f2ind'; auto.
-elim Hrec with ( y := stepf S1 ) ( x := x );
- [intro; left; (try assumption) | intros H1; apply dst_tmp2_stepf; auto |
- apply stepf1_dec; auto |
- apply (dstep_inv S1); auto; unfold S1; apply f2ind'; auto | try assumption].
-Qed.
-
-Lemma getdst_lists2moves:
- forall srcs dsts,
- length srcs = length dsts ->
- getsrc (listsLoc2Moves srcs dsts) = srcs /\
- getdst (listsLoc2Moves srcs dsts) = dsts.
-Proof.
-induction srcs; intros dsts; destruct dsts; simpl; auto; intro;
- (try discriminate).
-inversion H.
-elim IHsrcs with ( dsts := dsts ); auto; intros.
-rewrite H2; rewrite H0; auto.
-Qed.
-
-Lemma stepInv_pnilnil:
- forall p,
- simpleDest p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- no_overlap_list p -> stepInv (p, nil, nil).
-Proof.
-unfold stepInv; simpl; (repeat split); auto.
-rewrite app_nil; auto.
-generalize (no_overlap_noOverlap (p, nil, nil)).
-simpl; (intros H3; apply H3).
-generalize H2; unfold no_overlap_state, no_overlap_list; simpl; intro.
-rewrite app_nil; auto.
-apply disjoint_tmp__noTmp; auto.
-Qed.
-
-Lemma noO_list_pnilnil:
- forall (p : Moves),
- simpleDest p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- no_overlap_list p -> no_overlap_list (StateDone (Pmov (p, nil, nil))).
-Proof.
-intros; generalize (no_overlapD_res (p, nil, nil));
- unfold no_overlap_stateD, no_overlap_list.
-rewrite STM_Pmov; rewrite SB_Pmov; simpl; rewrite app_nil; intro.
-apply H3; auto.
-apply stepInv_pnilnil; auto.
-Qed.
-
-Lemma dis_srctmp1_pnilnil:
- forall (p : Moves),
- simpleDest p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- no_overlap_list p ->
- Loc.disjoint (getsrc (StateDone (Pmov (p, nil, nil)))) temporaries1.
-Proof.
-intros; generalize (Done_notmp1src_res (p, nil, nil)); simpl.
-rewrite STM_Pmov; rewrite SB_Pmov; simpl; rewrite app_nil; intro.
-unfold temporaries1.
-apply Loc.notin_disjoint; auto.
-simpl; intros.
-assert (HsI: stepInv (p, nil, nil)); (try apply stepInv_pnilnil); auto.
-elim H3 with x; (try assumption).
-intros; (repeat split); auto.
-intros; split;
- (apply Loc.in_notin_diff with ( ll := temporaries );
- [apply Loc.disjoint_notin with (getsrc p) | simpl]; auto).
-Qed.
-
-Lemma dis_dsttmp1_pnilnil:
- forall (p : Moves),
- simpleDest p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- no_overlap_list p ->
- Loc.disjoint (getdst (StateDone (Pmov (p, nil, nil)))) temporaries1.
-Proof.
-intros; generalize (Done_notmp1_res (p, nil, nil)); simpl.
-rewrite STM_Pmov; rewrite SB_Pmov; simpl; rewrite app_nil; intro.
-unfold temporaries1.
-apply Loc.notin_disjoint; auto.
-simpl; intros.
-assert (HsI: stepInv (p, nil, nil)); (try apply stepInv_pnilnil); auto.
-elim H3 with x; (try assumption).
-intros; (repeat split); auto.
-intros; split;
- (apply Loc.in_notin_diff with ( ll := temporaries );
- [apply Loc.disjoint_notin with (getdst p) | simpl]; auto).
-Qed.
-
-Lemma parallel_move_correct':
- forall p k rs m,
- Loc.norepet (getdst p) ->
- no_overlap_list p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- (exists rs' ,
- exec_instrs ge sp (p_move p k) rs m k rs' m /\
- (List.map rs' (getdst p) = List.map rs (getsrc p) /\
- (rs' (R IT3) = rs (R IT3) /\
- (forall l,
- Loc.notin l (getdst p) -> Loc.notin l temporaries -> rs' l = rs l)))
- ).
-Proof.
-unfold p_move, P_move.
-intros p k rs m Hnorepet HnoOverlap Hdisjointsrc Hdisjointdst.
-assert (HsD: simpleDest p); (try (apply norepet_SD; assumption)).
-assert (HsI: stepInv (p, nil, nil)); (try (apply stepInv_pnilnil; assumption)).
-generalize HsI; unfold stepInv; simpl StateToMove; simpl StateBeing;
- (repeat rewrite app_nil); intros [_ [_ [HnoO [Hnotmp _]]]].
-elim (exec_instrs_pmov (StateDone (Pmov (p, nil, nil))) k rs m); auto;
- (try apply noO_list_pnilnil); (try apply dis_dsttmp1_pnilnil);
- (try apply dis_srctmp1_pnilnil); auto.
-intros rs' [Hexec R]; exists rs'; (repeat split); auto.
-rewrite <- (Fpmov_correct_map p rs); auto.
-apply list_map_exten; intros; rewrite <- R; auto;
- (try
- (apply Loc.in_notin_diff with ( ll := temporaries );
- [apply Loc.disjoint_notin with (getdst p) | simpl]; auto)).
-generalize (dst_tmp2_res (p, nil, nil)); intros.
-elim H0 with ( x := r );
- [intros H2; right |
- simpl; rewrite app_nil; intros H2; apply In_norepet with (getdst p); auto |
- try assumption | rewrite STM_Pmov; rewrite SB_Pmov; auto].
-apply Loc.diff_sym; apply Loc.in_notin_diff with ( ll := temporaries );
- (try assumption).
-apply Loc.disjoint_notin with (getdst p); auto.
-generalize H2; unfold temporaries, temporaries2; intros; in_tac.
-rewrite <- (Fpmov_correct_IT3 p rs); auto; rewrite <- R;
- (try (simpl; intro; discriminate)); auto.
-intros r H; right; apply (Done_notmp3_res (p, nil, nil)); auto;
- (try (rewrite STM_Pmov; rewrite SB_Pmov; auto)).
-simpl; rewrite app_nil; intros; apply Loc.in_notin_diff with temporaries; auto.
-apply Loc.disjoint_notin with (getdst p); auto.
-simpl; right; right; left; trivial.
-intros; rewrite <- (Fpmov_correct_ext p rs); auto; rewrite <- R; auto;
- (try (apply Loc.in_notin_diff with temporaries; simpl; auto)).
-intros r H1; right; generalize (dst_tmp2_res (p, nil, nil)); intros.
-elim H2 with ( x := r );
- [intros H3 | simpl; rewrite app_nil; intros H3 | assumption
- | rewrite STM_Pmov; rewrite SB_Pmov; auto].
-apply Loc.diff_sym; apply Loc.in_notin_diff with temporaries; auto.
-generalize H3; unfold temporaries, temporaries2; intros; in_tac.
-apply Loc.diff_sym; apply Loc.in_notin_diff with ( ll := getdst p ); auto.
-Qed.
-
-Lemma parallel_move_correctX:
- forall srcs dsts k rs m,
- List.length srcs = List.length dsts ->
- no_overlap srcs dsts ->
- Loc.norepet dsts ->
- Loc.disjoint srcs temporaries ->
- Loc.disjoint dsts temporaries ->
- (exists rs' ,
- exec_instrs ge sp (parallel_move srcs dsts k) rs m k rs' m /\
- (List.map rs' dsts = List.map rs srcs /\
- (rs' (R IT3) = rs (R IT3) /\
- (forall l, Loc.notin l dsts -> Loc.notin l temporaries -> rs' l = rs l))) ).
-Proof.
-intros; unfold parallel_move, parallel_move_order;
- generalize (parallel_move_correct' (listsLoc2Moves srcs dsts) k rs m).
-elim (getdst_lists2moves srcs dsts); auto.
-intros heq1 heq2; rewrite heq1; rewrite heq2; unfold p_move.
-intros H4; apply H4; auto.
-unfold no_overlap_list; rewrite heq1; rewrite heq2; auto.
-Qed.
-
-End parallel_move_correction.
diff --git a/backend/Alloctyping.v b/backend/Alloctyping.v
index 39e53eef..e4f9f964 100644
--- a/backend/Alloctyping.v
+++ b/backend/Alloctyping.v
@@ -15,7 +15,7 @@ Require Import Allocproof.
Require Import RTLtyping.
Require Import LTLtyping.
Require Import Conventions.
-Require Import Alloctyping_aux.
+Require Import Parallelmove.
(** This file proves that register allocation (the translation from
RTL to LTL defined in file [Allocation]) preserves typing:
@@ -30,13 +30,13 @@ Variable live: PMap.t Regset.t.
Variable alloc: reg -> loc.
Variable tf: LTL.function.
-Hypothesis TYPE_RTL: type_rtl_function f = Some env.
+Hypothesis TYPE_RTL: type_function f = Some env.
Hypothesis ALLOC: regalloc f live (live0 f live) env = Some alloc.
Hypothesis TRANSL: transf_function f = Some tf.
-Lemma wt_rtl_function: RTLtyping.wt_function env f.
+Lemma wt_rtl_function: RTLtyping.wt_function f env.
Proof.
- apply type_rtl_function_correct; auto.
+ apply type_function_correct; auto.
Qed.
Lemma alloc_type: forall r, Loc.type (alloc r) = env r.
@@ -213,15 +213,34 @@ Proof.
auto.
Qed.
+Lemma wt_add_moves:
+ forall b moves,
+ (forall s d, In (s, d) moves ->
+ loc_read_ok s /\ loc_write_ok d /\ Loc.type s = Loc.type d) ->
+ wt_block tf b ->
+ wt_block tf
+ (List.fold_right (fun p k => add_move (fst p) (snd p) k) b moves).
+Proof.
+ induction moves; simpl; intros.
+ auto.
+ destruct a as [s d]. simpl.
+ elim (H s d). intros A [B C]. apply wt_add_move; auto. auto.
+Qed.
+
Theorem wt_parallel_move:
forall srcs dsts b,
List.map Loc.type srcs = List.map Loc.type dsts ->
locs_read_ok srcs ->
locs_write_ok dsts -> wt_block tf b -> wt_block tf (parallel_move srcs dsts b).
Proof.
- unfold locs_read_ok, locs_write_ok.
- apply wt_parallel_moveX; simpl; auto.
- exact wt_add_move.
+ intros. unfold parallel_move. apply wt_add_moves; auto.
+ intros.
+ elim (parmove_prop_2 _ _ _ _ H3); intros A B.
+ split. destruct A as [C|[C|C]].
+ apply H0; auto. subst s; exact I. subst s; exact I.
+ split. destruct B as [C|[C|C]].
+ apply H1; auto. subst d; exact I. subst d; exact I.
+ eapply parmove_prop_3; eauto.
Qed.
Lemma wt_add_op_move:
@@ -340,6 +359,18 @@ Proof.
rewrite loc_result_type. auto. constructor.
Qed.
+Lemma wt_add_alloc:
+ forall arg res s,
+ Loc.type arg = Tint -> Loc.type res = Tint ->
+ loc_read_ok arg -> loc_write_ok res ->
+ wt_block tf (add_alloc arg res s).
+Proof.
+ intros.
+ unfold add_alloc. apply wt_add_reload. auto. rewrite H; reflexivity.
+ apply wt_Balloc. apply wt_add_spill. auto. rewrite H0; reflexivity.
+ apply wt_Bgoto.
+Qed.
+
Lemma wt_add_cond:
forall cond args ifso ifnot,
List.map Loc.type args = type_of_condition cond ->
@@ -387,7 +418,10 @@ Lemma wt_add_entry:
Proof.
set (sig := RTL.fn_sig f).
assert (sig = tf.(fn_sig)).
- unfold sig. symmetry. apply Allocproof.sig_function_translated. auto.
+ unfold sig.
+ assert (transf_fundef (Internal f) = Some (Internal tf)).
+ unfold transf_fundef; rewrite TRANSL; auto.
+ generalize (Allocproof.sig_function_translated _ _ H). simpl; auto.
assert (locs_read_ok (loc_parameters sig)).
red; unfold loc_parameters. intros.
generalize (list_in_map_inv _ _ _ H0). intros [l1 [EQ IN]].
@@ -406,7 +440,7 @@ Qed.
Lemma wt_transf_instr:
forall pc instr,
- RTLtyping.wt_instr env f instr ->
+ RTLtyping.wt_instr env f.(RTL.fn_sig) instr ->
wt_block tf (transf_instr f live alloc pc instr).
Proof.
intros. inversion H; simpl.
@@ -441,6 +475,9 @@ Proof.
auto with allocty.
destruct ros; simpl; auto with allocty.
auto with allocty.
+ (* alloc *)
+ apply wt_add_alloc; auto with allocty.
+ rewrite alloc_type; auto. rewrite alloc_type; auto.
(* cond *)
apply wt_add_cond. rewrite alloc_types; auto. auto with allocty.
(* return *)
@@ -483,7 +520,7 @@ Lemma wt_transf_function:
transf_function f = Some tf -> wt_function tf.
Proof.
intros. generalize H; unfold transf_function.
- caseEq (type_rtl_function f). intros env TYP.
+ caseEq (type_function f). intros env TYP.
caseEq (analyze f). intros live ANL.
change (transfer f (RTL.fn_entrypoint f)
live!!(RTL.fn_entrypoint f))
@@ -497,13 +534,26 @@ Proof.
intros; discriminate.
Qed.
+Lemma wt_transf_fundef:
+ forall f tf,
+ transf_fundef f = Some tf -> wt_fundef tf.
+Proof.
+ intros until tf; destruct f; simpl.
+ caseEq (transf_function f). intros g TF EQ. inversion EQ.
+ constructor. eapply wt_transf_function; eauto.
+ congruence.
+ caseEq (type_external_function e); intros.
+ inversion H0. constructor. apply type_external_function_correct. auto.
+ congruence.
+Qed.
+
Lemma program_typing_preserved:
forall (p: RTL.program) (tp: LTL.program),
transf_program p = Some tp ->
LTLtyping.wt_program tp.
Proof.
intros; red; intros.
- generalize (transform_partial_program_function transf_function p i f H H0).
+ generalize (transform_partial_program_function transf_fundef p i f H H0).
intros [f0 [IN TRANSF]].
- apply wt_transf_function with f0; auto.
+ apply wt_transf_fundef with f0; auto.
Qed.
diff --git a/backend/Alloctyping_aux.v b/backend/Alloctyping_aux.v
deleted file mode 100644
index 0013c240..00000000
--- a/backend/Alloctyping_aux.v
+++ /dev/null
@@ -1,895 +0,0 @@
-(** Type preservation for parallel move compilation. *)
-
-(** This file, contributed by Laurence Rideau, shows that the parallel
- move compilation algorithm (file [Parallelmove]) generates a well-typed
- sequence of LTL instructions, provided the original problem is well-typed:
- the types of source and destination locations match pairwise. *)
-
-Require Import Coqlib.
-Require Import Locations.
-Require Import LTL.
-Require Import Allocation.
-Require Import LTLtyping.
-Require Import Allocproof_aux.
-Require Import Parallelmove.
-Require Import Inclusion.
-
-Section wt_move_correction.
-Variable tf : LTL.function.
-Variable loc_read_ok : loc -> Prop.
-Hypothesis loc_read_ok_R : forall r, loc_read_ok (R r).
-Variable loc_write_ok : loc -> Prop.
-Hypothesis loc_write_ok_R : forall r, loc_write_ok (R r).
-
-Let locs_read_ok (ll : list loc) : Prop :=
- forall l, In l ll -> loc_read_ok l.
-
-Let locs_write_ok (ll : list loc) : Prop :=
- forall l, In l ll -> loc_write_ok l.
-
-Hypothesis
- wt_add_move :
- forall (src dst : loc) (b : LTL.block),
- loc_read_ok src ->
- loc_write_ok dst ->
- Loc.type src = Loc.type dst ->
- wt_block tf b -> wt_block tf (add_move src dst b).
-
-Lemma in_move__in_srcdst:
- forall m p, In m p -> In (fst m) (getsrc p) /\ In (snd m) (getdst p).
-Proof.
-intros; induction p.
-inversion H.
-destruct a as [a1 a2]; destruct m as [m1 m2]; simpl.
-elim H; intro.
-inversion H0.
-subst a2; subst a1.
-split; [left; trivial | left; trivial].
-split; right; (elim IHp; simpl; intros; auto).
-Qed.
-
-Lemma T_type: forall r, Loc.type r = Loc.type (T r).
-Proof.
-intro; unfold T.
-case (Loc.type r); auto.
-Qed.
-
-Theorem incl_nil: forall A (l : list A), incl nil l.
-Proof.
-intros A l a; simpl; intros H; case H.
-Qed.
-Hint Resolve incl_nil :datatypes.
-
-Lemma split_move_incl:
- forall (l t1 t2 : Moves) (s d : Reg),
- split_move l s = Some (t1, d, t2) -> incl t1 l /\ incl t2 l.
-Proof.
-induction l.
-simpl; (intros; discriminate).
-intros t1 t2 s d; destruct a as [a1 a2]; simpl.
-case (Loc.eq a1 s); intro.
-intros.
-inversion H.
-split; auto.
-apply incl_nil.
-apply incl_tl; apply incl_refl; auto.
-caseEq (split_move l s); intro; (try (intros; discriminate)).
-destruct p as [[p1 p2] p3].
-intros.
-inversion H0.
-elim (IHl p1 p3 s p2); intros; auto.
-subst p3.
-split; auto.
-generalize H1; unfold incl; simpl.
-intros H4 a [H7|H6]; [try exact H7 | idtac].
-left; (try assumption).
-right; apply H4; auto.
-apply incl_tl; auto.
-Qed.
-
-Lemma in_split_move:
- forall (l t1 t2 : Moves) (s d : Reg),
- split_move l s = Some (t1, d, t2) -> In (s, d) l.
-Proof.
-induction l.
-simpl; intros; discriminate.
-intros t1 t2 s d; simpl.
-destruct a as [a1 a2].
-case (Loc.eq a1 s).
-intros.
-inversion H.
-subst a1; left; auto.
-intro; caseEq (split_move l s); (intros; (try discriminate)).
-destruct p as [[p1 p2] p3].
-right; inversion H0.
-subst p2.
-apply (IHl p1 p3); auto.
-Qed.
-
-Lemma move_types_stepf:
- forall S1,
- (forall x1 x2,
- In (x1, x2) (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1)) ->
- Loc.type x1 = Loc.type x2) ->
- forall x1 x2,
- In
- (x1, x2)
- (StateToMove (stepf S1) ++ (StateBeing (stepf S1) ++ StateDone (stepf S1))) ->
- Loc.type x1 = Loc.type x2.
-Proof.
-intros S1 H x1 x2.
-destruct S1 as [[t1 b1] d1]; set (S1:=(t1, b1, d1)); destruct t1; destruct b1;
- auto; simpl StateToMove in H |-; simpl StateBeing in H |-;
- simpl StateDone in H |-; simpl app in H |-.
-intro;
- elim
- (in_app_or
- (StateToMove (stepf S1)) (StateBeing (stepf S1) ++ StateDone (stepf S1))
- (x1, x2)); auto.
-assert (StateToMove (stepf S1) = nil).
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq d (fst (last b1))); case b1; simpl; auto.
-rewrite H1; intros H2; inversion H2.
-intro; elim (in_app_or (StateBeing (stepf S1)) (StateDone (stepf S1)) (x1, x2));
- auto.
-assert
- (StateBeing (stepf S1) = nil \/
- (StateBeing (stepf S1) = b1 \/ StateBeing (stepf S1) = replace_last_s b1)).
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq d (fst (last b1))); case b1; simpl; auto.
-elim H2; [intros H3; (try clear H2); (try exact H3) | intros H3; (try clear H2)].
-rewrite H3; intros H4; inversion H4.
-elim H3; [intros H2; (try clear H3); (try exact H2) | intros H2; (try clear H3)].
-rewrite H2; intros H4.
-apply H; (try in_tac).
-rewrite H2; intros H4.
-caseEq b1; intro; simpl; auto.
-rewrite H3 in H4; simpl in H4 |-; inversion H4.
-intros l H5; rewrite H5 in H4.
-generalize (app_rewriter _ l m0).
-intros [y [r H3]]; (try exact H3).
-rewrite H3 in H4.
-destruct y.
-rewrite last_replace in H4.
-elim (in_app_or r ((T r0, r1) :: nil) (x1, x2)); auto.
-intro; apply H.
-rewrite H5.
-rewrite H3; in_tac.
-intros H6; inversion H6.
-inversion H7.
-rewrite <- T_type.
-rewrite <- H10; apply H.
-rewrite H5; rewrite H3; (try in_tac).
-assert (In (r0, r1) ((r0, r1) :: nil)); [simpl; auto | in_tac].
-inversion H7.
-intro.
-destruct m as [s d].
-assert
- (StateDone (stepf S1) = (s, d) :: d1 \/
- StateDone (stepf S1) = (s, d) :: ((d, T d) :: d1)).
-simpl.
-case (Loc.eq d (fst (last b1))); case b1; simpl; auto.
-elim H3; [intros H4; (try clear H3); (try exact H4) | intros H4; (try clear H3)].
-apply H; (try in_tac).
-rewrite H4 in H2; in_tac.
-rewrite H4 in H2.
-simpl in H2 |-.
-elim H2; [intros H3; apply H | intros H3; elim H3; intros; [idtac | apply H]];
- (try in_tac).
-simpl; left; auto.
-inversion H5; apply T_type.
-intro;
- elim
- (in_app_or
- (StateToMove (stepf S1)) (StateBeing (stepf S1) ++ StateDone (stepf S1))
- (x1, x2)); auto.
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq s d); simpl; intros; apply H; in_tac.
-intro; elim (in_app_or (StateBeing (stepf S1)) (StateDone (stepf S1)) (x1, x2));
- auto.
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq s d); intros; apply H; (try in_tac).
-inversion H2.
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq s d); intros; apply H; (try in_tac).
-simpl in H2 |-; in_tac.
-simpl in H2 |-; in_tac.
-intro;
- elim
- (in_app_or
- (StateToMove (stepf S1)) (StateBeing (stepf S1) ++ StateDone (stepf S1))
- (x1, x2)); auto.
-simpl stepf.
-destruct m as [s d].
-destruct m0 as [s0 d0].
-case (Loc.eq s d0); [simpl; intros; apply H; in_tac | idtac].
-caseEq (split_move t1 d0); intro.
-destruct p as [[t2 b2] d2].
-intros Hsplit Hd; simpl StateToMove; intro.
-elim (split_move_incl t1 t2 d2 d0 b2 Hsplit); auto.
-intros; apply H.
-assert (In (x1, x2) ((s, d) :: (t1 ++ t1))).
-generalize H1; simpl; intros.
-elim H4; [intros H5; left; (try exact H5) | intros H5; right].
-elim (in_app_or t2 d2 (x1, x2)); auto; intro; apply in_or_app; left.
-unfold incl in H2 |-.
-apply H2; auto.
-unfold incl in H3 |-; apply H3; auto.
-in_tac.
-intro; case (Loc.eq d0 (fst (last b1))); case b1; auto; simpl StateToMove;
- intros; apply H; in_tac.
-intro; elim (in_app_or (StateBeing (stepf S1)) (StateDone (stepf S1)) (x1, x2));
- auto.
-simpl stepf.
-destruct m as [s d].
-destruct m0 as [s0 d0].
-case (Loc.eq s d0).
-intros e; rewrite <- e; simpl StateBeing.
-rewrite <- e in H.
-intro; apply H; in_tac.
-caseEq (split_move t1 d0); intro.
-destruct p as [[t2 b2] d2].
-simpl StateBeing.
-intros.
-apply H.
-generalize (in_split_move t1 t2 d2 d0 b2 H2).
-intros.
-elim H3; intros.
-rewrite <- H5.
-in_tac.
-in_tac.
-caseEq b1.
-simpl; intros e n F; elim F.
-intros m l H3 H4.
-case (Loc.eq d0 (fst (last (m :: l)))).
-generalize (app_rewriter Move l m).
-intros [y [r H5]]; rewrite H5.
-simpl StateBeing.
-destruct y as [y1 y2]; generalize (last_replace r y1 y2).
-simpl; intros heq H6.
-unfold Move in heq |-; unfold Move.
-rewrite heq.
-intro.
-elim (in_app_or r ((T y1, y2) :: nil) (x1, x2)); auto.
-intro; apply H.
-rewrite H3; rewrite H5; in_tac.
-simpl; intros [H8|H8]; inversion H8.
-rewrite <- T_type.
-apply H.
-rewrite H3; rewrite H5.
-rewrite <- H11; assert (In (y1, y2) ((y1, y2) :: nil)); auto.
-simpl; auto.
-in_tac.
-simpl StateBeing; intros.
-apply H; rewrite H3; (try in_tac).
-simpl stepf.
-destruct m as [s d].
-destruct m0 as [s0 d0].
-case (Loc.eq s d0); [simpl; intros; apply H; in_tac | idtac].
-caseEq (split_move t1 d0); intro.
-destruct p as [[t2 b2] d2].
-intros Hsplit Hd; simpl StateDone; intro.
-apply H; (try in_tac).
-case (Loc.eq d0 (fst (last b1))); case b1; simpl StateDone; intros;
- (try (apply H; in_tac)).
-elim H3; intros.
-apply H.
-assert (In (x1, x2) ((s0, d0) :: nil)); auto.
-rewrite H4; auto.
-simpl; left; auto.
-in_tac.
-elim H4; intros.
-inversion H5; apply T_type.
-apply H; in_tac.
-Qed.
-
-Lemma move_types_res:
- forall S1,
- (forall x1 x2,
- In (x1, x2) (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1)) ->
- Loc.type x1 = Loc.type x2) ->
- forall x1 x2,
- In
- (x1, x2)
- (StateToMove (Pmov S1) ++ (StateBeing (Pmov S1) ++ StateDone (Pmov S1))) ->
- Loc.type x1 = Loc.type x2.
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d]; set (S1:=(t, b, d)).
-unfold S1; rewrite Pmov_equation; intros.
-destruct t; auto.
-destruct b; auto.
-apply (Hrec (stepf S1)).
-apply stepf1_dec; auto.
-apply move_types_stepf; auto.
-unfold S1; auto.
-apply (Hrec (stepf S1)).
-apply stepf1_dec; auto.
-apply move_types_stepf; auto.
-unfold S1; auto.
-Qed.
-
-Lemma srcdst_tmp2_stepf:
- forall S1 x1 x2,
- In
- (x1, x2)
- (StateToMove (stepf S1) ++ (StateBeing (stepf S1) ++ StateDone (stepf S1))) ->
- (In x1 temporaries2 \/
- In x1 (getsrc (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1)))) /\
- (In x2 temporaries2 \/
- In x2 (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1)))).
-Proof.
-intros S1 x1 x2 H.
-(repeat rewrite getsrc_app); (repeat rewrite getdst_app).
-destruct S1 as [[t1 b1] d1]; set (S1:=(t1, b1, d1)); destruct t1; destruct b1;
- auto.
-simpl in H |-.
-elim (in_move__in_srcdst (x1, x2) d1); intros; auto.
-elim
- (in_app_or
- (StateToMove (stepf S1)) (StateBeing (stepf S1) ++ StateDone (stepf S1))
- (x1, x2)); auto.
-assert (StateToMove (stepf S1) = nil).
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq d (fst (last b1))); case b1; simpl; auto.
-rewrite H0; intros H2; inversion H2.
-intro; elim (in_app_or (StateBeing (stepf S1)) (StateDone (stepf S1)) (x1, x2));
- auto.
-simpl stepf.
-destruct m as [s d].
-caseEq b1.
-simpl.
-intros h1 h2; inversion h2.
-intros m l heq; generalize (app_rewriter _ l m).
-intros [y [r H3]]; (try exact H3).
-rewrite H3.
-destruct y.
-rewrite last_app; simpl fst.
-case (Loc.eq d r0).
-intros heqd.
-rewrite last_replace.
-simpl.
-intro; elim (in_app_or r ((T r0, r1) :: nil) (x1, x2)); auto.
-rewrite heq; rewrite H3.
-rewrite getsrc_app; simpl; rewrite getdst_app; simpl.
-intro; elim (in_move__in_srcdst (x1, x2) r); auto; simpl; intros; split; right;
- right; in_tac.
-intro.
-inversion H2; inversion H4.
-split.
-unfold T; case (Loc.type r0); left; [left | right]; auto.
-right; right; (try assumption).
-rewrite heq; rewrite H3.
-rewrite H7; simpl.
-rewrite getdst_app; simpl.
-assert (In x2 (x2 :: nil)); simpl; auto.
-in_tac.
-simpl StateBeing.
-intros; elim (in_move__in_srcdst (x1, x2) (r ++ ((r0, r1) :: nil))); auto;
- intros; split; right; right.
-unfold snd in H4 |-; unfold fst in H2 |-; rewrite heq; rewrite H3; (try in_tac).
-unfold snd in H4 |-; unfold fst in H2 |-; rewrite heq; rewrite H3; (try in_tac).
-simpl stepf.
-destruct m as [s d].
-caseEq b1; intro.
-simpl StateDone; intro.
-unfold S1, StateToMove, StateBeing.
-simpl app.
-elim (in_move__in_srcdst (x1, x2) ((s, d) :: d1)); auto; intros; split; right.
-simpl snd in H4 |-; simpl fst in H3 |-; simpl getdst in H4 |-;
- simpl getsrc in H3 |-; (try in_tac).
-simpl snd in H4 |-; simpl fst in H3 |-; simpl getdst in H4 |-;
- simpl getsrc in H3 |-; (try in_tac).
-intros; generalize (app_rewriter _ l m).
-intros [y [r H4]].
-generalize H2; rewrite H4; rewrite last_app.
-destruct y as [y1 y2].
-simpl fst.
-case (Loc.eq d y1).
-simpl StateDone; intros.
-elim H3; [intros H6; inversion H6; (try exact H6) | intros H6; (try clear H5)].
-simpl; split; right; left; auto.
-elim H6; [intros H5; inversion H5; (try exact H5) | intros H5; (try clear H6)].
-split; [right; simpl; right | left].
-rewrite H1; rewrite H4; rewrite getsrc_app; simpl getsrc.
-rewrite <- e; rewrite H8; assert (In x1 (x1 :: nil)); simpl; auto; (try in_tac).
-unfold T; case (Loc.type x1); simpl; auto.
-elim (in_move__in_srcdst (x1, x2) d1); auto; intros; split; right; right;
- (try in_tac).
-intro; simpl StateDone.
-unfold S1, StateToMove, StateBeing, StateDone.
-simpl getsrc; simpl app; (try in_tac).
-intro; elim (in_move__in_srcdst (x1, x2) ((s, d) :: d1));
- (auto; (simpl fst; simpl snd; simpl getsrc; simpl getdst); intros);
- (split; right; (try in_tac)).
-unfold S1, StateToMove, StateBeing, StateDone.
-elim
- (in_app_or
- (StateToMove (stepf S1)) (StateBeing (stepf S1) ++ StateDone (stepf S1))
- (x1, x2)); auto.
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq s d).
-simpl StateToMove.
-intros; elim (in_move__in_srcdst (x1, x2) t1); auto;
- (repeat (rewrite getsrc_app; simpl getsrc));
- (repeat (rewrite getdst_app; simpl getdst)); simpl fst; simpl snd; intros;
- split; right; simpl; right; (try in_tac).
-simpl StateToMove.
-intros; elim (in_move__in_srcdst (x1, x2) t1); auto;
- (repeat (rewrite getsrc_app; simpl getsrc));
- (repeat (rewrite getdst_app; simpl getdst)); simpl fst; simpl snd; intros;
- split; right; simpl; right; (try in_tac).
-intro; elim (in_app_or (StateBeing (stepf S1)) (StateDone (stepf S1)) (x1, x2));
- auto.
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq s d).
-simpl StateBeing; intros h1 h2; inversion h2.
-simpl StateBeing; intros h1 h2.
-elim (in_move__in_srcdst (x1, x2) ((s, d) :: nil)); auto; simpl fst; simpl snd;
- simpl; intros; split; right; (try in_tac).
-elim H1; [intros H3; left; (try exact H3) | intros H3; inversion H3].
-elim H2; [intros H3; left; (try exact H3) | intros H3; inversion H3].
-simpl stepf.
-destruct m as [s d].
-case (Loc.eq s d).
-simpl StateDone; intros h1 h2.
-elim (in_move__in_srcdst (x1, x2) d1); auto; simpl fst; simpl snd; simpl;
- intros; split; right; right; (try in_tac).
-simpl StateDone; intros h1 h2.
-elim (in_move__in_srcdst (x1, x2) d1); auto; simpl fst; simpl snd; simpl;
- intros; split; right; right; (try in_tac).
-elim
- (in_app_or
- (StateToMove (stepf S1)) (StateBeing (stepf S1) ++ StateDone (stepf S1))
- (x1, x2)); auto.
-simpl stepf.
-destruct m as [s d].
-destruct m0 as [s0 d0].
-case (Loc.eq s d0).
-unfold S1, StateToMove, StateBeing, StateDone.
-simpl app at 1.
-intros; elim (in_move__in_srcdst (x1, x2) t1);
- (auto; simpl; intros; (split; right; right; (try in_tac))).
-intro; caseEq (split_move t1 d0); intro.
-destruct p as [[t2 b2] d2].
-intros Hsplit; unfold S1, StateToMove, StateBeing, StateDone; intro.
-elim (split_move_incl t1 t2 d2 d0 b2 Hsplit); auto.
-intros.
-assert (In (x1, x2) ((s, d) :: (t1 ++ t1))).
-generalize H0; simpl; intros.
-elim H3; [intros H5; left; (try exact H5) | intros H5; right].
-elim (in_app_or t2 d2 (x1, x2)); auto; intro; apply in_or_app; left.
-unfold incl in H1 |-.
-apply H1; auto.
-unfold incl in H2 |-; apply H2; auto.
-split; right.
-elim (in_move__in_srcdst (x1, x2) ((s, d) :: (t1 ++ t1)));
- (auto; simpl; intros; (try in_tac)).
-elim H4; [intros H6; (try clear H4); (try exact H6) | intros H6; (try clear H4)].
-left; (try assumption).
-right; (try in_tac).
-rewrite getsrc_app in H6; (try in_tac).
-elim (in_move__in_srcdst (x1, x2) ((s, d) :: (t1 ++ t1)));
- (auto; simpl; intros; (try in_tac)).
-elim H5; [intros H6; (try clear H5); (try exact H6) | intros H6; (try clear H5)].
-left; (try assumption).
-right; rewrite getdst_app in H6; (try in_tac).
-caseEq b1; intro.
-unfold S1, StateToMove, StateBeing, StateDone.
-intro; elim (in_move__in_srcdst (x1, x2) ((s, d) :: t1)); (auto; intros).
-simpl snd in H4 |-; simpl fst in H3 |-; split; right; (try in_tac).
-intros l heq; generalize (app_rewriter _ l m).
-intros [y [r H1]]; rewrite H1.
-destruct y as [y1 y2].
-rewrite last_app; simpl fst.
-case (Loc.eq d0 y1).
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) ((s, d) :: t1)); auto; intros.
-simpl snd in H4 |-; simpl fst in H3 |-; (split; right; (try in_tac)).
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) ((s, d) :: t1)); auto; intros.
-simpl snd in H4 |-; simpl fst in H3 |-; (split; right; (try in_tac)).
-intro; elim (in_app_or (StateBeing (stepf S1)) (StateDone (stepf S1)) (x1, x2));
- auto.
-simpl stepf.
-destruct m as [s d].
-destruct m0 as [s0 d0].
-case (Loc.eq s d0).
-intros e; rewrite <- e; simpl StateBeing.
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) ((s, d) :: ((s0, s) :: b1))); auto;
- simpl; intros.
-split; right; (try in_tac).
-elim H2; [intros H4; left; (try exact H4) | intros H4; (try clear H2)].
-elim H4; [intros H2; right; (try exact H2) | intros H2; (try clear H4)].
-assert (In x1 (s0 :: nil)); auto; (try in_tac).
-simpl; auto.
-right; (try in_tac).
-elim H3; [intros H4; left; (try exact H4) | intros H4; (try clear H3)].
-elim H4; [intros H3; right; (try exact H3) | intros H3; (try clear H4)].
-rewrite <- e; (try in_tac).
-assert (In x2 (s :: nil)); [simpl; auto | try in_tac].
-right; (try in_tac).
-intro; caseEq (split_move t1 d0); intro.
-destruct p as [[t2 b2] d2].
-simpl StateBeing.
-intros.
-generalize (in_split_move t1 t2 d2 d0 b2 H1).
-intros.
-split; right; elim H2; intros.
-rewrite H4 in H3; elim (in_move__in_srcdst (x1, x2) t1); auto; intros.
-simpl snd in H6 |-; simpl fst in H5 |-; (try in_tac).
-unfold S1, StateToMove, StateBeing, StateDone.
-simpl getsrc; (try in_tac).
-elim (in_move__in_srcdst (x1, x2) ((s0, d0) :: b1)); (auto; intros).
-simpl snd in H6 |-; simpl fst in H5 |-; (try in_tac).
-unfold S1, StateToMove, StateBeing, StateDone.
-simpl.
-simpl in H5 |-.
-elim H5; [intros H7; (try clear H5); (try exact H7) | intros H7; (try clear H5)].
-assert (In x1 (s0 :: nil)); simpl; auto.
-right; in_tac.
-right; in_tac.
-inversion H4.
-simpl.
-subst b2.
-rewrite H4 in H3.
-elim (in_move__in_srcdst (x1, x2) t1); (auto; intros).
-simpl snd in H7 |-.
-right; in_tac.
-unfold S1, StateToMove, StateBeing, StateDone.
-elim (in_move__in_srcdst (x1, x2) ((s0, d0) :: b1)); auto; intros.
-simpl snd in H6 |-; (try in_tac).
-apply
- (in_or_app (getdst ((s, d) :: t1)) (getdst ((s0, d0) :: b1) ++ getdst d1) x2);
- right; (try in_tac).
-caseEq b1.
-intros h1 h2; inversion h2.
-intros m l heq.
-generalize (app_rewriter _ l m); intros [y [r H2]]; rewrite H2.
-destruct y as [y1 y2].
-rewrite last_app; simpl fst.
-case (Loc.eq d0 y1).
-unfold S1, StateToMove, StateBeing, StateDone.
-generalize (last_replace r y1 y2).
-unfold Move; intros H3 H6.
-rewrite H3.
-intro.
-elim (in_app_or r ((T y1, y2) :: nil) (x1, x2)); auto.
-intro.
-rewrite heq; rewrite H2; (split; right).
-elim (in_move__in_srcdst (x1, x2) r); auto; simpl fst; simpl snd; intros;
- (try in_tac).
-simpl.
-rewrite getsrc_app; (right; (try in_tac)).
-elim (in_move__in_srcdst (x1, x2) r); auto; simpl fst; simpl snd; intros;
- (try in_tac).
-simpl.
-rewrite getdst_app; right; (try in_tac).
-intros h; inversion h; inversion H5.
-split; [left; simpl; auto | right].
-unfold T; case (Loc.type y1); auto.
-subst y2.
-rewrite heq; rewrite H2.
-simpl.
-rewrite getdst_app; simpl.
-assert (In x2 (x2 :: nil)); [simpl; auto | right; (try in_tac)].
-unfold S1, StateToMove, StateBeing, StateDone.
-intro; rewrite heq; rewrite H2; (split; right).
-intros; elim (in_move__in_srcdst (x1, x2) (r ++ ((y1, y2) :: nil))); auto;
- intros.
-simpl snd in H5 |-; simpl fst in H4 |-.
-simpl.
-right; (try in_tac).
-apply in_or_app; right; simpl; right; (try in_tac).
-elim (in_move__in_srcdst (x1, x2) (r ++ ((y1, y2) :: nil))); auto; intros.
-simpl snd in H5 |-.
-simpl.
-right; (try in_tac).
-apply in_or_app; right; simpl; right; (try in_tac).
-simpl stepf.
-destruct m as [s d].
-destruct m0 as [s0 d0].
-case (Loc.eq s d0).
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) d1); auto; intros.
-simpl in H3 |-; simpl in H2 |-.
-split; right; (try in_tac).
-intro; caseEq (split_move t1 d0); intro.
-destruct p as [[t2 b2] d2].
-simpl StateDone.
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) d1); auto; intros.
-simpl in H3 |-; simpl in H4 |-.
-split; right; (try in_tac).
-caseEq b1.
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) ((s0, d0) :: d1)); auto; intros.
-simpl in H5 |-; simpl in H4 |-; split; right; (try in_tac).
-simpl.
-elim H4; [intros H6; right; (try exact H6) | intros H6; (try clear H4)].
-assert (In x1 (x1 :: nil)); [simpl; auto | rewrite H6; (try in_tac)].
-right; (try in_tac).
-elim H5; [intros H6; right; simpl; (try exact H6) | intros H6; (try clear H5)].
-assert (In x2 (x2 :: nil)); [simpl; auto | rewrite H6; (try in_tac)].
-try in_tac.
-intros m l heq.
-generalize (app_rewriter _ l m); intros [y [r H2]]; rewrite H2.
-destruct y as [y1 y2].
-rewrite last_app; simpl fst.
-case (Loc.eq d0 y1).
-unfold S1, StateToMove, StateBeing, StateDone.
-unfold S1, StateToMove, StateBeing, StateDone.
-intros.
-elim H3; intros.
-inversion H4.
-simpl; split; right; auto.
-right; apply in_or_app; right; simpl; auto.
-right; apply in_or_app; right; simpl; auto.
-elim H4; intros.
-inversion H5.
-simpl; split; [right | left].
-rewrite heq; rewrite H2; simpl.
-rewrite <- e; rewrite H7.
-rewrite getsrc_app; simpl.
-right; assert (In x1 (x1 :: nil)); [simpl; auto | try in_tac].
-unfold T; case (Loc.type x1); auto.
-elim (in_move__in_srcdst (x1, x2) d1); (auto; intros).
-simpl snd in H7 |-; simpl fst in H6 |-; split; right; (try in_tac).
-unfold S1, StateToMove, StateBeing, StateDone.
-intros; elim (in_move__in_srcdst (x1, x2) ((s0, d0) :: d1));
- (auto; simpl; intros).
-split; right.
-elim H4; [intros H6; right; (try exact H6) | intros H6; (try clear H4)].
-apply in_or_app; right; simpl; auto.
-right; (try in_tac).
-elim H5; [intros H6; right; (try exact H6) | intros H6; (try clear H5)].
-apply in_or_app; right; simpl; auto.
-right; (try in_tac).
-Qed.
-
-Lemma getsrc_f: forall s l, In s (getsrc l) -> (exists d , In (s, d) l ).
-Proof.
-induction l; simpl getsrc.
-simpl; (intros h; elim h).
-intros; destruct a as [a1 a2].
-simpl in H |-.
-elim H; [intros H0; (try clear H); (try exact H0) | intros H0; (try clear H)].
-subst a1.
-exists a2; simpl; auto.
-simpl.
-elim IHl; [intros d H; (try clear IHl); (try exact H) | idtac]; auto.
-exists d; [right; (try assumption)].
-Qed.
-
-Lemma incl_src: forall l1 l2, incl l1 l2 -> incl (getsrc l1) (getsrc l2).
-Proof.
-intros.
-unfold incl in H |-.
-unfold incl.
-intros a H0; (try assumption).
-generalize (getsrc_f a).
-intros H1; elim H1 with ( l := l1 );
- [intros d H2; (try clear H1); (try exact H2) | idtac]; auto.
-assert (In (a, d) l2).
-apply H; auto.
-elim (in_move__in_srcdst (a, d) l2); auto.
-Qed.
-
-Lemma getdst_f: forall d l, In d (getdst l) -> (exists s , In (s, d) l ).
-Proof.
-induction l; simpl getdst.
-simpl; (intros h; elim h).
-intros; destruct a as [a1 a2].
-simpl in H |-.
-elim H; [intros H0; (try clear H); (try exact H0) | intros H0; (try clear H)].
-subst a2.
-exists a1; simpl; auto.
-simpl.
-elim IHl; [intros s H; (try clear IHl); (try exact H) | idtac]; auto.
-exists s; [right; (try assumption)].
-Qed.
-
-Lemma incl_dst: forall l1 l2, incl l1 l2 -> incl (getdst l1) (getdst l2).
-Proof.
-intros.
-unfold incl in H |-.
-unfold incl.
-intros a H0; (try assumption).
-generalize (getdst_f a).
-intros H1; elim H1 with ( l := l1 );
- [intros d H2; (try clear H1); (try exact H2) | idtac]; auto.
-assert (In (d, a) l2).
-apply H; auto.
-elim (in_move__in_srcdst (d, a) l2); auto.
-Qed.
-
-Lemma src_tmp2_res:
- forall S1 x1 x2,
- In
- (x1, x2)
- (StateToMove (Pmov S1) ++ (StateBeing (Pmov S1) ++ StateDone (Pmov S1))) ->
- (In x1 temporaries2 \/
- In x1 (getsrc (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1)))) /\
- (In x2 temporaries2 \/
- In x2 (getdst (StateToMove S1 ++ (StateBeing S1 ++ StateDone S1)))).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1 Hrec.
-destruct S1 as [[t b] d]; set (S1:=(t, b, d)).
-unfold S1; rewrite Pmov_equation; intros.
-destruct t.
-destruct b.
-apply srcdst_tmp2_stepf; auto.
-elim Hrec with ( y := stepf S1 ) ( x1 := x1 ) ( x2 := x2 );
- [idtac | apply stepf1_dec; auto | auto].
-intros.
-elim H1; [intros H2; (try clear H1); (try exact H2) | intros H2; (try clear H1)].
-elim H0; [intros H1; (try clear H0); (try exact H1) | intros H1; (try clear H0)].
-split; [left; (try assumption) | idtac].
-left; (try assumption).
-elim (getsrc_f x1) with ( 1 := H1 ); intros x3 H3.
-split; auto.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-elim H0; [intros H1; (try clear H0); (try exact H1) | intros H1; (try clear H0)].
-elim (getdst_f x2) with ( 1 := H2 ); intros x3 H3.
-split; auto.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-elim (getsrc_f x1) with ( 1 := H1 ); intros x3 H3.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-clear H3.
-elim (getdst_f x2) with ( 1 := H2 ); intros x4 H3.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-elim Hrec with ( y := stepf S1 ) ( x1 := x1 ) ( x2 := x2 );
- [idtac | apply stepf1_dec; auto | auto].
-intros.
-elim H1; [intros H2; (try clear H1); (try exact H2) | intros H2; (try clear H1)].
-elim H0; [intros H1; (try clear H0); (try exact H1) | intros H1; (try clear H0)].
-split; [left; (try assumption) | idtac].
-left; (try assumption).
-elim (getsrc_f x1) with ( 1 := H1 ); intros x3 H3.
-split; auto.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-elim H0; [intros H1; (try clear H0); (try exact H1) | intros H1; (try clear H0)].
-elim (getdst_f x2) with ( 1 := H2 ); intros x3 H3.
-split; auto.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-elim (getsrc_f x1) with ( 1 := H1 ); intros x3 H3.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-clear H3.
-elim (getdst_f x2) with ( 1 := H2 ); intros x4 H3.
-elim srcdst_tmp2_stepf with ( 1 := H3 ); auto.
-Qed.
-
-Lemma wt_add_moves:
- forall p b,
- List.map Loc.type (getsrc p) = List.map Loc.type (getdst p) ->
- locs_read_ok (getsrc p) ->
- locs_write_ok (getdst p) ->
- wt_block tf b ->
- wt_block
- tf
- (fold_left
- (fun (k0 : LTL.block) =>
- fun (p0 : loc * loc) => add_move (fst p0) (snd p0) k0) p b).
-Proof.
-induction p.
-intros; simpl; auto.
-intros; destruct a as [a1 a2]; simpl.
-apply IHp; auto.
-inversion H; auto.
-simpl in H0 |-.
-unfold locs_read_ok in H0 |-.
-simpl in H0 |-.
-unfold locs_read_ok; auto.
-generalize H1; unfold locs_write_ok; simpl; auto.
-apply wt_add_move; (try assumption).
-simpl in H0 |-.
-unfold locs_read_ok in H0 |-.
-apply H0.
-simpl; left; trivial.
-unfold locs_write_ok in H1 |-; apply H1.
-simpl; left; trivial.
-inversion H; auto.
-Qed.
-
-Lemma map_f_getsrc_getdst:
- forall (b : Set) (f : Reg -> b) p,
- map f (getsrc p) = map f (getdst p) ->
- forall x1 x2, In (x1, x2) p -> f x1 = f x2.
-Proof.
-intros b f0 p; induction p; simpl; auto.
-intros; contradiction.
-destruct a.
-simpl.
-intros heq; injection heq.
-intros h1 h2.
-intros x1 x2 [H3|H3].
-injection H3.
-intros; subst; auto.
-apply IHp; auto.
-Qed.
-
-Lemma wt_parallel_move':
- forall p b,
- List.map Loc.type (getsrc p) = List.map Loc.type (getdst p) ->
- locs_read_ok (getsrc p) ->
- locs_write_ok (getdst p) -> wt_block tf b -> wt_block tf (p_move p b).
-Proof.
-unfold p_move.
-unfold P_move.
-intros; apply wt_add_moves; auto.
-rewrite getsrc_map; rewrite getdst_map.
-rewrite list_map_compose.
-rewrite list_map_compose.
-apply list_map_exten.
-generalize (move_types_res (p, nil, nil)); auto.
-destruct x as [x1 x2]; simpl; intros; auto.
-symmetry; apply H3.
-simpl.
-rewrite app_nil.
-apply map_f_getsrc_getdst; auto.
-in_tac.
-unfold locs_read_ok.
-intros l H3.
-elim getsrc_f with ( 1 := H3 ); intros x3 H4.
-elim (src_tmp2_res (p, nil, nil) l x3).
-simpl.
-rewrite app_nil.
-intros [[H'|[H'|H']]|H'] _.
-subst l; hnf; auto.
-subst l; hnf; auto.
-contradiction.
-apply H0; auto.
-in_tac.
-intros l H3.
-elim getdst_f with ( 1 := H3 ); intros x3 H4.
-elim (src_tmp2_res (p, nil, nil) x3 l).
-simpl.
-rewrite app_nil.
-intros _ [[H'|[H'|H']]|H'].
-subst l; hnf; auto.
-subst l; hnf; auto.
-contradiction.
-apply H1; auto.
-in_tac.
-Qed.
-
-Theorem wt_parallel_moveX:
- forall srcs dsts b,
- List.map Loc.type srcs = List.map Loc.type dsts ->
- locs_read_ok srcs ->
- locs_write_ok dsts -> wt_block tf b -> wt_block tf (parallel_move srcs dsts b).
-Proof.
-unfold parallel_move, parallel_move_order, P_move.
-intros.
-generalize (wt_parallel_move' (listsLoc2Moves srcs dsts)); intros H'.
-unfold p_move, P_move in H' |-.
-apply H'; auto.
-elim (getdst_lists2moves srcs dsts); auto.
-unfold Allocation.listsLoc2Moves, listsLoc2Moves.
-intros heq1 heq2; rewrite heq1; rewrite heq2; auto.
-repeat rewrite <- (list_length_map Loc.type).
-rewrite H; auto.
-elim (getdst_lists2moves srcs dsts); auto.
-unfold Allocation.listsLoc2Moves, listsLoc2Moves.
-intros heq1 heq2; rewrite heq1; auto.
-repeat rewrite <- (list_length_map Loc.type).
-rewrite H; auto.
-elim (getdst_lists2moves srcs dsts); auto.
-unfold Allocation.listsLoc2Moves, listsLoc2Moves.
-intros heq1 heq2; rewrite heq2; auto.
-repeat rewrite <- (list_length_map Loc.type).
-rewrite H; auto.
-Qed.
-
-End wt_move_correction.
diff --git a/backend/CSE.v b/backend/CSE.v
index 243f6dd7..68010133 100644
--- a/backend/CSE.v
+++ b/backend/CSE.v
@@ -178,6 +178,15 @@ Definition add_load (n: numbering) (rd: reg)
let (n1, vs) := valnum_regs n rs in
add_rhs n1 rd (Load chunk addr vs).
+(** [add_unknown n rd] returns a numbering where [rd] is mapped to a
+ fresh value number, and no equations are added. This is useful
+ to model instructions with unpredictable results such as [Ialloc]. *)
+
+Definition add_unknown (n: numbering) (rd: reg) :=
+ mknumbering (Psucc n.(num_next))
+ n.(num_eqs)
+ (PTree.set rd n.(num_next) n.(num_reg)).
+
(** [kill_load n] removes all equations involving memory loads.
It is used to reflect the effect of a memory store, which can
potentially invalidate all such equations. *)
@@ -328,6 +337,8 @@ Definition transfer (f: function) (pc: node) (before: numbering) :=
kill_loads before
| Icall sig ros args res s =>
empty_numbering
+ | Ialloc arg res s =>
+ add_unknown before res
| Icond cond args ifso ifnot =>
before
| Ireturn optarg =>
@@ -415,6 +426,8 @@ Definition transf_function (f: function) : function :=
f.(fn_nextpc)
(transf_code_wf f approxs f.(fn_code_wf)).
+Definition transf_fundef := AST.transf_fundef transf_function.
+
Definition transf_program (p: program) : program :=
- transform_program transf_function p.
+ transform_program transf_fundef p.
diff --git a/backend/CSEproof.v b/backend/CSEproof.v
index db8a973b..4420269e 100644
--- a/backend/CSEproof.v
+++ b/backend/CSEproof.v
@@ -7,6 +7,7 @@ Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Registers.
@@ -177,6 +178,18 @@ Proof.
apply wf_add_rhs; auto.
Qed.
+Lemma wf_add_unknown:
+ forall n rd,
+ wf_numbering n ->
+ wf_numbering (add_unknown n rd).
+Proof.
+ intros. inversion H. unfold add_unknown. constructor; simpl.
+ intros. eapply wf_equation_increasing; eauto. auto with coqlib.
+ intros until v. rewrite PTree.gsspec. case (peq r rd); intros.
+ inversion H2. auto with coqlib.
+ apply Plt_trans_succ. eauto.
+Qed.
+
Lemma kill_load_eqs_incl:
forall eqs, List.incl (kill_load_eqs eqs) eqs.
Proof.
@@ -205,6 +218,7 @@ Proof.
apply wf_add_load; auto.
apply wf_kill_loads; auto.
apply wf_empty_numbering.
+ apply wf_add_unknown; auto.
Qed.
(** As a consequence, the numberings computed by the static analysis
@@ -497,6 +511,47 @@ Proof.
simpl. exists a; split; congruence.
Qed.
+(** [add_unknown] returns a numbering that is satisfiable in the
+ register set after setting the target register to any value. *)
+
+Lemma add_unknown_satisfiable:
+ forall n rs dst v,
+ wf_numbering n ->
+ numbering_satisfiable ge sp rs m n ->
+ numbering_satisfiable ge sp
+ (rs#dst <- v) m (add_unknown n dst).
+Proof.
+ intros. destruct H0 as [valu A].
+ set (valu' := VMap.set n.(num_next) v valu).
+ assert (numbering_holds valu' ge sp rs m n).
+ eapply numbering_holds_exten; eauto.
+ unfold valu'; red; intros. apply VMap.gso. auto with coqlib.
+ destruct H0 as [B C].
+ exists valu'; split; simpl; intros.
+ eauto.
+ rewrite PTree.gsspec in H0. rewrite Regmap.gsspec.
+ destruct (peq r dst). inversion H0. unfold valu'. rewrite VMap.gss; auto.
+ eauto.
+Qed.
+
+(** Allocation of a fresh memory block preserves satisfiability. *)
+
+Lemma alloc_satisfiable:
+ forall lo hi b m' rs n,
+ Mem.alloc m lo hi = (m', b) ->
+ numbering_satisfiable ge sp rs m n ->
+ numbering_satisfiable ge sp rs m' n.
+Proof.
+ intros. destruct H0 as [valu [A B]].
+ exists valu; split; intros.
+ generalize (A _ _ H0). destruct rh; simpl.
+ auto.
+ intros [addr [C D]]. exists addr; split. auto.
+ destruct addr; simpl in *; try discriminate.
+ eapply Mem.load_alloc_other; eauto.
+ eauto.
+Qed.
+
(** [kill_load] preserves satisfiability. Moreover, the resulting numbering
is satisfiable in any concrete memory state. *)
@@ -612,8 +667,8 @@ End SATISFIABILITY.
(** The transfer function preserves satisfiability of numberings. *)
Lemma transfer_correct:
- forall ge c sp pc rs m pc' rs' m' f n,
- exec_instr ge c sp pc rs m pc' rs' m' ->
+ forall ge c sp pc rs m t pc' rs' m' f n,
+ exec_instr ge c sp pc rs m t pc' rs' m' ->
c = f.(fn_code) ->
wf_numbering n ->
numbering_satisfiable ge sp rs m n ->
@@ -628,6 +683,8 @@ Proof.
eapply kill_load_satisfiable; eauto.
(* Icall *)
apply empty_numbering_satisfiable.
+ (* Ialloc *)
+ apply add_unknown_satisfiable; auto. eapply alloc_satisfiable; eauto.
Qed.
(** The numberings associated to each instruction by the static analysis
@@ -637,8 +694,8 @@ Qed.
satisfiability at [pc']. *)
Theorem analysis_correct_1:
- forall ge c sp pc rs m pc' rs' m' f,
- exec_instr ge c sp pc rs m pc' rs' m' ->
+ forall ge c sp pc rs m t pc' rs' m' f,
+ exec_instr ge c sp pc rs m t pc' rs' m' ->
c = f.(fn_code) ->
numbering_satisfiable ge sp rs m (analyze f)!!pc ->
numbering_satisfiable ge sp rs' m' (analyze f)!!pc'.
@@ -655,15 +712,15 @@ Proof.
eapply Solver.fixpoint_solution; eauto.
elim (fn_code_wf f pc); intro. auto.
rewrite <- CODE in H0.
- elim (exec_instr_present _ _ _ _ _ _ _ _ _ EXEC H0).
+ elim (exec_instr_present _ _ _ _ _ _ _ _ _ _ EXEC H0).
rewrite CODE in EXEC. eapply successors_correct; eauto.
apply H0. auto.
intros. rewrite PMap.gi. apply empty_numbering_satisfiable.
Qed.
Theorem analysis_correct_N:
- forall ge c sp pc rs m pc' rs' m' f,
- exec_instrs ge c sp pc rs m pc' rs' m' ->
+ forall ge c sp pc rs m t pc' rs' m' f,
+ exec_instrs ge c sp pc rs m t pc' rs' m' ->
c = f.(fn_code) ->
numbering_satisfiable ge sp rs m (analyze f)!!pc ->
numbering_satisfiable ge sp rs' m' (analyze f)!!pc'.
@@ -702,19 +759,26 @@ Let tge := Genv.globalenv tprog.
Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof (Genv.find_symbol_transf transf_function prog).
+Proof (Genv.find_symbol_transf transf_fundef prog).
Lemma functions_translated:
- forall (v: val) (f: RTL.function),
+ forall (v: val) (f: RTL.fundef),
Genv.find_funct ge v = Some f ->
- Genv.find_funct tge v = Some (transf_function f).
-Proof (@Genv.find_funct_transf _ _ transf_function prog).
+ Genv.find_funct tge v = Some (transf_fundef f).
+Proof (@Genv.find_funct_transf _ _ transf_fundef prog).
Lemma funct_ptr_translated:
- forall (b: block) (f: RTL.function),
+ forall (b: block) (f: RTL.fundef),
Genv.find_funct_ptr ge b = Some f ->
- Genv.find_funct_ptr tge b = Some (transf_function f).
-Proof (@Genv.find_funct_ptr_transf _ _ transf_function prog).
+ Genv.find_funct_ptr tge b = Some (transf_fundef f).
+Proof (@Genv.find_funct_ptr_transf _ _ transf_fundef prog).
+
+Lemma sig_translated:
+ forall (f: RTL.fundef),
+ funsig (transf_fundef f) = funsig f.
+Proof.
+ intros; case f; intros; reflexivity.
+Qed.
(** The proof of semantic preservation is a simulation argument using
diagrams of the following form:
@@ -732,26 +796,26 @@ Proof (@Genv.find_funct_ptr_transf _ _ transf_function prog).
Definition exec_instr_prop
(c: code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
forall f
(CF: c = f.(RTL.fn_code))
(SAT: numbering_satisfiable ge sp rs m (analyze f)!!pc),
- exec_instr tge (transf_code (analyze f) c) sp pc rs m pc' rs' m'.
+ exec_instr tge (transf_code (analyze f) c) sp pc rs m t pc' rs' m'.
Definition exec_instrs_prop
(c: code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
forall f
(CF: c = f.(RTL.fn_code))
(SAT: numbering_satisfiable ge sp rs m (analyze f)!!pc),
- exec_instrs tge (transf_code (analyze f) c) sp pc rs m pc' rs' m'.
+ exec_instrs tge (transf_code (analyze f) c) sp pc rs m t pc' rs' m'.
Definition exec_function_prop
- (f: RTL.function) (args: list val) (m: mem)
+ (f: RTL.fundef) (args: list val) (m: mem) (t: trace)
(res: val) (m': mem) : Prop :=
- exec_function tge (transf_function f) args m res m'.
+ exec_function tge (transf_fundef f) args m t res m'.
Ltac TransfInstr :=
match goal with
@@ -766,9 +830,9 @@ Ltac TransfInstr :=
derivation for the source program. *)
Lemma transf_function_correct:
- forall f args m res m',
- exec_function ge f args m res m' ->
- exec_function_prop f args m res m'.
+ forall f args m t res m',
+ exec_function ge f args m t res m' ->
+ exec_function_prop f args m t res m'.
Proof.
apply (exec_function_ind_3 ge
exec_instr_prop exec_instrs_prop exec_function_prop);
@@ -804,12 +868,15 @@ Proof.
rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
intro; eapply exec_Istore; eauto.
(* Icall *)
- assert (find_function tge ros rs = Some (transf_function f)).
+ assert (find_function tge ros rs = Some (transf_fundef f)).
destruct ros; simpl in H0; simpl.
apply functions_translated; auto.
rewrite symbols_preserved. destruct (Genv.find_symbol ge i).
apply funct_ptr_translated; auto. discriminate.
- intro; eapply exec_Icall with (f := transf_function f); eauto.
+ intro; eapply exec_Icall with (f := transf_fundef f); eauto.
+ generalize (sig_translated f); congruence.
+ (* Ialloc *)
+ intro; eapply exec_Ialloc; eauto.
(* Icond true *)
intro; eapply exec_Icond_true; eauto.
(* Icond false *)
@@ -821,22 +888,23 @@ Proof.
(* trans *)
eapply exec_trans; eauto. apply H2; auto.
eapply analysis_correct_N; eauto.
- (* function *)
- intro. unfold transf_function; eapply exec_funct; simpl; eauto.
+ (* internal function *)
+ intro. unfold transf_function; simpl; eapply exec_funct_internal; simpl; eauto.
eapply H1; eauto. eapply analysis_correct_entry; eauto.
+ (* external function *)
+ unfold transf_function; simpl. apply exec_funct_external; auto.
Qed.
Theorem transf_program_correct:
- forall (r: val),
- exec_program prog r -> exec_program tprog r.
+ forall (t: trace) (r: val),
+ exec_program prog t r -> exec_program tprog t r.
Proof.
- intros r [fptr [f [m [FINDS [FINDF [SIG EXEC]]]]]].
- red. exists fptr; exists (transf_function f); exists m.
+ intros t r [fptr [f [m [FINDS [FINDF [SIG EXEC]]]]]].
+ red. exists fptr; exists (transf_fundef f); exists m.
split. change (prog_main tprog) with (prog_main prog).
rewrite symbols_preserved. assumption.
split. apply funct_ptr_translated; auto.
- split. unfold transf_function.
- rewrite <- SIG. destruct (analyze f); reflexivity.
+ split. generalize (sig_translated f); congruence.
apply transf_function_correct.
unfold tprog, transf_program. rewrite Genv.init_mem_transf.
exact EXEC.
diff --git a/backend/Cmconstr.v b/backend/Cmconstr.v
index e6168d1a..f3a63fae 100644
--- a/backend/Cmconstr.v
+++ b/backend/Cmconstr.v
@@ -57,6 +57,8 @@ Fixpoint lift_expr (p: nat) (a: expr) {struct a}: expr :=
| Elet b c => Elet (lift_expr p b) (lift_expr (S p) c)
| Eletvar n =>
if le_gt_dec p n then Eletvar (S n) else Eletvar n
+ | Ealloc b =>
+ Ealloc (lift_expr p b)
end
with lift_condexpr (p: nat) (a: condexpr) {struct a}: condexpr :=
@@ -795,11 +797,6 @@ Definition divf (e1 e2: expr) := Eop Odivf (e1:::e2:::Enil).
(** ** Truncations and sign extensions *)
-Definition cast8unsigned (e: expr) :=
- rolm e Int.zero (Int.repr 255).
-Definition cast16unsigned (e: expr) :=
- rolm e Int.zero (Int.repr 65535).
-
Inductive cast8signed_cases: forall (e1: expr), Set :=
| cast8signed_case1:
forall (e2: expr),
@@ -822,6 +819,28 @@ Definition cast8signed (e: expr) :=
| cast8signed_default e1 => Eop Ocast8signed (e1 ::: Enil)
end.
+Inductive cast8unsigned_cases: forall (e1: expr), Set :=
+ | cast8unsigned_case1:
+ forall (e2: expr),
+ cast8unsigned_cases (Eop Ocast8unsigned (e2 ::: Enil))
+ | cast8unsigned_default:
+ forall (e1: expr),
+ cast8unsigned_cases e1.
+
+Definition cast8unsigned_match (e1: expr) :=
+ match e1 as z1 return cast8unsigned_cases z1 with
+ | Eop Ocast8unsigned (e2 ::: Enil) =>
+ cast8unsigned_case1 e2
+ | e1 =>
+ cast8unsigned_default e1
+ end.
+
+Definition cast8unsigned (e: expr) :=
+ match cast8unsigned_match e with
+ | cast8unsigned_case1 e1 => e
+ | cast8unsigned_default e1 => Eop Ocast8unsigned (e1 ::: Enil)
+ end.
+
Inductive cast16signed_cases: forall (e1: expr), Set :=
| cast16signed_case1:
forall (e2: expr),
@@ -844,6 +863,28 @@ Definition cast16signed (e: expr) :=
| cast16signed_default e1 => Eop Ocast16signed (e1 ::: Enil)
end.
+Inductive cast16unsigned_cases: forall (e1: expr), Set :=
+ | cast16unsigned_case1:
+ forall (e2: expr),
+ cast16unsigned_cases (Eop Ocast16unsigned (e2 ::: Enil))
+ | cast16unsigned_default:
+ forall (e1: expr),
+ cast16unsigned_cases e1.
+
+Definition cast16unsigned_match (e1: expr) :=
+ match e1 as z1 return cast16unsigned_cases z1 with
+ | Eop Ocast16unsigned (e2 ::: Enil) =>
+ cast16unsigned_case1 e2
+ | e1 =>
+ cast16unsigned_default e1
+ end.
+
+Definition cast16unsigned (e: expr) :=
+ match cast16unsigned_match e with
+ | cast16unsigned_case1 e1 => e
+ | cast16unsigned_default e1 => Eop Ocast16unsigned (e1 ::: Enil)
+ end.
+
Inductive singleoffloat_cases: forall (e1: expr), Set :=
| singleoffloat_case1:
forall (e2: expr),
diff --git a/backend/Cmconstrproof.v b/backend/Cmconstrproof.v
index f27fa73c..b9976eec 100644
--- a/backend/Cmconstrproof.v
+++ b/backend/Cmconstrproof.v
@@ -16,6 +16,7 @@ Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Op.
Require Import Globalenvs.
Require Import Cminor.
@@ -67,34 +68,67 @@ Scheme eval_expr_ind_3 := Minimality for eval_expr Sort Prop
with eval_exprlist_ind_3 := Minimality for eval_exprlist Sort Prop.
Hint Resolve eval_Evar eval_Eassign eval_Eop eval_Eload eval_Estore
- eval_Ecall eval_Econdition
+ eval_Ecall eval_Econdition eval_Ealloc
eval_Elet eval_Eletvar
eval_CEtrue eval_CEfalse eval_CEcond
eval_CEcondition eval_Enil eval_Econs: evalexpr.
+Lemma eval_list_one:
+ forall sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_exprlist ge sp le e1 m1 (a ::: Enil) t e2 m2 (v :: nil).
+Proof.
+ intros. econstructor. eauto. constructor. traceEq.
+Qed.
+
+Lemma eval_list_two:
+ forall sp le e1 m1 a1 t1 e2 m2 v1 a2 t2 e3 m3 v2 t,
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 ->
+ eval_expr ge sp le e2 m2 a2 t2 e3 m3 v2 ->
+ t = t1 ** t2 ->
+ eval_exprlist ge sp le e1 m1 (a1 ::: a2 ::: Enil) t e3 m3 (v1 :: v2 :: nil).
+Proof.
+ intros. econstructor. eauto. econstructor. eauto. constructor.
+ reflexivity. traceEq.
+Qed.
+
+Lemma eval_list_three:
+ forall sp le e1 m1 a1 t1 e2 m2 v1 a2 t2 e3 m3 v2 a3 t3 e4 m4 v3 t,
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 ->
+ eval_expr ge sp le e2 m2 a2 t2 e3 m3 v2 ->
+ eval_expr ge sp le e3 m3 a3 t3 e4 m4 v3 ->
+ t = t1 ** t2 ** t3 ->
+ eval_exprlist ge sp le e1 m1 (a1 ::: a2 ::: a3 ::: Enil) t e4 m4 (v1 :: v2 :: v3 :: nil).
+Proof.
+ intros. econstructor. eauto. econstructor. eauto. econstructor. eauto. constructor.
+ reflexivity. reflexivity. traceEq.
+Qed.
+
+Hint Resolve eval_list_one eval_list_two eval_list_three: evalexpr.
+
Lemma eval_lift_expr:
- forall w sp le e1 m1 a e2 m2 v,
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall w sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
forall p le', insert_lenv le p w le' ->
- eval_expr ge sp le' e1 m1 (lift_expr p a) e2 m2 v.
+ eval_expr ge sp le' e1 m1 (lift_expr p a) t e2 m2 v.
Proof.
intros w.
apply (eval_expr_ind_3 ge
- (fun sp le e1 m1 a e2 m2 v =>
+ (fun sp le e1 m1 a t e2 m2 v =>
forall p le', insert_lenv le p w le' ->
- eval_expr ge sp le' e1 m1 (lift_expr p a) e2 m2 v)
- (fun sp le e1 m1 a e2 m2 vb =>
+ eval_expr ge sp le' e1 m1 (lift_expr p a) t e2 m2 v)
+ (fun sp le e1 m1 a t e2 m2 vb =>
forall p le', insert_lenv le p w le' ->
- eval_condexpr ge sp le' e1 m1 (lift_condexpr p a) e2 m2 vb)
- (fun sp le e1 m1 al e2 m2 vl =>
+ eval_condexpr ge sp le' e1 m1 (lift_condexpr p a) t e2 m2 vb)
+ (fun sp le e1 m1 al t e2 m2 vl =>
forall p le', insert_lenv le p w le' ->
- eval_exprlist ge sp le' e1 m1 (lift_exprlist p al) e2 m2 vl));
+ eval_exprlist ge sp le' e1 m1 (lift_exprlist p al) t e2 m2 vl));
simpl; intros; eauto with evalexpr.
destruct v1; eapply eval_Econdition;
eauto with evalexpr; simpl; eauto with evalexpr.
- eapply eval_Elet. eauto. apply H2. apply insert_lenv_S; auto.
+ eapply eval_Elet. eauto. apply H2. apply insert_lenv_S; auto. auto.
case (le_gt_dec p n); intro.
apply eval_Eletvar. eapply insert_lenv_lookup2; eauto.
@@ -105,9 +139,9 @@ Proof.
Qed.
Lemma eval_lift:
- forall sp le e1 m1 a e2 m2 v w,
- eval_expr ge sp le e1 m1 a e2 m2 v ->
- eval_expr ge sp (w::le) e1 m1 (lift a) e2 m2 v.
+ forall sp le e1 m1 a t e2 m2 v w,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_expr ge sp (w::le) e1 m1 (lift a) t e2 m2 v.
Proof.
intros. unfold lift. eapply eval_lift_expr.
eexact H. apply insert_lenv_0.
@@ -126,67 +160,69 @@ Ltac TrivialOp cstr :=
of operator applications. *)
Lemma inv_eval_Eop_0:
- forall sp le e1 m1 op e2 m2 v,
- eval_expr ge sp le e1 m1 (Eop op Enil) e2 m2 v ->
- e2 = e1 /\ m2 = m1 /\ eval_operation ge sp op nil = Some v.
+ forall sp le e1 m1 op t e2 m2 v,
+ eval_expr ge sp le e1 m1 (Eop op Enil) t e2 m2 v ->
+ t = E0 /\ e2 = e1 /\ m2 = m1 /\ eval_operation ge sp op nil = Some v.
Proof.
intros. inversion H. inversion H6.
intuition. congruence.
Qed.
Lemma inv_eval_Eop_1:
- forall sp le e1 m1 op a1 e2 m2 v,
- eval_expr ge sp le e1 m1 (Eop op (a1 ::: Enil)) e2 m2 v ->
+ forall sp le e1 m1 op t a1 e2 m2 v,
+ eval_expr ge sp le e1 m1 (Eop op (a1 ::: Enil)) t e2 m2 v ->
exists v1,
- eval_expr ge sp le e1 m1 a1 e2 m2 v1 /\
+ eval_expr ge sp le e1 m1 a1 t e2 m2 v1 /\
eval_operation ge sp op (v1 :: nil) = Some v.
Proof.
intros.
- inversion H. inversion H6. inversion H21.
- subst e4; subst m4. subst e6; subst m6.
- exists v1; intuition. congruence.
+ inversion H. inversion H6. inversion H19.
+ subst. exists v1; intuition. rewrite E0_right. auto.
Qed.
Lemma inv_eval_Eop_2:
- forall sp le e1 m1 op a1 a2 e3 m3 v,
- eval_expr ge sp le e1 m1 (Eop op (a1 ::: a2 ::: Enil)) e3 m3 v ->
- exists e2, exists m2, exists v1, exists v2,
- eval_expr ge sp le e1 m1 a1 e2 m2 v1 /\
- eval_expr ge sp le e2 m2 a2 e3 m3 v2 /\
+ forall sp le e1 m1 op a1 a2 t3 e3 m3 v,
+ eval_expr ge sp le e1 m1 (Eop op (a1 ::: a2 ::: Enil)) t3 e3 m3 v ->
+ exists t1, exists t2, exists e2, exists m2, exists v1, exists v2,
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 /\
+ eval_expr ge sp le e2 m2 a2 t2 e3 m3 v2 /\
+ t3 = t1 ** t2 /\
eval_operation ge sp op (v1 :: v2 :: nil) = Some v.
Proof.
intros.
- inversion H. inversion H6. inversion H21. inversion H32.
- subst e7; subst m7. subst e9; subst m9.
- exists e4; exists m4; exists v1; exists v2. intuition. congruence.
+ inversion H. subst. inversion H6. subst. inversion H8. subst.
+ inversion H10. subst.
+ exists t1; exists t0; exists e0; exists m0; exists v0; exists v1.
+ intuition. traceEq.
Qed.
Ltac SimplEval :=
match goal with
- | [ |- (eval_expr _ ?sp ?le ?e1 ?m1 (Eop ?op Enil) ?e2 ?m2 ?v) -> _] =>
+ | [ |- (eval_expr _ ?sp ?le ?e1 ?m1 (Eop ?op Enil) ?t ?e2 ?m2 ?v) -> _] =>
intro XX1;
- generalize (inv_eval_Eop_0 sp le e1 m1 op e2 m2 v XX1);
+ generalize (inv_eval_Eop_0 sp le e1 m1 op t e2 m2 v XX1);
clear XX1;
- intros [XX1 [XX2 XX3]];
- subst e2; subst m2; simpl in XX3;
- try (simplify_eq XX3; clear XX3;
+ intros [XX1 [XX2 [XX3 XX4]]];
+ subst t e2 m2; simpl in XX4;
+ try (simplify_eq XX4; clear XX4;
let EQ := fresh "EQ" in (intro EQ; rewrite EQ))
- | [ |- (eval_expr _ ?sp ?le ?e1 ?m1 (Eop ?op (?a1 ::: Enil)) ?e2 ?m2 ?v) -> _] =>
+ | [ |- (eval_expr _ ?sp ?le ?e1 ?m1 (Eop ?op (?a1 ::: Enil)) ?t ?e2 ?m2 ?v) -> _] =>
intro XX1;
- generalize (inv_eval_Eop_1 sp le e1 m1 op a1 e2 m2 v XX1);
+ generalize (inv_eval_Eop_1 sp le e1 m1 op t a1 e2 m2 v XX1);
clear XX1;
let v1 := fresh "v" in let EV := fresh "EV" in
let EQ := fresh "EQ" in
(intros [v1 [EV EQ]]; simpl in EQ)
- | [ |- (eval_expr _ ?sp ?le ?e1 ?m1 (Eop ?op (?a1 ::: ?a2 ::: Enil)) ?e2 ?m2 ?v) -> _] =>
+ | [ |- (eval_expr _ ?sp ?le ?e1 ?m1 (Eop ?op (?a1 ::: ?a2 ::: Enil)) ?t ?e2 ?m2 ?v) -> _] =>
intro XX1;
- generalize (inv_eval_Eop_2 sp le e1 m1 op a1 a2 e2 m2 v XX1);
+ generalize (inv_eval_Eop_2 sp le e1 m1 op a1 a2 t e2 m2 v XX1);
clear XX1;
+ let t1 := fresh "t" in let t2 := fresh "t" in
let e := fresh "e" in let m := fresh "m" in
let v1 := fresh "v" in let v2 := fresh "v" in
let EV1 := fresh "EV" in let EV2 := fresh "EV" in
- let EQ := fresh "EQ" in
- (intros [e [m [v1 [v2 [EV1 [EV2 EQ]]]]]]; simpl in EQ)
+ let EQ := fresh "EQ" in let TR := fresh "TR" in
+ (intros [t1 [t2 [e [m [v1 [v2 [EV1 [EV2 [TR EQ]]]]]]]]]; simpl in EQ)
| _ => idtac
end.
@@ -209,57 +245,57 @@ Ltac InvEval H :=
*)
Theorem eval_negint:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (negint a) e2 m2 (Vint (Int.neg x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (negint a) t e2 m2 (Vint (Int.neg x)).
Proof.
TrivialOp negint.
Qed.
Theorem eval_negfloat:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e1 m1 (negfloat a) e2 m2 (Vfloat (Float.neg x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e1 m1 (negfloat a) t e2 m2 (Vfloat (Float.neg x)).
Proof.
TrivialOp negfloat.
Qed.
Theorem eval_absfloat:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e1 m1 (absfloat a) e2 m2 (Vfloat (Float.abs x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e1 m1 (absfloat a) t e2 m2 (Vfloat (Float.abs x)).
Proof.
TrivialOp absfloat.
Qed.
Theorem eval_intoffloat:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e1 m1 (intoffloat a) e2 m2 (Vint (Float.intoffloat x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e1 m1 (intoffloat a) t e2 m2 (Vint (Float.intoffloat x)).
Proof.
TrivialOp intoffloat.
Qed.
Theorem eval_floatofint:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (floatofint a) e2 m2 (Vfloat (Float.floatofint x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (floatofint a) t e2 m2 (Vfloat (Float.floatofint x)).
Proof.
TrivialOp floatofint.
Qed.
Theorem eval_floatofintu:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (floatofintu a) e2 m2 (Vfloat (Float.floatofintu x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (floatofintu a) t e2 m2 (Vfloat (Float.floatofintu x)).
Proof.
TrivialOp floatofintu.
Qed.
Theorem eval_notint:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (notint a) e2 m2 (Vint (Int.not x)).
+ forall sp le e1 m1 a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (notint a) t e2 m2 (Vint (Int.not x)).
Proof.
unfold notint; intros until x; case (notint_match a); intros.
InvEval H. FuncInv. EvalOp. simpl. congruence.
@@ -269,15 +305,15 @@ Proof.
eapply eval_Eop.
eapply eval_Econs. apply eval_Eletvar. simpl. reflexivity.
eapply eval_Econs. apply eval_Eletvar. simpl. reflexivity.
- apply eval_Enil.
- simpl. rewrite Int.or_idem. auto.
+ apply eval_Enil. reflexivity. reflexivity.
+ simpl. rewrite Int.or_idem. auto. traceEq.
Qed.
Lemma eval_notbool_base:
- forall sp le e1 m1 a e2 m2 v b,
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall sp le e1 m1 a t e2 m2 v b,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
Val.bool_of_val v b ->
- eval_expr ge sp le e1 m1 (notbool_base a) e2 m2 (Val.of_bool (negb b)).
+ eval_expr ge sp le e1 m1 (notbool_base a) t e2 m2 (Val.of_bool (negb b)).
Proof.
TrivialOp notbool_base. simpl.
inversion H0.
@@ -290,10 +326,10 @@ Hint Resolve Val.bool_of_true_val Val.bool_of_false_val
Val.bool_of_true_val_inv Val.bool_of_false_val_inv: valboolof.
Theorem eval_notbool:
- forall a sp le e1 m1 e2 m2 v b,
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall a sp le e1 m1 t e2 m2 v b,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
Val.bool_of_val v b ->
- eval_expr ge sp le e1 m1 (notbool a) e2 m2 (Val.of_bool (negb b)).
+ eval_expr ge sp le e1 m1 (notbool a) t e2 m2 (Val.of_bool (negb b)).
Proof.
assert (N1: forall v b, Val.is_false v -> Val.bool_of_val v b -> Val.is_true (Val.of_bool (negb b))).
intros. inversion H0; simpl; auto; subst v; simpl in H.
@@ -305,34 +341,33 @@ Proof.
induction a; simpl; intros; try (eapply eval_notbool_base; eauto).
destruct o; try (eapply eval_notbool_base; eauto).
- destruct e. inversion H. inversion H7. subst vl.
- simpl in H11. inversion H11. subst v0; subst v.
+ destruct e. InvEval H. injection XX4; clear XX4; intro; subst v.
inversion H0. rewrite Int.eq_false; auto.
simpl; eauto with evalexpr.
rewrite Int.eq_true; simpl; eauto with evalexpr.
eapply eval_notbool_base; eauto.
- inversion H. subst sp0 le0 e0 m op al e3 m0 v0.
- simpl in H11. eapply eval_Eop; eauto.
+ inversion H. subst.
+ simpl in H12. eapply eval_Eop; eauto.
simpl. caseEq (eval_condition c vl); intros.
- rewrite H1 in H11.
+ rewrite H1 in H12.
assert (b0 = b).
- destruct b0; inversion H11; subst v; inversion H0; auto.
+ destruct b0; inversion H12; subst v; inversion H0; auto.
subst b0. rewrite (Op.eval_negate_condition _ _ H1).
destruct b; reflexivity.
- rewrite H1 in H11; discriminate.
+ rewrite H1 in H12; discriminate.
inversion H; eauto 10 with evalexpr valboolof.
inversion H; eauto 10 with evalexpr valboolof.
- inversion H. eapply eval_Econdition. eexact H11.
- destruct v1; eauto.
+ inversion H. subst. eapply eval_Econdition with (t2 := t8). eexact H36.
+ destruct v4; eauto. auto.
Qed.
Theorem eval_addimm:
- forall sp le e1 m1 n a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (addimm n a) e2 m2 (Vint (Int.add x n)).
+ forall sp le e1 m1 n a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (addimm n a) t e2 m2 (Vint (Int.add x n)).
Proof.
unfold addimm; intros until x.
generalize (Int.eq_spec n Int.zero). case (Int.eq n Int.zero); intro.
@@ -341,16 +376,16 @@ Proof.
InvEval H0. EvalOp. simpl. rewrite Int.add_commut. auto.
InvEval H0. destruct (Genv.find_symbol ge s); discriminate.
InvEval H0.
- destruct sp; simpl in XX3; discriminate.
+ destruct sp; simpl in XX4; discriminate.
InvEval H0. FuncInv. EvalOp. simpl. subst x.
rewrite Int.add_assoc. decEq; decEq; decEq. apply Int.add_commut.
EvalOp.
Qed.
Theorem eval_addimm_ptr:
- forall sp le e1 m1 n a e2 m2 b ofs,
- eval_expr ge sp le e1 m1 a e2 m2 (Vptr b ofs) ->
- eval_expr ge sp le e1 m1 (addimm n a) e2 m2 (Vptr b (Int.add ofs n)).
+ forall sp le e1 m1 n t a e2 m2 b ofs,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vptr b ofs) ->
+ eval_expr ge sp le e1 m1 (addimm n a) t e2 m2 (Vptr b (Int.add ofs n)).
Proof.
unfold addimm; intros until ofs.
generalize (Int.eq_spec n Int.zero). case (Int.eq n Int.zero); intro.
@@ -361,8 +396,8 @@ Proof.
destruct (Genv.find_symbol ge s).
rewrite Int.add_commut. congruence.
discriminate.
- InvEval H0. destruct sp; simpl in XX3; try discriminate.
- inversion XX3. EvalOp. simpl. decEq. decEq.
+ InvEval H0. destruct sp; simpl in XX4; try discriminate.
+ inversion XX4. EvalOp. simpl. decEq. decEq.
rewrite Int.add_assoc. decEq. apply Int.add_commut.
InvEval H0. FuncInv. subst b0; subst ofs. EvalOp. simpl.
rewrite (Int.add_commut n m). rewrite Int.add_assoc. auto.
@@ -370,13 +405,14 @@ Proof.
Qed.
Theorem eval_add:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (add a b) e3 m3 (Vint (Int.add x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (add a b) (t1**t2) e3 m3 (Vint (Int.add x y)).
Proof.
intros until y. unfold add; case (add_match a b); intros.
- InvEval H. rewrite Int.add_commut. apply eval_addimm. assumption.
+ InvEval H. rewrite Int.add_commut. apply eval_addimm.
+ rewrite E0_left; assumption.
InvEval H. FuncInv. InvEval H0. FuncInv.
replace (Int.add x y) with (Int.add (Int.add i i0) (Int.add n1 n2)).
apply eval_addimm. EvalOp.
@@ -387,7 +423,7 @@ Proof.
apply eval_addimm. EvalOp.
subst x. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
InvEval H0. FuncInv.
- apply eval_addimm. auto.
+ apply eval_addimm. rewrite E0_right. auto.
InvEval H0. FuncInv.
replace (Int.add x y) with (Int.add (Int.add x i) n2).
apply eval_addimm. EvalOp.
@@ -396,10 +432,10 @@ Proof.
Qed.
Theorem eval_add_ptr:
- forall sp le e1 m1 a e2 m2 p x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vptr p x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (add a b) e3 m3 (Vptr p (Int.add x y)).
+ forall sp le e1 m1 a t1 e2 m2 p x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vptr p x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (add a b) (t1**t2) e3 m3 (Vptr p (Int.add x y)).
Proof.
intros until y. unfold add; case (add_match a b); intros.
InvEval H.
@@ -412,7 +448,7 @@ Proof.
replace (Int.add x y) with (Int.add (Int.add i y) n1).
apply eval_addimm_ptr. subst b0. EvalOp.
subst x. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
- InvEval H0. apply eval_addimm_ptr. auto.
+ InvEval H0. apply eval_addimm_ptr. rewrite E0_right. auto.
InvEval H0. FuncInv.
replace (Int.add x y) with (Int.add (Int.add x i) n2).
apply eval_addimm_ptr. EvalOp.
@@ -421,14 +457,14 @@ Proof.
Qed.
Theorem eval_add_ptr_2:
- forall sp le e1 m1 a e2 m2 p x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vptr p y) ->
- eval_expr ge sp le e1 m1 (add a b) e3 m3 (Vptr p (Int.add y x)).
+ forall sp le e1 m1 a t1 e2 m2 p x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vptr p y) ->
+ eval_expr ge sp le e1 m1 (add a b) (t1**t2) e3 m3 (Vptr p (Int.add y x)).
Proof.
intros until y. unfold add; case (add_match a b); intros.
InvEval H.
- apply eval_addimm_ptr. auto.
+ apply eval_addimm_ptr. rewrite E0_left. auto.
InvEval H. FuncInv. InvEval H0. FuncInv.
replace (Int.add y x) with (Int.add (Int.add i0 i) (Int.add n1 n2)).
apply eval_addimm_ptr. subst b0. EvalOp.
@@ -448,15 +484,15 @@ Proof.
Qed.
Theorem eval_sub:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (sub a b) e3 m3 (Vint (Int.sub x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (sub a b) (t1**t2) e3 m3 (Vint (Int.sub x y)).
Proof.
intros until y.
unfold sub; case (sub_match a b); intros.
InvEval H0. rewrite Int.sub_add_opp.
- apply eval_addimm. assumption.
+ apply eval_addimm. rewrite E0_right. assumption.
InvEval H. FuncInv. InvEval H0. FuncInv.
replace (Int.sub x y) with (Int.add (Int.sub i i0) (Int.sub n1 n2)).
apply eval_addimm. EvalOp.
@@ -476,15 +512,15 @@ Proof.
Qed.
Theorem eval_sub_ptr_int:
- forall sp le e1 m1 a e2 m2 p x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vptr p x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (sub a b) e3 m3 (Vptr p (Int.sub x y)).
+ forall sp le e1 m1 a t1 e2 m2 p x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vptr p x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (sub a b) (t1**t2) e3 m3 (Vptr p (Int.sub x y)).
Proof.
intros until y.
unfold sub; case (sub_match a b); intros.
InvEval H0. rewrite Int.sub_add_opp.
- apply eval_addimm_ptr. assumption.
+ apply eval_addimm_ptr. rewrite E0_right. assumption.
InvEval H. FuncInv. InvEval H0. FuncInv.
subst b0.
replace (Int.sub x y) with (Int.add (Int.sub i i0) (Int.sub n1 n2)).
@@ -505,10 +541,10 @@ Proof.
Qed.
Theorem eval_sub_ptr_ptr:
- forall sp le e1 m1 a e2 m2 p x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vptr p x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vptr p y) ->
- eval_expr ge sp le e1 m1 (sub a b) e3 m3 (Vint (Int.sub x y)).
+ forall sp le e1 m1 a t1 e2 m2 p x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vptr p x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vptr p y) ->
+ eval_expr ge sp le e1 m1 (sub a b) (t1**t2) e3 m3 (Vint (Int.sub x y)).
Proof.
intros until y.
unfold sub; case (sub_match a b); intros.
@@ -535,9 +571,9 @@ Proof.
Qed.
Lemma eval_rolm:
- forall sp le e1 m1 a amount mask e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (rolm a amount mask) e2 m2 (Vint (Int.rolm x amount mask)).
+ forall sp le e1 m1 a amount mask t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (rolm a amount mask) t e2 m2 (Vint (Int.rolm x amount mask)).
Proof.
intros until x. unfold rolm; case (rolm_match a); intros.
InvEval H. eauto with evalexpr.
@@ -554,10 +590,10 @@ Proof.
Qed.
Theorem eval_shlimm:
- forall sp le e1 m1 a n e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
+ forall sp le e1 m1 a n t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
Int.ltu n (Int.repr 32) = true ->
- eval_expr ge sp le e1 m1 (shlimm a n) e2 m2 (Vint (Int.shl x n)).
+ eval_expr ge sp le e1 m1 (shlimm a n) t e2 m2 (Vint (Int.shl x n)).
Proof.
intros. unfold shlimm.
generalize (Int.eq_spec n Int.zero); case (Int.eq n Int.zero); intro.
@@ -568,10 +604,10 @@ Proof.
Qed.
Theorem eval_shruimm:
- forall sp le e1 m1 a n e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
+ forall sp le e1 m1 a n t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
Int.ltu n (Int.repr 32) = true ->
- eval_expr ge sp le e1 m1 (shruimm a n) e2 m2 (Vint (Int.shru x n)).
+ eval_expr ge sp le e1 m1 (shruimm a n) t e2 m2 (Vint (Int.shru x n)).
Proof.
intros. unfold shruimm.
generalize (Int.eq_spec n Int.zero); case (Int.eq n Int.zero); intro.
@@ -582,9 +618,9 @@ Proof.
Qed.
Lemma eval_mulimm_base:
- forall sp le e1 m1 a n e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (mulimm_base n a) e2 m2 (Vint (Int.mul x n)).
+ forall sp le e1 m1 a t n e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (mulimm_base n a) t e2 m2 (Vint (Int.mul x n)).
Proof.
intros; unfold mulimm_base.
generalize (Int.one_bits_decomp n).
@@ -597,7 +633,7 @@ Proof.
rewrite Int.add_zero. rewrite <- Int.shl_mul.
apply eval_shlimm. auto. auto with coqlib.
destruct l.
- intros. apply eval_Elet with e2 m2 (Vint x). auto.
+ intros. apply eval_Elet with t e2 m2 (Vint x) E0. auto.
rewrite H1. simpl. rewrite Int.add_zero.
rewrite Int.mul_add_distr_r.
rewrite <- Int.shl_mul.
@@ -609,18 +645,19 @@ Proof.
apply eval_shlimm. apply eval_Eletvar. simpl. reflexivity.
auto with coqlib.
auto with evalexpr.
- reflexivity.
+ reflexivity. traceEq. reflexivity. traceEq.
intros. EvalOp.
Qed.
Theorem eval_mulimm:
- forall sp le e1 m1 a n e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (mulimm n a) e2 m2 (Vint (Int.mul x n)).
+ forall sp le e1 m1 a n t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (mulimm n a) t e2 m2 (Vint (Int.mul x n)).
Proof.
intros until x; unfold mulimm.
generalize (Int.eq_spec n Int.zero); case (Int.eq n Int.zero); intro.
- subst n. rewrite Int.mul_zero. eauto with evalexpr.
+ subst n. rewrite Int.mul_zero.
+ intro. eapply eval_Elet; eauto with evalexpr. traceEq.
generalize (Int.eq_spec n Int.one); case (Int.eq n Int.one); intro.
subst n. rewrite Int.mul_one. auto.
case (mulimm_match a); intros.
@@ -633,24 +670,25 @@ Proof.
Qed.
Theorem eval_mul:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (mul a b) e3 m3 (Vint (Int.mul x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (mul a b) (t1**t2) e3 m3 (Vint (Int.mul x y)).
Proof.
intros until y.
unfold mul; case (mul_match a b); intros.
- InvEval H. rewrite Int.mul_commut. apply eval_mulimm. auto.
- InvEval H0. apply eval_mulimm. auto.
+ InvEval H. rewrite Int.mul_commut. apply eval_mulimm.
+ rewrite E0_left; auto.
+ InvEval H0. rewrite E0_right. apply eval_mulimm. auto.
EvalOp.
Qed.
Theorem eval_divs:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
y <> Int.zero ->
- eval_expr ge sp le e1 m1 (divs a b) e3 m3 (Vint (Int.divs x y)).
+ eval_expr ge sp le e1 m1 (divs a b) (t1**t2) e3 m3 (Vint (Int.divs x y)).
Proof.
TrivialOp divs. simpl.
predSpec Int.eq Int.eq_spec y Int.zero. contradiction. auto.
@@ -662,11 +700,11 @@ Lemma eval_mod_aux:
y <> Int.zero ->
eval_operation ge sp divop (Vint x :: Vint y :: nil) =
Some (Vint (semdivop x y))) ->
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
y <> Int.zero ->
- eval_expr ge sp le e1 m1 (mod_aux divop a b) e3 m3
+ eval_expr ge sp le e1 m1 (mod_aux divop a b) (t1**t2) e3 m3
(Vint (Int.sub x (Int.mul (semdivop x y) y))).
Proof.
intros; unfold mod_aux.
@@ -678,18 +716,21 @@ Proof.
eapply eval_Econs. eapply eval_Eop.
eapply eval_Econs. apply eval_Eletvar. simpl; reflexivity.
eapply eval_Econs. apply eval_Eletvar. simpl; reflexivity.
- apply eval_Enil. apply H. assumption.
+ apply eval_Enil. reflexivity. reflexivity.
+ apply H. assumption.
eapply eval_Econs. apply eval_Eletvar. simpl; reflexivity.
- apply eval_Enil. simpl; reflexivity. apply eval_Enil.
- reflexivity.
+ apply eval_Enil. reflexivity. reflexivity.
+ simpl; reflexivity. apply eval_Enil.
+ reflexivity. reflexivity. reflexivity.
+ reflexivity. traceEq.
Qed.
Theorem eval_mods:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
y <> Int.zero ->
- eval_expr ge sp le e1 m1 (mods a b) e3 m3 (Vint (Int.mods x y)).
+ eval_expr ge sp le e1 m1 (mods a b) (t1**t2) e3 m3 (Vint (Int.mods x y)).
Proof.
intros; unfold mods.
rewrite Int.mods_divs.
@@ -699,44 +740,44 @@ Proof.
Qed.
Lemma eval_divu_base:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
y <> Int.zero ->
- eval_expr ge sp le e1 m1 (Eop Odivu (a ::: b ::: Enil)) e3 m3 (Vint (Int.divu x y)).
+ eval_expr ge sp le e1 m1 (Eop Odivu (a ::: b ::: Enil)) (t1**t2) e3 m3 (Vint (Int.divu x y)).
Proof.
intros. EvalOp. simpl.
predSpec Int.eq Int.eq_spec y Int.zero. contradiction. auto.
Qed.
Theorem eval_divu:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
y <> Int.zero ->
- eval_expr ge sp le e1 m1 (divu a b) e3 m3 (Vint (Int.divu x y)).
+ eval_expr ge sp le e1 m1 (divu a b) (t1**t2) e3 m3 (Vint (Int.divu x y)).
Proof.
intros until y.
unfold divu; case (divu_match b); intros.
InvEval H0. caseEq (Int.is_power2 y).
intros. rewrite (Int.divu_pow2 x y i H0).
- apply eval_shruimm. auto.
+ apply eval_shruimm. rewrite E0_right. auto.
apply Int.is_power2_range with y. auto.
intros. subst n2. eapply eval_divu_base. eexact H. EvalOp. auto.
eapply eval_divu_base; eauto.
Qed.
Theorem eval_modu:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
y <> Int.zero ->
- eval_expr ge sp le e1 m1 (modu a b) e3 m3 (Vint (Int.modu x y)).
+ eval_expr ge sp le e1 m1 (modu a b) (t1**t2) e3 m3 (Vint (Int.modu x y)).
Proof.
intros until y; unfold modu; case (divu_match b); intros.
InvEval H0. caseEq (Int.is_power2 y).
intros. rewrite (Int.modu_and x y i H0).
- rewrite <- Int.rolm_zero. apply eval_rolm. auto.
+ rewrite <- Int.rolm_zero. apply eval_rolm. rewrite E0_right; auto.
intro. rewrite Int.modu_divu. eapply eval_mod_aux.
intros. simpl. predSpec Int.eq Int.eq_spec y0 Int.zero.
contradiction. auto.
@@ -748,9 +789,9 @@ Proof.
Qed.
Theorem eval_andimm:
- forall sp le e1 m1 n a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (andimm n a) e2 m2 (Vint (Int.and x n)).
+ forall sp le e1 m1 n a t e2 m2 x,
+ eval_expr ge sp le e1 m1 a t e2 m2 (Vint x) ->
+ eval_expr ge sp le e1 m1 (andimm n a) t e2 m2 (Vint (Int.and x n)).
Proof.
intros. unfold andimm. case (Int.is_rlw_mask n).
rewrite <- Int.rolm_zero. apply eval_rolm; auto.
@@ -758,25 +799,26 @@ Proof.
Qed.
Theorem eval_and:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (and a b) e3 m3 (Vint (Int.and x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (and a b) (t1**t2) e3 m3 (Vint (Int.and x y)).
Proof.
intros until y; unfold and; case (mul_match a b); intros.
- InvEval H. rewrite Int.and_commut. apply eval_andimm; auto.
- InvEval H0. apply eval_andimm; auto.
+ InvEval H. rewrite Int.and_commut.
+ rewrite E0_left; apply eval_andimm; auto.
+ InvEval H0. rewrite E0_right; apply eval_andimm; auto.
EvalOp.
Qed.
Remark eval_same_expr_pure:
- forall a1 a2 sp le e1 m1 e2 m2 v1 e3 m3 v2,
+ forall a1 a2 sp le e1 m1 t1 e2 m2 v1 t2 e3 m3 v2,
same_expr_pure a1 a2 = true ->
- eval_expr ge sp le e1 m1 a1 e2 m2 v1 ->
- eval_expr ge sp le e2 m2 a2 e3 m3 v2 ->
- a2 = a1 /\ v2 = v1 /\ e2 = e1 /\ m2 = m1.
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 ->
+ eval_expr ge sp le e2 m2 a2 t2 e3 m3 v2 ->
+ t1 = E0 /\ t2 = E0 /\ a2 = a1 /\ v2 = v1 /\ e2 = e1 /\ m2 = m1.
Proof.
- intros until v1.
+ intros until v2.
destruct a1; simpl; try (intros; discriminate).
destruct a2; simpl; try (intros; discriminate).
case (ident_eq i i0); intros.
@@ -786,22 +828,20 @@ Proof.
Qed.
Lemma eval_or:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (or a b) e3 m3 (Vint (Int.or x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (or a b) (t1**t2) e3 m3 (Vint (Int.or x y)).
Proof.
intros until y; unfold or; case (or_match a b); intros.
generalize (Int.eq_spec amount1 amount2); case (Int.eq amount1 amount2); intro.
case (Int.is_rlw_mask (Int.or mask1 mask2)).
- caseEq (same_expr_pure t1 t2); intro.
+ caseEq (same_expr_pure t0 t3); intro.
simpl. InvEval H. FuncInv. InvEval H0. FuncInv.
- generalize (eval_same_expr_pure _ _ _ _ _ _ _ _ _ _ _ _ H2 EV EV0).
- intros [EQ1 [EQ2 [EQ3 EQ4]]].
- injection EQ2; intro EQ5.
- subst t2; subst e2; subst m2; subst i0.
- EvalOp. simpl. subst x; subst y; subst amount2.
- rewrite Int.or_rolm. auto.
+ generalize (eval_same_expr_pure _ _ _ _ _ _ _ _ _ _ _ _ _ _ H2 EV EV0).
+ intros [EQ1 [EQ2 [EQ3 [EQ4 [EQ5 EQ6]]]]].
+ injection EQ4; intro EQ7. subst.
+ EvalOp. simpl. rewrite Int.or_rolm. auto.
simpl. EvalOp.
simpl. EvalOp.
simpl. EvalOp.
@@ -809,173 +849,184 @@ Proof.
Qed.
Theorem eval_xor:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (xor a b) e3 m3 (Vint (Int.xor x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (xor a b) (t1**t2) e3 m3 (Vint (Int.xor x y)).
Proof. TrivialOp xor. Qed.
Theorem eval_shl:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
Int.ltu y (Int.repr 32) = true ->
- eval_expr ge sp le e1 m1 (shl a b) e3 m3 (Vint (Int.shl x y)).
+ eval_expr ge sp le e1 m1 (shl a b) (t1**t2) e3 m3 (Vint (Int.shl x y)).
Proof.
intros until y; unfold shl; case (shift_match b); intros.
- InvEval H0. apply eval_shlimm; auto.
+ InvEval H0. rewrite E0_right. apply eval_shlimm; auto.
EvalOp. simpl. rewrite H1. auto.
Qed.
Theorem eval_shr:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
Int.ltu y (Int.repr 32) = true ->
- eval_expr ge sp le e1 m1 (shr a b) e3 m3 (Vint (Int.shr x y)).
+ eval_expr ge sp le e1 m1 (shr a b) (t1**t2) e3 m3 (Vint (Int.shr x y)).
Proof.
TrivialOp shr. simpl. rewrite H1. auto.
Qed.
Theorem eval_shru:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
Int.ltu y (Int.repr 32) = true ->
- eval_expr ge sp le e1 m1 (shru a b) e3 m3 (Vint (Int.shru x y)).
+ eval_expr ge sp le e1 m1 (shru a b) (t1**t2) e3 m3 (Vint (Int.shru x y)).
Proof.
intros until y; unfold shru; case (shift_match b); intros.
- InvEval H0. apply eval_shruimm; auto.
+ InvEval H0. rewrite E0_right; apply eval_shruimm; auto.
EvalOp. simpl. rewrite H1. auto.
Qed.
Theorem eval_addf:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vfloat y) ->
- eval_expr ge sp le e1 m1 (addf a b) e3 m3 (Vfloat (Float.add x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vfloat y) ->
+ eval_expr ge sp le e1 m1 (addf a b) (t1**t2) e3 m3 (Vfloat (Float.add x y)).
Proof.
intros until y; unfold addf; case (addf_match a b); intros.
- InvEval H. FuncInv. EvalOp. simpl. subst x. reflexivity.
+ InvEval H. FuncInv. EvalOp.
+ econstructor; eauto. econstructor; eauto. econstructor; eauto. constructor.
+ traceEq. simpl. subst x. reflexivity.
InvEval H0. FuncInv. eapply eval_Elet. eexact H. EvalOp.
- eapply eval_Econs. apply eval_lift. eexact EV.
- eapply eval_Econs. apply eval_lift. eexact EV0.
- eapply eval_Econs. apply eval_Eletvar. simpl; reflexivity.
- apply eval_Enil.
- subst y. rewrite Float.addf_commut. reflexivity.
+ econstructor; eauto with evalexpr.
+ econstructor; eauto with evalexpr.
+ econstructor. apply eval_Eletvar. simpl; reflexivity.
+ constructor. reflexivity. traceEq.
+ subst y. rewrite Float.addf_commut. reflexivity. auto.
EvalOp.
Qed.
Theorem eval_subf:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vfloat y) ->
- eval_expr ge sp le e1 m1 (subf a b) e3 m3 (Vfloat (Float.sub x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vfloat y) ->
+ eval_expr ge sp le e1 m1 (subf a b) (t1**t2) e3 m3 (Vfloat (Float.sub x y)).
Proof.
intros until y; unfold subf; case (subf_match a b); intros.
- InvEval H. FuncInv. EvalOp. subst x. reflexivity.
+ InvEval H. FuncInv. EvalOp.
+ econstructor; eauto. econstructor; eauto. econstructor; eauto. constructor.
+ traceEq. subst x. reflexivity.
EvalOp.
Qed.
Theorem eval_mulf:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vfloat y) ->
- eval_expr ge sp le e1 m1 (mulf a b) e3 m3 (Vfloat (Float.mul x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vfloat y) ->
+ eval_expr ge sp le e1 m1 (mulf a b) (t1**t2) e3 m3 (Vfloat (Float.mul x y)).
Proof. TrivialOp mulf. Qed.
Theorem eval_divf:
- forall sp le e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vfloat y) ->
- eval_expr ge sp le e1 m1 (divf a b) e3 m3 (Vfloat (Float.div x y)).
+ forall sp le e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vfloat y) ->
+ eval_expr ge sp le e1 m1 (divf a b) (t1**t2) e3 m3 (Vfloat (Float.div x y)).
Proof. TrivialOp divf. Qed.
Theorem eval_cast8signed:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (cast8signed a) e2 m2 (Vint (Int.cast8signed x)).
+ forall sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_expr ge sp le e1 m1 (cast8signed a) t e2 m2 (Val.cast8signed v).
Proof.
- intros until x; unfold cast8signed; case (cast8signed_match a); intros.
- InvEval H. FuncInv. EvalOp. subst x. rewrite Int.cast8_signed_idem. reflexivity.
+ intros until v; unfold cast8signed; case (cast8signed_match a); intros.
+ replace (Val.cast8signed v) with v. auto.
+ InvEval H. inversion EQ. destruct v0; simpl; auto. rewrite Int.cast8_signed_idem. reflexivity.
EvalOp.
Qed.
Theorem eval_cast8unsigned:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (cast8unsigned a) e2 m2 (Vint (Int.cast8unsigned x)).
+ forall sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_expr ge sp le e1 m1 (cast8unsigned a) t e2 m2 (Val.cast8unsigned v).
Proof.
- intros. unfold cast8unsigned. rewrite Int.cast8unsigned_and.
- rewrite <- Int.rolm_zero. eapply eval_rolm; eauto.
+ intros until v; unfold cast8unsigned; case (cast8unsigned_match a); intros.
+ replace (Val.cast8unsigned v) with v. auto.
+ InvEval H. inversion EQ. destruct v0; simpl; auto. rewrite Int.cast8_unsigned_idem. reflexivity.
+ EvalOp.
Qed.
-
+
Theorem eval_cast16signed:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (cast16signed a) e2 m2 (Vint (Int.cast16signed x)).
+ forall sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_expr ge sp le e1 m1 (cast16signed a) t e2 m2 (Val.cast16signed v).
Proof.
- intros until x; unfold cast16signed; case (cast16signed_match a); intros.
- InvEval H. FuncInv. EvalOp. subst x. rewrite Int.cast16_signed_idem. reflexivity.
+ intros until v; unfold cast16signed; case (cast16signed_match a); intros.
+ replace (Val.cast16signed v) with v. auto.
+ InvEval H. inversion EQ. destruct v0; simpl; auto. rewrite Int.cast16_signed_idem. reflexivity.
EvalOp.
Qed.
Theorem eval_cast16unsigned:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e1 m1 (cast16unsigned a) e2 m2 (Vint (Int.cast16unsigned x)).
+ forall sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_expr ge sp le e1 m1 (cast16unsigned a) t e2 m2 (Val.cast16unsigned v).
Proof.
- intros. unfold cast16unsigned. rewrite Int.cast16unsigned_and.
- rewrite <- Int.rolm_zero. eapply eval_rolm; eauto.
+ intros until v; unfold cast16unsigned; case (cast16unsigned_match a); intros.
+ replace (Val.cast16unsigned v) with v. auto.
+ InvEval H. inversion EQ. destruct v0; simpl; auto. rewrite Int.cast16_unsigned_idem. reflexivity.
+ EvalOp.
Qed.
Theorem eval_singleoffloat:
- forall sp le e1 m1 a e2 m2 x,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e1 m1 (singleoffloat a) e2 m2 (Vfloat (Float.singleoffloat x)).
+ forall sp le e1 m1 a t e2 m2 v,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
+ eval_expr ge sp le e1 m1 (singleoffloat a) t e2 m2 (Val.singleoffloat v).
Proof.
- intros until x; unfold singleoffloat; case (singleoffloat_match a); intros.
- InvEval H. FuncInv. EvalOp. subst x. rewrite Float.singleoffloat_idem. reflexivity.
+ intros until v; unfold singleoffloat; case (singleoffloat_match a); intros.
+ replace (Val.singleoffloat v) with v. auto.
+ InvEval H. inversion EQ. destruct v0; simpl; auto. rewrite Float.singleoffloat_idem. reflexivity.
EvalOp.
Qed.
Theorem eval_cmp:
- forall sp le c e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (cmp c a b) e3 m3 (Val.of_bool (Int.cmp c x y)).
+ forall sp le c e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (cmp c a b) (t1**t2) e3 m3 (Val.of_bool (Int.cmp c x y)).
Proof.
TrivialOp cmp.
simpl. case (Int.cmp c x y); auto.
Qed.
Theorem eval_cmp_null_r:
- forall sp le c e1 m1 a e2 m2 p x b e3 m3 v,
- eval_expr ge sp le e1 m1 a e2 m2 (Vptr p x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint Int.zero) ->
+ forall sp le c e1 m1 a t1 e2 m2 p x b t2 e3 m3 v,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vptr p x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint Int.zero) ->
(c = Ceq /\ v = Vfalse) \/ (c = Cne /\ v = Vtrue) ->
- eval_expr ge sp le e1 m1 (cmp c a b) e3 m3 v.
+ eval_expr ge sp le e1 m1 (cmp c a b) (t1**t2) e3 m3 v.
Proof.
TrivialOp cmp.
simpl. elim H1; intros [EQ1 EQ2]; subst c; subst v; reflexivity.
Qed.
Theorem eval_cmp_null_l:
- forall sp le c e1 m1 a e2 m2 p x b e3 m3 v,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint Int.zero) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vptr p x) ->
+ forall sp le c e1 m1 a t1 e2 m2 p x b t2 e3 m3 v,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint Int.zero) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vptr p x) ->
(c = Ceq /\ v = Vfalse) \/ (c = Cne /\ v = Vtrue) ->
- eval_expr ge sp le e1 m1 (cmp c a b) e3 m3 v.
+ eval_expr ge sp le e1 m1 (cmp c a b) (t1**t2) e3 m3 v.
Proof.
TrivialOp cmp.
simpl. elim H1; intros [EQ1 EQ2]; subst c; subst v; reflexivity.
Qed.
Theorem eval_cmp_ptr:
- forall sp le c e1 m1 a e2 m2 p x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vptr p x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vptr p y) ->
- eval_expr ge sp le e1 m1 (cmp c a b) e3 m3 (Val.of_bool (Int.cmp c x y)).
+ forall sp le c e1 m1 a t1 e2 m2 p x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vptr p x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vptr p y) ->
+ eval_expr ge sp le e1 m1 (cmp c a b) (t1**t2) e3 m3 (Val.of_bool (Int.cmp c x y)).
Proof.
TrivialOp cmp.
simpl. unfold eq_block. rewrite zeq_true.
@@ -983,32 +1034,32 @@ Proof.
Qed.
Theorem eval_cmpu:
- forall sp le c e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vint x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vint y) ->
- eval_expr ge sp le e1 m1 (cmpu c a b) e3 m3 (Val.of_bool (Int.cmpu c x y)).
+ forall sp le c e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vint x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vint y) ->
+ eval_expr ge sp le e1 m1 (cmpu c a b) (t1**t2) e3 m3 (Val.of_bool (Int.cmpu c x y)).
Proof.
TrivialOp cmpu.
simpl. case (Int.cmpu c x y); auto.
Qed.
Theorem eval_cmpf:
- forall sp le c e1 m1 a e2 m2 x b e3 m3 y,
- eval_expr ge sp le e1 m1 a e2 m2 (Vfloat x) ->
- eval_expr ge sp le e2 m2 b e3 m3 (Vfloat y) ->
- eval_expr ge sp le e1 m1 (cmpf c a b) e3 m3 (Val.of_bool (Float.cmp c x y)).
+ forall sp le c e1 m1 a t1 e2 m2 x b t2 e3 m3 y,
+ eval_expr ge sp le e1 m1 a t1 e2 m2 (Vfloat x) ->
+ eval_expr ge sp le e2 m2 b t2 e3 m3 (Vfloat y) ->
+ eval_expr ge sp le e1 m1 (cmpf c a b) (t1**t2) e3 m3 (Val.of_bool (Float.cmp c x y)).
Proof.
TrivialOp cmpf.
simpl. case (Float.cmp c x y); auto.
Qed.
Lemma eval_base_condition_of_expr:
- forall sp le a e1 m1 e2 m2 v (b: bool),
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall sp le a e1 m1 t e2 m2 v (b: bool),
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
Val.bool_of_val v b ->
eval_condexpr ge sp le e1 m1
(CEcond (Ccompuimm Cne Int.zero) (a ::: Enil))
- e2 m2 b.
+ t e2 m2 b.
Proof.
intros.
eapply eval_CEcond. eauto with evalexpr.
@@ -1016,62 +1067,60 @@ Proof.
Qed.
Lemma eval_condition_of_expr:
- forall a sp le e1 m1 e2 m2 v (b: bool),
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall a sp le e1 m1 t e2 m2 v (b: bool),
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
Val.bool_of_val v b ->
- eval_condexpr ge sp le e1 m1 (condexpr_of_expr a) e2 m2 b.
+ eval_condexpr ge sp le e1 m1 (condexpr_of_expr a) t e2 m2 b.
Proof.
induction a; simpl; intros;
try (eapply eval_base_condition_of_expr; eauto; fail).
destruct o; try (eapply eval_base_condition_of_expr; eauto; fail).
- destruct e. inversion H. subst sp0 le0 e m op al e0 m0 v0.
- inversion H7. subst sp0 le0 e m e1 m1 vl.
- simpl in H11. inversion H11; subst v.
+ destruct e. InvEval H. inversion XX4; subst v.
inversion H0.
rewrite Int.eq_false; auto. constructor.
subst i; rewrite Int.eq_true. constructor.
eapply eval_base_condition_of_expr; eauto.
- inversion H. eapply eval_CEcond; eauto. simpl in H11.
+ inversion H. subst. eapply eval_CEcond; eauto. simpl in H12.
destruct (eval_condition c vl); try discriminate.
- destruct b0; inversion H11; subst v0; subst v; inversion H0; congruence.
+ destruct b0; inversion H12; subst; inversion H0; congruence.
- inversion H.
+ inversion H. subst.
destruct v1; eauto with evalexpr.
Qed.
Theorem eval_conditionalexpr_true:
- forall sp le e1 m1 a1 e2 m2 v1 a2 e3 m3 v2 a3,
- eval_expr ge sp le e1 m1 a1 e2 m2 v1 ->
+ forall sp le e1 m1 a1 t1 e2 m2 v1 t2 a2 e3 m3 v2 a3,
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 ->
Val.is_true v1 ->
- eval_expr ge sp le e2 m2 a2 e3 m3 v2 ->
- eval_expr ge sp le e1 m1 (conditionalexpr a1 a2 a3) e3 m3 v2.
+ eval_expr ge sp le e2 m2 a2 t2 e3 m3 v2 ->
+ eval_expr ge sp le e1 m1 (conditionalexpr a1 a2 a3) (t1**t2) e3 m3 v2.
Proof.
intros; unfold conditionalexpr.
- apply eval_Econdition with e2 m2 true; auto.
+ apply eval_Econdition with t1 e2 m2 true t2; auto.
eapply eval_condition_of_expr; eauto with valboolof.
Qed.
Theorem eval_conditionalexpr_false:
- forall sp le e1 m1 a1 e2 m2 v1 a2 e3 m3 v2 a3,
- eval_expr ge sp le e1 m1 a1 e2 m2 v1 ->
+ forall sp le e1 m1 a1 t1 e2 m2 v1 a2 t2 e3 m3 v2 a3,
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 ->
Val.is_false v1 ->
- eval_expr ge sp le e2 m2 a3 e3 m3 v2 ->
- eval_expr ge sp le e1 m1 (conditionalexpr a1 a2 a3) e3 m3 v2.
+ eval_expr ge sp le e2 m2 a3 t2 e3 m3 v2 ->
+ eval_expr ge sp le e1 m1 (conditionalexpr a1 a2 a3) (t1**t2) e3 m3 v2.
Proof.
intros; unfold conditionalexpr.
- apply eval_Econdition with e2 m2 false; auto.
+ apply eval_Econdition with t1 e2 m2 false t2; auto.
eapply eval_condition_of_expr; eauto with valboolof.
Qed.
Lemma eval_addressing:
- forall sp le e1 m1 a e2 m2 v b ofs,
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall sp le e1 m1 a t e2 m2 v b ofs,
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
v = Vptr b ofs ->
match addressing a with (mode, args) =>
exists vl,
- eval_exprlist ge sp le e1 m1 args e2 m2 vl /\
+ eval_exprlist ge sp le e1 m1 args t e2 m2 vl /\
eval_addressing ge sp mode vl = Some v
end.
Proof.
@@ -1080,19 +1129,13 @@ Proof.
simpl. auto.
InvEval H. exists (@nil val). split. eauto with evalexpr.
simpl. auto.
- InvEval H. FuncInv.
+ InvEval H. InvEval EV. rewrite E0_left in TR. subst t1. FuncInv.
congruence.
- InvEval EV.
+ destruct (Genv.find_symbol ge s); congruence.
exists (Vint i0 :: nil). split. eauto with evalexpr.
simpl. subst v. destruct (Genv.find_symbol ge s). congruence.
discriminate.
InvEval H. FuncInv.
- destruct (Genv.find_symbol ge s); congruence.
- InvEval EV.
- exists (Vint i0 :: nil). split. eauto with evalexpr.
- simpl. destruct (Genv.find_symbol ge s). congruence.
- discriminate.
- InvEval H. FuncInv.
congruence.
exists (Vptr b0 i :: nil). split. eauto with evalexpr.
simpl. congruence.
@@ -1108,56 +1151,56 @@ Proof.
Qed.
Theorem eval_load:
- forall sp le e1 m1 a e2 m2 v chunk v',
- eval_expr ge sp le e1 m1 a e2 m2 v ->
+ forall sp le e1 m1 a t e2 m2 v chunk v',
+ eval_expr ge sp le e1 m1 a t e2 m2 v ->
Mem.loadv chunk m2 v = Some v' ->
- eval_expr ge sp le e1 m1 (load chunk a) e2 m2 v'.
+ eval_expr ge sp le e1 m1 (load chunk a) t e2 m2 v'.
Proof.
intros. generalize H0; destruct v; simpl; intro; try discriminate.
unfold load.
- generalize (eval_addressing _ _ _ _ _ _ _ _ _ _ H (refl_equal _)).
+ generalize (eval_addressing _ _ _ _ _ _ _ _ _ _ _ H (refl_equal _)).
destruct (addressing a). intros [vl [EV EQ]].
eapply eval_Eload; eauto.
Qed.
Theorem eval_store:
- forall sp le e1 m1 a1 e2 m2 v1 a2 e3 m3 v2 chunk m4,
- eval_expr ge sp le e1 m1 a1 e2 m2 v1 ->
- eval_expr ge sp le e2 m2 a2 e3 m3 v2 ->
+ forall sp le e1 m1 a1 t1 e2 m2 v1 a2 t2 e3 m3 v2 chunk m4,
+ eval_expr ge sp le e1 m1 a1 t1 e2 m2 v1 ->
+ eval_expr ge sp le e2 m2 a2 t2 e3 m3 v2 ->
Mem.storev chunk m3 v1 v2 = Some m4 ->
- eval_expr ge sp le e1 m1 (store chunk a1 a2) e3 m4 v2.
+ eval_expr ge sp le e1 m1 (store chunk a1 a2) (t1**t2) e3 m4 v2.
Proof.
intros. generalize H1; destruct v1; simpl; intro; try discriminate.
unfold store.
- generalize (eval_addressing _ _ _ _ _ _ _ _ _ _ H (refl_equal _)).
+ generalize (eval_addressing _ _ _ _ _ _ _ _ _ _ _ H (refl_equal _)).
destruct (addressing a1). intros [vl [EV EQ]].
eapply eval_Estore; eauto.
Qed.
Theorem exec_ifthenelse_true:
- forall sp e1 m1 a e2 m2 v ifso ifnot e3 m3 out,
- eval_expr ge sp nil e1 m1 a e2 m2 v ->
+ forall sp e1 m1 a t1 e2 m2 v ifso ifnot t2 e3 m3 out,
+ eval_expr ge sp nil e1 m1 a t1 e2 m2 v ->
Val.is_true v ->
- exec_stmt ge sp e2 m2 ifso e3 m3 out ->
- exec_stmt ge sp e1 m1 (ifthenelse a ifso ifnot) e3 m3 out.
+ exec_stmt ge sp e2 m2 ifso t2 e3 m3 out ->
+ exec_stmt ge sp e1 m1 (ifthenelse a ifso ifnot) (t1**t2) e3 m3 out.
Proof.
intros. unfold ifthenelse.
- apply exec_Sifthenelse with e2 m2 true.
+ apply exec_Sifthenelse with t1 e2 m2 true t2.
eapply eval_condition_of_expr; eauto with valboolof.
- auto.
+ auto. auto.
Qed.
Theorem exec_ifthenelse_false:
- forall sp e1 m1 a e2 m2 v ifso ifnot e3 m3 out,
- eval_expr ge sp nil e1 m1 a e2 m2 v ->
+ forall sp e1 m1 a t1 e2 m2 v ifso ifnot t2 e3 m3 out,
+ eval_expr ge sp nil e1 m1 a t1 e2 m2 v ->
Val.is_false v ->
- exec_stmt ge sp e2 m2 ifnot e3 m3 out ->
- exec_stmt ge sp e1 m1 (ifthenelse a ifso ifnot) e3 m3 out.
+ exec_stmt ge sp e2 m2 ifnot t2 e3 m3 out ->
+ exec_stmt ge sp e1 m1 (ifthenelse a ifso ifnot) (t1**t2) e3 m3 out.
Proof.
intros. unfold ifthenelse.
- apply exec_Sifthenelse with e2 m2 false.
+ apply exec_Sifthenelse with t1 e2 m2 false t2.
eapply eval_condition_of_expr; eauto with valboolof.
- auto.
+ auto. auto.
Qed.
End CMCONSTR.
diff --git a/backend/Cminor.v b/backend/Cminor.v
index 826c5298..9eed0091 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -5,6 +5,7 @@ Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Floats.
+Require Import Events.
Require Import Values.
Require Import Mem.
Require Import Op.
@@ -35,6 +36,7 @@ Inductive expr : Set :=
| Econdition : condexpr -> expr -> expr -> expr
| Elet : expr -> expr -> expr
| Eletvar : nat -> expr
+ | Ealloc : expr -> expr
with condexpr : Set :=
| CEtrue: condexpr
@@ -77,7 +79,14 @@ Record function : Set := mkfunction {
fn_body: stmt
}.
-Definition program := AST.program function.
+Definition fundef := AST.fundef function.
+Definition program := AST.program fundef.
+
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => f.(fn_sig)
+ | External ef => ef.(ef_sig)
+ end.
(** * Operational semantics *)
@@ -120,7 +129,7 @@ Fixpoint switch_target (n: int) (dfl: nat) (cases: list (int * nat))
- [lenv]: let environments, map de Bruijn indices to values.
*)
-Definition genv := Genv.t function.
+Definition genv := Genv.t fundef.
Definition env := PTree.t val.
Definition letenv := list val.
@@ -157,56 +166,65 @@ Variable ge: genv.
Inductive eval_expr:
val -> letenv ->
- env -> mem -> expr ->
+ env -> mem -> expr -> trace ->
env -> mem -> val -> Prop :=
| eval_Evar:
forall sp le e m id v,
PTree.get id e = Some v ->
- eval_expr sp le e m (Evar id) e m v
+ eval_expr sp le e m (Evar id) E0 e m v
| eval_Eassign:
- forall sp le e m id a e1 m1 v,
- eval_expr sp le e m a e1 m1 v ->
- eval_expr sp le e m (Eassign id a) (PTree.set id v e1) m1 v
+ forall sp le e m id a t e1 m1 v,
+ eval_expr sp le e m a t e1 m1 v ->
+ eval_expr sp le e m (Eassign id a) t (PTree.set id v e1) m1 v
| eval_Eop:
- forall sp le e m op al e1 m1 vl v,
- eval_exprlist sp le e m al e1 m1 vl ->
+ forall sp le e m op al t e1 m1 vl v,
+ eval_exprlist sp le e m al t e1 m1 vl ->
eval_operation ge sp op vl = Some v ->
- eval_expr sp le e m (Eop op al) e1 m1 v
+ eval_expr sp le e m (Eop op al) t e1 m1 v
| eval_Eload:
- forall sp le e m chunk addr al e1 m1 v vl a,
- eval_exprlist sp le e m al e1 m1 vl ->
+ forall sp le e m chunk addr al t e1 m1 v vl a,
+ eval_exprlist sp le e m al t e1 m1 vl ->
eval_addressing ge sp addr vl = Some a ->
Mem.loadv chunk m1 a = Some v ->
- eval_expr sp le e m (Eload chunk addr al) e1 m1 v
+ eval_expr sp le e m (Eload chunk addr al) t e1 m1 v
| eval_Estore:
- forall sp le e m chunk addr al b e1 m1 vl e2 m2 m3 v a,
- eval_exprlist sp le e m al e1 m1 vl ->
- eval_expr sp le e1 m1 b e2 m2 v ->
+ forall sp le e m chunk addr al b t t1 e1 m1 vl t2 e2 m2 m3 v a,
+ eval_exprlist sp le e m al t1 e1 m1 vl ->
+ eval_expr sp le e1 m1 b t2 e2 m2 v ->
eval_addressing ge sp addr vl = Some a ->
Mem.storev chunk m2 a v = Some m3 ->
- eval_expr sp le e m (Estore chunk addr al b) e2 m3 v
+ t = t1 ** t2 ->
+ eval_expr sp le e m (Estore chunk addr al b) t e2 m3 v
| eval_Ecall:
- forall sp le e m sig a bl e1 e2 m1 m2 m3 vf vargs vres f,
- eval_expr sp le e m a e1 m1 vf ->
- eval_exprlist sp le e1 m1 bl e2 m2 vargs ->
+ forall sp le e m sig a bl t t1 e1 m1 t2 e2 m2 t3 m3 vf vargs vres f,
+ eval_expr sp le e m a t1 e1 m1 vf ->
+ eval_exprlist sp le e1 m1 bl t2 e2 m2 vargs ->
Genv.find_funct ge vf = Some f ->
- f.(fn_sig) = sig ->
- eval_funcall m2 f vargs m3 vres ->
- eval_expr sp le e m (Ecall sig a bl) e2 m3 vres
+ funsig f = sig ->
+ eval_funcall m2 f vargs t3 m3 vres ->
+ t = t1 ** t2 ** t3 ->
+ eval_expr sp le e m (Ecall sig a bl) t e2 m3 vres
| eval_Econdition:
- forall sp le e m a b c e1 m1 v1 e2 m2 v2,
- eval_condexpr sp le e m a e1 m1 v1 ->
- eval_expr sp le e1 m1 (if v1 then b else c) e2 m2 v2 ->
- eval_expr sp le e m (Econdition a b c) e2 m2 v2
+ forall sp le e m a b c t t1 e1 m1 v1 t2 e2 m2 v2,
+ eval_condexpr sp le e m a t1 e1 m1 v1 ->
+ eval_expr sp le e1 m1 (if v1 then b else c) t2 e2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr sp le e m (Econdition a b c) t e2 m2 v2
| eval_Elet:
- forall sp le e m a b e1 m1 v1 e2 m2 v2,
- eval_expr sp le e m a e1 m1 v1 ->
- eval_expr sp (v1::le) e1 m1 b e2 m2 v2 ->
- eval_expr sp le e m (Elet a b) e2 m2 v2
+ forall sp le e m a b t t1 e1 m1 v1 t2 e2 m2 v2,
+ eval_expr sp le e m a t1 e1 m1 v1 ->
+ eval_expr sp (v1::le) e1 m1 b t2 e2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr sp le e m (Elet a b) t e2 m2 v2
| eval_Eletvar:
forall sp le e m n v,
nth_error le n = Some v ->
- eval_expr sp le e m (Eletvar n) e m v
+ eval_expr sp le e m (Eletvar n) E0 e m v
+ | eval_Ealloc:
+ forall sp le e m a t e1 m1 n m2 b,
+ eval_expr sp le e m a t e1 m1 (Vint n) ->
+ Mem.alloc m1 0 (Int.signed n) = (m2, b) ->
+ eval_expr sp le e m (Ealloc a) t e1 m2 (Vptr b Int.zero)
(** Evaluation of a condition expression:
[eval_condexpr ge sp le e m a e' m' b]
@@ -216,24 +234,25 @@ Inductive eval_expr:
with eval_condexpr:
val -> letenv ->
- env -> mem -> condexpr ->
+ env -> mem -> condexpr -> trace ->
env -> mem -> bool -> Prop :=
| eval_CEtrue:
forall sp le e m,
- eval_condexpr sp le e m CEtrue e m true
+ eval_condexpr sp le e m CEtrue E0 e m true
| eval_CEfalse:
forall sp le e m,
- eval_condexpr sp le e m CEfalse e m false
+ eval_condexpr sp le e m CEfalse E0 e m false
| eval_CEcond:
- forall sp le e m cond al e1 m1 vl b,
- eval_exprlist sp le e m al e1 m1 vl ->
+ forall sp le e m cond al t1 e1 m1 vl b,
+ eval_exprlist sp le e m al t1 e1 m1 vl ->
eval_condition cond vl = Some b ->
- eval_condexpr sp le e m (CEcond cond al) e1 m1 b
+ eval_condexpr sp le e m (CEcond cond al) t1 e1 m1 b
| eval_CEcondition:
- forall sp le e m a b c e1 m1 vb1 e2 m2 vb2,
- eval_condexpr sp le e m a e1 m1 vb1 ->
- eval_condexpr sp le e1 m1 (if vb1 then b else c) e2 m2 vb2 ->
- eval_condexpr sp le e m (CEcondition a b c) e2 m2 vb2
+ forall sp le e m a b c t t1 e1 m1 vb1 t2 e2 m2 vb2,
+ eval_condexpr sp le e m a t1 e1 m1 vb1 ->
+ eval_condexpr sp le e1 m1 (if vb1 then b else c) t2 e2 m2 vb2 ->
+ t = t1 ** t2 ->
+ eval_condexpr sp le e m (CEcondition a b c) t e2 m2 vb2
(** Evaluation of a list of expressions:
[eval_exprlist ge sp le al m a e' m' vl]
@@ -244,16 +263,17 @@ with eval_condexpr:
with eval_exprlist:
val -> letenv ->
- env -> mem -> exprlist ->
+ env -> mem -> exprlist -> trace ->
env -> mem -> list val -> Prop :=
| eval_Enil:
forall sp le e m,
- eval_exprlist sp le e m Enil e m nil
+ eval_exprlist sp le e m Enil E0 e m nil
| eval_Econs:
- forall sp le e m a bl e1 m1 v e2 m2 vl,
- eval_expr sp le e m a e1 m1 v ->
- eval_exprlist sp le e1 m1 bl e2 m2 vl ->
- eval_exprlist sp le e m (Econs a bl) e2 m2 (v :: vl)
+ forall sp le e m a bl t t1 e1 m1 v t2 e2 m2 vl,
+ eval_expr sp le e m a t1 e1 m1 v ->
+ eval_exprlist sp le e1 m1 bl t2 e2 m2 vl ->
+ t = t1 ** t2 ->
+ eval_exprlist sp le e m (Econs a bl) t e2 m2 (v :: vl)
(** Evaluation of a function invocation: [eval_funcall ge m f args m' res]
means that the function [f], applied to the arguments [args] in
@@ -261,15 +281,19 @@ with eval_exprlist:
*)
with eval_funcall:
- mem -> function -> list val ->
+ mem -> fundef -> list val -> trace ->
mem -> val -> Prop :=
- | eval_funcall_intro:
- forall m f vargs m1 sp e e2 m2 out vres,
+ | eval_funcall_internal:
+ forall m f vargs m1 sp e t e2 m2 out vres,
Mem.alloc m 0 f.(fn_stackspace) = (m1, sp) ->
set_locals f.(fn_vars) (set_params vargs f.(fn_params)) = e ->
- exec_stmt (Vptr sp Int.zero) e m1 f.(fn_body) e2 m2 out ->
+ exec_stmt (Vptr sp Int.zero) e m1 f.(fn_body) t e2 m2 out ->
outcome_result_value out f.(fn_sig).(sig_res) vres ->
- eval_funcall m f vargs (Mem.free m2 sp) vres
+ eval_funcall m (Internal f) vargs t (Mem.free m2 sp) vres
+ | eval_funcall_external:
+ forall ef m args t res,
+ event_match ef args t res ->
+ eval_funcall m (External ef) args t m res
(** Execution of a statement: [exec_stmt ge sp e m s e' m' out]
means that statement [s] executes with outcome [out].
@@ -277,59 +301,62 @@ with eval_funcall:
with exec_stmt:
val ->
- env -> mem -> stmt ->
+ env -> mem -> stmt -> trace ->
env -> mem -> outcome -> Prop :=
| exec_Sskip:
forall sp e m,
- exec_stmt sp e m Sskip e m Out_normal
+ exec_stmt sp e m Sskip E0 e m Out_normal
| exec_Sexpr:
- forall sp e m a e1 m1 v,
- eval_expr sp nil e m a e1 m1 v ->
- exec_stmt sp e m (Sexpr a) e1 m1 Out_normal
+ forall sp e m a t e1 m1 v,
+ eval_expr sp nil e m a t e1 m1 v ->
+ exec_stmt sp e m (Sexpr a) t e1 m1 Out_normal
| exec_Sifthenelse:
- forall sp e m a s1 s2 e1 m1 v1 e2 m2 out,
- eval_condexpr sp nil e m a e1 m1 v1 ->
- exec_stmt sp e1 m1 (if v1 then s1 else s2) e2 m2 out ->
- exec_stmt sp e m (Sifthenelse a s1 s2) e2 m2 out
+ forall sp e m a s1 s2 t t1 e1 m1 v1 t2 e2 m2 out,
+ eval_condexpr sp nil e m a t1 e1 m1 v1 ->
+ exec_stmt sp e1 m1 (if v1 then s1 else s2) t2 e2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt sp e m (Sifthenelse a s1 s2) t e2 m2 out
| exec_Sseq_continue:
- forall sp e m s1 e1 m1 s2 e2 m2 out,
- exec_stmt sp e m s1 e1 m1 Out_normal ->
- exec_stmt sp e1 m1 s2 e2 m2 out ->
- exec_stmt sp e m (Sseq s1 s2) e2 m2 out
+ forall sp e m t s1 t1 e1 m1 s2 t2 e2 m2 out,
+ exec_stmt sp e m s1 t1 e1 m1 Out_normal ->
+ exec_stmt sp e1 m1 s2 t2 e2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt sp e m (Sseq s1 s2) t e2 m2 out
| exec_Sseq_stop:
- forall sp e m s1 s2 e1 m1 out,
- exec_stmt sp e m s1 e1 m1 out ->
+ forall sp e m t s1 s2 e1 m1 out,
+ exec_stmt sp e m s1 t e1 m1 out ->
out <> Out_normal ->
- exec_stmt sp e m (Sseq s1 s2) e1 m1 out
+ exec_stmt sp e m (Sseq s1 s2) t e1 m1 out
| exec_Sloop_loop:
- forall sp e m s e1 m1 e2 m2 out,
- exec_stmt sp e m s e1 m1 Out_normal ->
- exec_stmt sp e1 m1 (Sloop s) e2 m2 out ->
- exec_stmt sp e m (Sloop s) e2 m2 out
+ forall sp e m s t t1 e1 m1 t2 e2 m2 out,
+ exec_stmt sp e m s t1 e1 m1 Out_normal ->
+ exec_stmt sp e1 m1 (Sloop s) t2 e2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt sp e m (Sloop s) t e2 m2 out
| exec_Sloop_stop:
- forall sp e m s e1 m1 out,
- exec_stmt sp e m s e1 m1 out ->
+ forall sp e m t s e1 m1 out,
+ exec_stmt sp e m s t e1 m1 out ->
out <> Out_normal ->
- exec_stmt sp e m (Sloop s) e1 m1 out
+ exec_stmt sp e m (Sloop s) t e1 m1 out
| exec_Sblock:
- forall sp e m s e1 m1 out,
- exec_stmt sp e m s e1 m1 out ->
- exec_stmt sp e m (Sblock s) e1 m1 (outcome_block out)
+ forall sp e m s t e1 m1 out,
+ exec_stmt sp e m s t e1 m1 out ->
+ exec_stmt sp e m (Sblock s) t e1 m1 (outcome_block out)
| exec_Sexit:
forall sp e m n,
- exec_stmt sp e m (Sexit n) e m (Out_exit n)
+ exec_stmt sp e m (Sexit n) E0 e m (Out_exit n)
| exec_Sswitch:
- forall sp e m a cases default e1 m1 n,
- eval_expr sp nil e m a e1 m1 (Vint n) ->
+ forall sp e m a cases default t1 e1 m1 n,
+ eval_expr sp nil e m a t1 e1 m1 (Vint n) ->
exec_stmt sp e m (Sswitch a cases default)
- e1 m1 (Out_exit (switch_target n default cases))
+ t1 e1 m1 (Out_exit (switch_target n default cases))
| exec_Sreturn_none:
forall sp e m,
- exec_stmt sp e m (Sreturn None) e m (Out_return None)
+ exec_stmt sp e m (Sreturn None) E0 e m (Out_return None)
| exec_Sreturn_some:
- forall sp e m a e1 m1 v,
- eval_expr sp nil e m a e1 m1 v ->
- exec_stmt sp e m (Sreturn (Some a)) e1 m1 (Out_return (Some v)).
+ forall sp e m a t e1 m1 v,
+ eval_expr sp nil e m a t e1 m1 v ->
+ exec_stmt sp e m (Sreturn (Some a)) t e1 m1 (Out_return (Some v)).
End RELSEM.
@@ -337,12 +364,12 @@ End RELSEM.
holds if the application of [p]'s main function to no arguments
in the initial memory state for [p] eventually returns value [r]. *)
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
exists b, exists f, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- f.(fn_sig) = mksignature nil (Some Tint) /\
- eval_funcall ge m0 f nil m r.
+ funsig f = mksignature nil (Some Tint) /\
+ eval_funcall ge m0 f nil t m r.
diff --git a/backend/Cminorgen.v b/backend/Cminorgen.v
index cb889928..4c611b44 100644
--- a/backend/Cminorgen.v
+++ b/backend/Cminorgen.v
@@ -104,18 +104,27 @@ Definition make_cast (chunk: memory_chunk) (e: expr): expr :=
Definition make_load (chunk: memory_chunk) (e: expr): expr :=
Cmconstr.load chunk e.
-(** In Csharpminor, the value of a store expression is the stored data
- normalized to the memory chunk. In Cminor, it is the stored data.
- For the translation, we could normalize before storing. However,
- the memory model performs automatic normalization of the stored
- data. It is therefore correct to store the data as is, then
- normalize the result value of the store expression. This is more
- efficient in general because often the result value is ignored:
- the normalization code will therefore be eliminated later as dead
- code. *)
-
-Definition make_store (chunk: memory_chunk) (e1 e2: expr): expr :=
- make_cast chunk (Cmconstr.store chunk e1 e2).
+Definition store_arg (chunk: memory_chunk) (e: expr) : expr :=
+ match e with
+ | Eop op (Econs e1 Enil) =>
+ match op with
+ | Ocast8signed =>
+ match chunk with Mint8signed => e1 | _ => e end
+ | Ocast8unsigned =>
+ match chunk with Mint8unsigned => e1 | _ => e end
+ | Ocast16signed =>
+ match chunk with Mint16signed => e1 | _ => e end
+ | Ocast16unsigned =>
+ match chunk with Mint16unsigned => e1 | _ => e end
+ | Osingleoffloat =>
+ match chunk with Mfloat32 => e1 | _ => e end
+ | _ => e
+ end
+ | _ => e
+ end.
+
+Definition make_store (chunk: memory_chunk) (e1 e2: expr): stmt :=
+ Sexpr(Cmconstr.store chunk e1 (store_arg chunk e2)).
Definition make_stackaddr (ofs: Z): expr :=
Eop (Oaddrstack (Int.repr ofs)) Enil.
@@ -156,10 +165,10 @@ Definition var_get (cenv: compilenv) (id: ident): option expr :=
None
end.
-Definition var_set (cenv: compilenv) (id: ident) (rhs: expr): option expr :=
+Definition var_set (cenv: compilenv) (id: ident) (rhs: expr): option stmt :=
match PMap.get id cenv with
| Var_local chunk =>
- Some(Eassign id (make_cast chunk rhs))
+ Some(Sexpr(Eassign id (make_cast chunk rhs)))
| Var_stack_scalar chunk ofs =>
Some(make_store chunk (make_stackaddr ofs) rhs)
| Var_global_scalar chunk =>
@@ -199,16 +208,10 @@ Fixpoint transl_expr (cenv: compilenv) (e: Csharpminor.expr)
match e with
| Csharpminor.Evar id => var_get cenv id
| Csharpminor.Eaddrof id => var_addr cenv id
- | Csharpminor.Eassign id e =>
- do te <- transl_expr cenv e; var_set cenv id te
| Csharpminor.Eop op el =>
do tel <- transl_exprlist cenv el; make_op op tel
| Csharpminor.Eload chunk e =>
do te <- transl_expr cenv e; Some (make_load chunk te)
- | Csharpminor.Estore chunk e1 e2 =>
- do te1 <- transl_expr cenv e1;
- do te2 <- transl_expr cenv e2;
- Some (make_store chunk te1 te2)
| Csharpminor.Ecall sig e el =>
do te <- transl_expr cenv e;
do tel <- transl_exprlist cenv el;
@@ -224,6 +227,9 @@ Fixpoint transl_expr (cenv: compilenv) (e: Csharpminor.expr)
Some (Elet te1 te2)
| Csharpminor.Eletvar n =>
Some (Eletvar n)
+ | Csharpminor.Ealloc e =>
+ do te <- transl_expr cenv e;
+ Some (Ealloc te)
end
with transl_exprlist (cenv: compilenv) (el: Csharpminor.exprlist)
@@ -246,6 +252,12 @@ Fixpoint transl_stmt (cenv: compilenv) (s: Csharpminor.stmt)
Some Sskip
| Csharpminor.Sexpr e =>
do te <- transl_expr cenv e; Some(Sexpr te)
+ | Csharpminor.Sassign id e =>
+ do te <- transl_expr cenv e; var_set cenv id te
+ | Csharpminor.Sstore chunk e1 e2 =>
+ do te1 <- transl_expr cenv e1;
+ do te2 <- transl_expr cenv e2;
+ Some (make_store chunk te1 te2)
| Csharpminor.Sseq s1 s2 =>
do ts1 <- transl_stmt cenv s1;
do ts2 <- transl_stmt cenv s2;
@@ -280,11 +292,8 @@ Fixpoint addr_taken_expr (e: Csharpminor.expr): Identset.t :=
match e with
| Csharpminor.Evar id => Identset.empty
| Csharpminor.Eaddrof id => Identset.add id Identset.empty
- | Csharpminor.Eassign id e => addr_taken_expr e
| Csharpminor.Eop op el => addr_taken_exprlist el
| Csharpminor.Eload chunk e => addr_taken_expr e
- | Csharpminor.Estore chunk e1 e2 =>
- Identset.union (addr_taken_expr e1) (addr_taken_expr e2)
| Csharpminor.Ecall sig e el =>
Identset.union (addr_taken_expr e) (addr_taken_exprlist el)
| Csharpminor.Econdition e1 e2 e3 =>
@@ -293,6 +302,7 @@ Fixpoint addr_taken_expr (e: Csharpminor.expr): Identset.t :=
| Csharpminor.Elet e1 e2 =>
Identset.union (addr_taken_expr e1) (addr_taken_expr e2)
| Csharpminor.Eletvar n => Identset.empty
+ | Csharpminor.Ealloc e => addr_taken_expr e
end
with addr_taken_exprlist (e: Csharpminor.exprlist): Identset.t :=
@@ -306,6 +316,9 @@ Fixpoint addr_taken_stmt (s: Csharpminor.stmt): Identset.t :=
match s with
| Csharpminor.Sskip => Identset.empty
| Csharpminor.Sexpr e => addr_taken_expr e
+ | Csharpminor.Sassign id e => addr_taken_expr e
+ | Csharpminor.Sstore chunk e1 e2 =>
+ Identset.union (addr_taken_expr e1) (addr_taken_expr e2)
| Csharpminor.Sseq s1 s2 =>
Identset.union (addr_taken_stmt s1) (addr_taken_stmt s2)
| Csharpminor.Sifthenelse e s1 s2 =>
@@ -362,10 +375,10 @@ Definition build_compilenv
(globenv, 0).
Definition assign_global_variable
- (ce: compilenv) (id_vi: ident * var_kind) : compilenv :=
- match id_vi with
- | (id, Vscalar chunk) => PMap.set id (Var_global_scalar chunk) ce
- | (id, Varray sz) => PMap.set id Var_global_array ce
+ (ce: compilenv) (info: ident * var_kind * list init_data) : compilenv :=
+ match info with
+ | (id, Vscalar chunk, _) => PMap.set id (Var_global_scalar chunk) ce
+ | (id, Varray _, _) => PMap.set id Var_global_array ce
end.
Definition build_global_compilenv (p: Csharpminor.program) : compilenv :=
@@ -387,7 +400,7 @@ Fixpoint store_parameters
Sseq (Sexpr (Eassign id (make_cast chunk (Evar id))))
(store_parameters cenv rem)
| Var_stack_scalar chunk ofs =>
- Sseq (Sexpr (make_store chunk (make_stackaddr ofs) (Evar id)))
+ Sseq (make_store chunk (make_stackaddr ofs) (Evar id))
(store_parameters cenv rem)
| _ =>
Sskip (* should never happen *)
@@ -412,6 +425,9 @@ Definition transl_function
(Sseq (store_parameters cenv f.(Csharpminor.fn_params)) tbody))
else None.
+Definition transl_fundef (gce: compilenv) (f: Csharpminor.fundef): option fundef :=
+ transf_partial_fundef (transl_function gce) f.
+
Definition transl_program (p: Csharpminor.program) : option program :=
let gce := build_global_compilenv p in
- transform_partial_program (transl_function gce) (program_of_program p).
+ transform_partial_program (transl_fundef gce) (program_of_program p).
diff --git a/backend/Cminorgenproof.v b/backend/Cminorgenproof.v
index 7b3bc9bb..7820095a 100644
--- a/backend/Cminorgenproof.v
+++ b/backend/Cminorgenproof.v
@@ -8,6 +8,7 @@ Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Csharpminor.
Require Import Op.
@@ -29,38 +30,51 @@ Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
intro. unfold ge, tge.
- apply Genv.find_symbol_transf_partial with (transl_function gce).
+ apply Genv.find_symbol_transf_partial with (transl_fundef gce).
exact TRANSL.
Qed.
Lemma function_ptr_translated:
- forall (b: block) (f: Csharpminor.function),
+ forall (b: block) (f: Csharpminor.fundef),
Genv.find_funct_ptr ge b = Some f ->
exists tf,
- Genv.find_funct_ptr tge b = Some tf /\ transl_function gce f = Some tf.
+ Genv.find_funct_ptr tge b = Some tf /\ transl_fundef gce f = Some tf.
Proof.
intros.
generalize
- (Genv.find_funct_ptr_transf_partial (transl_function gce) TRANSL H).
- case (transl_function gce f).
+ (Genv.find_funct_ptr_transf_partial (transl_fundef gce) TRANSL H).
+ case (transl_fundef gce f).
intros tf [A B]. exists tf. tauto.
intros [A B]. elim B. reflexivity.
Qed.
Lemma functions_translated:
- forall (v: val) (f: Csharpminor.function),
+ forall (v: val) (f: Csharpminor.fundef),
Genv.find_funct ge v = Some f ->
exists tf,
- Genv.find_funct tge v = Some tf /\ transl_function gce f = Some tf.
+ Genv.find_funct tge v = Some tf /\ transl_fundef gce f = Some tf.
Proof.
intros.
generalize
- (Genv.find_funct_transf_partial (transl_function gce) TRANSL H).
- case (transl_function gce f).
+ (Genv.find_funct_transf_partial (transl_fundef gce) TRANSL H).
+ case (transl_fundef gce f).
intros tf [A B]. exists tf. tauto.
intros [A B]. elim B. reflexivity.
Qed.
+Lemma sig_preserved:
+ forall f tf,
+ transl_fundef gce f = Some tf ->
+ Cminor.funsig tf = Csharpminor.funsig f.
+Proof.
+ intros until tf; destruct f; simpl.
+ unfold transl_function. destruct (build_compilenv gce f).
+ case (zle z Int.max_signed); try congruence.
+ intro. case (transl_stmt c (Csharpminor.fn_body f)); simpl; try congruence.
+ intros. inversion H. reflexivity.
+ intro. inversion H; reflexivity.
+Qed.
+
Definition global_compilenv_match (ce: compilenv) (gv: gvarenv) : Prop :=
forall id,
match ce!!id with
@@ -72,9 +86,9 @@ Definition global_compilenv_match (ce: compilenv) (gv: gvarenv) : Prop :=
Lemma global_compilenv_charact:
global_compilenv_match gce (global_var_env prog).
Proof.
- set (mkgve := fun gv vars =>
+ set (mkgve := fun gv (vars: list (ident * var_kind * list init_data)) =>
List.fold_left
- (fun gv (id_vi: ident * var_kind) => PTree.set (fst id_vi) (snd id_vi) gv)
+ (fun gve x => match x with (id, k, init) => PTree.set id k gve end)
vars gv).
assert (forall vars gv ce,
global_compilenv_match ce gv ->
@@ -83,7 +97,7 @@ Proof.
induction vars; simpl; intros.
auto.
apply IHvars. intro id1. unfold assign_global_variable.
- destruct a as [id2 lv2]. destruct lv2; simpl; rewrite PMap.gsspec; rewrite PTree.gsspec.
+ destruct a as [[id2 lv2] init2]. destruct lv2; simpl; rewrite PMap.gsspec; rewrite PTree.gsspec.
case (peq id1 id2); intro. auto. apply H.
case (peq id1 id2); intro. auto. apply H.
@@ -283,6 +297,7 @@ Qed.
must be normalized with respect to the memory chunk of the variable,
in the following sense. *)
+(*
Definition val_normalized (chunk: memory_chunk) (v: val) : Prop :=
exists v0, v = Val.load_result chunk v0.
@@ -305,34 +320,31 @@ Lemma load_result_normalized:
Proof.
intros chunk v [v0 EQ]. rewrite EQ. apply load_result_idem.
Qed.
-
+*)
Lemma match_env_store_local:
forall f cenv e m1 m2 te sp lo hi id b chunk v tv,
e!id = Some(b, Vscalar chunk) ->
- val_inject f v tv ->
- val_normalized chunk tv ->
+ val_inject f (Val.load_result chunk v) tv ->
store chunk m1 b 0 v = Some m2 ->
match_env f cenv e m1 te sp lo hi ->
match_env f cenv e m2 (PTree.set id tv te) sp lo hi.
Proof.
- intros. inversion H3. constructor; auto.
+ intros. inversion H2. constructor; auto.
intros. generalize (me_vars0 id0); intro.
- inversion H4; subst.
+ inversion H3; subst.
(* var_local *)
case (peq id id0); intro.
(* the stored variable *)
subst id0.
- change Csharpminor.var_kind with var_kind in H5.
- rewrite H in H6. injection H6; clear H6; intros; subst b0 chunk0.
+ change Csharpminor.var_kind with var_kind in H4.
+ rewrite H in H5. injection H5; clear H5; intros; subst b0 chunk0.
econstructor. eauto.
eapply load_store_same; eauto. auto.
rewrite PTree.gss. reflexivity.
- replace tv with (Val.load_result chunk tv).
- apply Mem.load_result_inject. constructor; auto.
- apply load_result_normalized; auto.
+ auto.
(* a different variable *)
econstructor; eauto.
- rewrite <- H7. eapply load_store_other; eauto.
+ rewrite <- H6. eapply load_store_other; eauto.
rewrite PTree.gso; auto.
(* var_stack_scalar *)
econstructor; eauto.
@@ -375,16 +387,15 @@ Qed.
Lemma match_callstack_store_local:
forall f cenv e te sp lo hi cs bound tbound m1 m2 id b chunk v tv,
e!id = Some(b, Vscalar chunk) ->
- val_inject f v tv ->
- val_normalized chunk tv ->
+ val_inject f (Val.load_result chunk v) tv ->
store chunk m1 b 0 v = Some m2 ->
match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m1 ->
match_callstack f (mkframe cenv e (PTree.set id tv te) sp lo hi :: cs) bound tbound m2.
Proof.
- intros. inversion H3. constructor; auto.
+ intros. inversion H2. constructor; auto.
eapply match_env_store_local; eauto.
eapply match_callstack_store_above; eauto.
- inversion H17.
+ inversion H16.
generalize (me_bounded0 _ _ _ H). omega.
Qed.
@@ -409,20 +420,19 @@ Qed.
Lemma match_callstack_store_local_unchanged:
forall f cenv e te sp lo hi cs bound tbound m1 m2 id b chunk v tv,
e!id = Some(b, Vscalar chunk) ->
- val_inject f v tv ->
- val_normalized chunk tv ->
+ val_inject f (Val.load_result chunk v) tv ->
store chunk m1 b 0 v = Some m2 ->
te!id = Some tv ->
match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m1 ->
match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m2.
Proof.
- intros. inversion H4. constructor; auto.
+ intros. inversion H3. constructor; auto.
apply match_env_extensional with (PTree.set id tv te).
eapply match_env_store_local; eauto.
intros. rewrite PTree.gsspec.
case (peq id0 id); intros. congruence. auto.
eapply match_callstack_store_above; eauto.
- inversion H18.
+ inversion H17.
generalize (me_bounded0 _ _ _ H). omega.
Qed.
@@ -698,6 +708,90 @@ Proof.
injection H; intros. rewrite <- H2; simpl. omega.
Qed.
+Lemma match_env_alloc:
+ forall m1 l h m2 b tm1 tm2 tb f1 ce e te sp lo hi,
+ alloc m1 l h = (m2, b) ->
+ alloc tm1 l h = (tm2, tb) ->
+ match_env f1 ce e m1 te sp lo hi ->
+ hi <= m1.(nextblock) ->
+ sp < tm1.(nextblock) ->
+ let f2 := extend_inject b (Some(tb, 0)) f1 in
+ inject_incr f1 f2 ->
+ match_env f2 ce e m2 te sp lo hi.
+Proof.
+ intros.
+ assert (BEQ: b = m1.(nextblock)). injection H; auto.
+ assert (TBEQ: tb = tm1.(nextblock)). injection H0; auto.
+ inversion H1. constructor; auto.
+ (* me_vars *)
+ intros. generalize (me_vars0 id); intro. inversion H5.
+ (* var_local *)
+ econstructor; eauto.
+ generalize (me_bounded0 _ _ _ H7). intro.
+ unfold f2, extend_inject. case (eq_block b0 b); intro.
+ subst b0. rewrite BEQ in H12. omegaContradiction.
+ auto.
+ (* var_stack_scalar *)
+ econstructor; eauto.
+ (* var_stack_array *)
+ econstructor; eauto.
+ (* var_global_scalar *)
+ econstructor; eauto.
+ (* var_global_array *)
+ econstructor; eauto.
+ (* me_bounded *)
+ intros until delta. unfold f2, extend_inject. case (eq_block b0 b); intro.
+ intro. injection H5; clear H5; intros.
+ rewrite H6 in TBEQ. rewrite TBEQ in H3. omegaContradiction.
+ eauto.
+ (* me_inj *)
+ intros until delta. unfold f2, extend_inject. case (eq_block b0 b); intros.
+ injection H5; clear H5; intros; subst b0 tb0 delta.
+ rewrite BEQ in H6. omegaContradiction.
+ eauto.
+Qed.
+
+Lemma match_callstack_alloc_rec:
+ forall f1 cs bound tbound m1,
+ match_callstack f1 cs bound tbound m1 ->
+ forall l h m2 b tm1 tm2 tb,
+ alloc m1 l h = (m2, b) ->
+ alloc tm1 l h = (tm2, tb) ->
+ bound <= m1.(nextblock) ->
+ tbound <= tm1.(nextblock) ->
+ let f2 := extend_inject b (Some(tb, 0)) f1 in
+ inject_incr f1 f2 ->
+ match_callstack f2 cs bound tbound m2.
+Proof.
+ induction 1; intros.
+ constructor.
+ inversion H. constructor.
+ intros. elim (mg_symbols0 _ _ H5); intros.
+ split; auto. elim (H4 b0); intros; congruence.
+ intros. generalize (mg_functions0 _ H5). elim (H4 b0); congruence.
+ constructor. auto. auto.
+ unfold f2. eapply match_env_alloc; eauto. omega. omega.
+ unfold f2; eapply IHmatch_callstack; eauto.
+ inversion H1; omega.
+ omega.
+Qed.
+
+Lemma match_callstack_alloc:
+ forall f1 cs m1 tm1 l h m2 b tm2 tb,
+ match_callstack f1 cs m1.(nextblock) tm1.(nextblock) m1 ->
+ alloc m1 l h = (m2, b) ->
+ alloc tm1 l h = (tm2, tb) ->
+ let f2 := extend_inject b (Some(tb, 0)) f1 in
+ inject_incr f1 f2 ->
+ match_callstack f2 cs m2.(nextblock) tm2.(nextblock) m2.
+Proof.
+ intros. unfold f2 in *.
+ apply match_callstack_incr_bound with m1.(nextblock) tm1.(nextblock).
+ eapply match_callstack_alloc_rec; eauto. omega. omega.
+ injection H0; intros; subst m2; simpl; omega.
+ injection H1; intros; subst tm2; simpl; omega.
+Qed.
+
(** [match_callstack] implies [match_globalenvs]. *)
Lemma match_callstack_match_globalenvs:
@@ -757,14 +851,14 @@ Qed.
provided arguments match pairwise ([val_list_inject f] hypothesis). *)
Lemma make_op_correct:
- forall al a op vl m2 v sp le te1 tm1 te2 tm2 tvl f,
+ forall al a op vl m2 v sp le te1 tm1 t te2 tm2 tvl f,
make_op op al = Some a ->
Csharpminor.eval_operation op vl m2 = Some v ->
- eval_exprlist tge (Vptr sp Int.zero) le te1 tm1 al te2 tm2 tvl ->
+ eval_exprlist tge (Vptr sp Int.zero) le te1 tm1 al t te2 tm2 tvl ->
val_list_inject f vl tvl ->
mem_inject f m2 tm2 ->
exists tv,
- eval_expr tge (Vptr sp Int.zero) le te1 tm1 a te2 tm2 tv
+ eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 tv
/\ val_inject f v tv.
Proof.
intros.
@@ -782,23 +876,27 @@ Proof.
(* floatconst *)
TrivialOp. econstructor. constructor. reflexivity.
(* Unary operators *)
- inversion H1. subst sp0 le0 e m a0 bl e2 m0 tvl.
- inversion H14. subst sp0 le0 e m e1 m1 vl0.
- inversion H2. subst vl v' vl'. inversion H8. subst vl0.
+ inversion H1; clear H1; subst.
+ inversion H11; clear H11; subst.
+ rewrite E0_right.
+ inversion H2; clear H2; subst. inversion H8; clear H8; subst.
destruct op; simplify_eq H; intro; subst a;
simpl in H0; destruct v1; simplify_eq H0; intro; subst v;
inversion H7; subst v0;
TrivialOp.
+ change (Vint (Int.cast8unsigned i)) with (Val.cast8unsigned (Vint i)). eauto with evalexpr.
+ change (Vint (Int.cast8signed i)) with (Val.cast8signed (Vint i)). eauto with evalexpr.
+ change (Vint (Int.cast16unsigned i)) with (Val.cast16unsigned (Vint i)). eauto with evalexpr.
+ change (Vint (Int.cast16signed i)) with (Val.cast16signed (Vint i)). eauto with evalexpr.
+ change (Vfloat (Float.singleoffloat f0)) with (Val.singleoffloat (Vfloat f0)). eauto with evalexpr.
(* Binary operations *)
- inversion H1. subst sp0 le0 e m a0 bl e2 m0 tvl.
- inversion H14. subst sp0 le0 e m a0 bl e2 m3 vl0.
- inversion H16. subst sp0 le0 e m e0 m0 vl1.
- inversion H2. subst vl v' vl'.
- inversion H8. subst vl0 v' vl'.
- inversion H12. subst vl.
+ inversion H1; clear H1; subst. inversion H11; clear H11; subst.
+ inversion H12; clear H12; subst. rewrite E0_right.
+ inversion H2; clear H2; subst. inversion H9; clear H9; subst.
+ inversion H10; clear H10; subst.
destruct op; simplify_eq H; intro; subst a;
simpl in H0; destruct v2; destruct v3; simplify_eq H0; intro; try subst v;
- inversion H7; inversion H9; subst v0; subst v1;
+ inversion H7; inversion H8; subst v0; subst v1;
TrivialOp.
(* add int ptr *)
exists (Vptr b2 (Int.add ofs2 i)); split.
@@ -815,7 +913,8 @@ Proof.
(* sub ptr ptr *)
destruct (eq_block b b0); simplify_eq H0; intro; subst v; subst b0.
assert (b4 = b2). congruence. subst b4.
- exists (Vint (Int.sub ofs2 ofs3)); split. eauto with evalexpr.
+ exists (Vint (Int.sub ofs3 ofs2)); split.
+ eauto with evalexpr.
subst ofs2 ofs3. replace x0 with x. rewrite Int.sub_shifted. constructor.
congruence.
(* divs *)
@@ -842,11 +941,11 @@ Proof.
(* cmp int ptr *)
elim (eval_compare_null_inv _ _ _ H0); intros; subst i1 i.
exists v; split. eauto with evalexpr.
- elim H18; intros [A B]; subst v; unfold Vtrue, Vfalse; constructor.
+ elim H12; intros [A B]; subst v; unfold Vtrue, Vfalse; constructor.
(* cmp ptr int *)
elim (eval_compare_null_inv _ _ _ H0); intros; subst i1 i0.
exists v; split. eauto with evalexpr.
- elim H18; intros [A B]; subst v; unfold Vtrue, Vfalse; constructor.
+ elim H12; intros [A B]; subst v; unfold Vtrue, Vfalse; constructor.
(* cmp ptr ptr *)
caseEq (valid_pointer m2 b (Int.signed i) && valid_pointer m2 b0 (Int.signed i0));
intro EQ; rewrite EQ in H0; try discriminate.
@@ -854,7 +953,7 @@ Proof.
assert (b4 = b2); [congruence|subst b4].
assert (x0 = x); [congruence|subst x0].
elim (andb_prop _ _ EQ); intros.
- exists (Val.of_bool (Int.cmp c ofs2 ofs3)); split.
+ exists (Val.of_bool (Int.cmp c ofs3 ofs2)); split.
eauto with evalexpr.
subst ofs2 ofs3. rewrite Int.translate_cmp.
apply val_inject_val_of_bool.
@@ -866,55 +965,51 @@ Qed.
normalized according to the given memory chunk. *)
Lemma make_cast_correct:
- forall f sp le te1 tm1 a te2 tm2 v chunk v' tv,
- eval_expr tge (Vptr sp Int.zero) le te1 tm1 a te2 tm2 tv ->
- cast chunk v = Some v' ->
+ forall f sp le te1 tm1 a t te2 tm2 v chunk tv,
+ eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 tv ->
val_inject f v tv ->
exists tv',
eval_expr tge (Vptr sp Int.zero) le
te1 tm1 (make_cast chunk a)
- te2 tm2 tv'
- /\ val_inject f v' tv'
- /\ val_normalized chunk tv'.
+ t te2 tm2 tv'
+ /\ val_inject f (Val.load_result chunk v) tv'.
Proof.
- intros. destruct chunk; destruct v; simplify_eq H0; intro; subst v'; simpl;
- inversion H1; subst tv.
+ intros. destruct chunk.
- exists (Vint (Int.cast8signed i)).
+ exists (Val.cast8signed tv).
split. apply eval_cast8signed; auto.
- split. constructor. exists (Vint i); reflexivity.
+ inversion H0; simpl; constructor.
- exists (Vint (Int.cast8unsigned i)).
+ exists (Val.cast8unsigned tv).
split. apply eval_cast8unsigned; auto.
- split. constructor. exists (Vint i); reflexivity.
+ inversion H0; simpl; constructor.
- exists (Vint (Int.cast16signed i)).
+ exists (Val.cast16signed tv).
split. apply eval_cast16signed; auto.
- split. constructor. exists (Vint i); reflexivity.
+ inversion H0; simpl; constructor.
- exists (Vint (Int.cast16unsigned i)).
+ exists (Val.cast16unsigned tv).
split. apply eval_cast16unsigned; auto.
- split. constructor. exists (Vint i); reflexivity.
-
- exists (Vint i).
- split. auto. split. auto. exists (Vint i); reflexivity.
+ inversion H0; simpl; constructor.
- exists (Vptr b2 ofs2).
- split. auto. split. auto. exists (Vptr b2 ofs2); reflexivity.
+ exists tv.
+ split. simpl; auto.
+ inversion H0; simpl; econstructor; eauto.
- exists (Vfloat (Float.singleoffloat f0)).
+ exists (Val.singleoffloat tv).
split. apply eval_singleoffloat; auto.
- split. constructor. exists (Vfloat f0); reflexivity.
+ inversion H0; simpl; constructor.
- exists (Vfloat f0).
- split. auto. split. auto. exists (Vfloat f0); reflexivity.
+ exists tv.
+ split. simpl; auto.
+ inversion H0; simpl; constructor.
Qed.
Lemma make_stackaddr_correct:
forall sp le te tm ofs,
eval_expr tge (Vptr sp Int.zero) le
te tm (make_stackaddr ofs)
- te tm (Vptr sp (Int.repr ofs)).
+ E0 te tm (Vptr sp (Int.repr ofs)).
Proof.
intros; unfold make_stackaddr.
eapply eval_Eop. econstructor. simpl. decEq. decEq.
@@ -926,7 +1021,7 @@ Lemma make_globaladdr_correct:
Genv.find_symbol tge id = Some b ->
eval_expr tge (Vptr sp Int.zero) le
te tm (make_globaladdr id)
- te tm (Vptr b Int.zero).
+ E0 te tm (Vptr b Int.zero).
Proof.
intros; unfold make_globaladdr.
eapply eval_Eop. econstructor. simpl. rewrite H. auto.
@@ -935,67 +1030,75 @@ Qed.
(** Correctness of [make_load] and [make_store]. *)
Lemma make_load_correct:
- forall sp le te1 tm1 a te2 tm2 va chunk v,
- eval_expr tge (Vptr sp Int.zero) le te1 tm1 a te2 tm2 va ->
+ forall sp le te1 tm1 a t te2 tm2 va chunk v,
+ eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 va ->
Mem.loadv chunk tm2 va = Some v ->
eval_expr tge (Vptr sp Int.zero) le
te1 tm1 (make_load chunk a)
- te2 tm2 v.
+ t te2 tm2 v.
Proof.
intros; unfold make_load.
eapply eval_load; eauto.
Qed.
-Lemma val_content_inject_cast:
- forall f chunk v1 v2 tv1,
- cast chunk v1 = Some v2 ->
- val_inject f v1 tv1 ->
- val_content_inject f (mem_chunk chunk) v2 tv1.
+Lemma store_arg_content_inject:
+ forall f sp le te1 tm1 a t te2 tm2 v va chunk,
+ eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 va ->
+ val_inject f v va ->
+ exists vb,
+ eval_expr tge (Vptr sp Int.zero) le te1 tm1 (store_arg chunk a) t te2 tm2 vb
+ /\ val_content_inject f (mem_chunk chunk) v vb.
Proof.
- intros. destruct chunk; destruct v1; simplify_eq H; intro; subst v2;
- inversion H0; simpl.
- apply val_content_inject_8. apply Int.cast8_unsigned_signed.
- apply val_content_inject_8. apply Int.cast8_unsigned_idem.
- apply val_content_inject_16. apply Int.cast16_unsigned_signed.
- apply val_content_inject_16. apply Int.cast16_unsigned_idem.
- constructor; constructor.
- constructor; econstructor; eauto.
- apply val_content_inject_32. apply Float.singleoffloat_idem.
- constructor; constructor.
+ intros.
+ assert (exists vb,
+ eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 vb
+ /\ val_content_inject f (mem_chunk chunk) v vb).
+ exists va; split. assumption. constructor. assumption.
+ inversion H; clear H; subst; simpl; trivial.
+ inversion H2; clear H2; subst; trivial.
+ inversion H4; clear H4; subst; trivial.
+ rewrite E0_right. rewrite E0_right in H1.
+ destruct op; trivial; destruct chunk; trivial;
+ exists v0; (split; [auto|
+ simpl in H3; inversion H3; clear H3; subst va;
+ destruct v0; simpl in H0; inversion H0; subst; try (constructor; constructor)]).
+ apply val_content_inject_8. apply Int.cast8_unsigned_signed.
+ apply val_content_inject_8. apply Int.cast8_unsigned_idem.
+ apply val_content_inject_16. apply Int.cast16_unsigned_signed.
+ apply val_content_inject_16. apply Int.cast16_unsigned_idem.
+ apply val_content_inject_32. apply Float.singleoffloat_idem.
Qed.
Lemma make_store_correct:
- forall f sp le te1 tm1 addr te2 tm2 tvaddr rhs te3 tm3 tvrhs
- chunk vrhs v m3 vaddr m4,
- eval_expr tge (Vptr sp Int.zero) le
- te1 tm1 addr te2 tm2 tvaddr ->
- eval_expr tge (Vptr sp Int.zero) le
- te2 tm2 rhs te3 tm3 tvrhs ->
- cast chunk vrhs = Some v ->
- Mem.storev chunk m3 vaddr v = Some m4 ->
+ forall f sp te1 tm1 addr te2 tm2 tvaddr rhs te3 tm3 tvrhs
+ chunk vrhs m3 vaddr m4 t1 t2,
+ eval_expr tge (Vptr sp Int.zero) nil
+ te1 tm1 addr t1 te2 tm2 tvaddr ->
+ eval_expr tge (Vptr sp Int.zero) nil
+ te2 tm2 rhs t2 te3 tm3 tvrhs ->
+ Mem.storev chunk m3 vaddr vrhs = Some m4 ->
mem_inject f m3 tm3 ->
val_inject f vaddr tvaddr ->
val_inject f vrhs tvrhs ->
- exists tm4, exists tv,
- eval_expr tge (Vptr sp Int.zero) le
+ exists tm4,
+ exec_stmt tge (Vptr sp Int.zero)
te1 tm1 (make_store chunk addr rhs)
- te3 tm4 tv
+ (t1**t2) te3 tm4 Out_normal
/\ mem_inject f m4 tm4
- /\ val_inject f v tv
/\ nextblock tm4 = nextblock tm3.
Proof.
intros. unfold make_store.
- assert (val_content_inject f (mem_chunk chunk) v tvrhs).
- eapply val_content_inject_cast; eauto.
- elim (storev_mapped_inject_1 _ _ _ _ _ _ _ _ _ H3 H2 H4 H6).
- intros tm4 [STORE MEMINJ].
- generalize (eval_store _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ H H0 STORE).
+ exploit store_arg_content_inject. eexact H0. eauto.
+ intros [tv [EVAL VCINJ]].
+ exploit storev_mapped_inject_1; eauto.
+ intros [tm4 [STORE MEMINJ]].
+ exploit eval_store. eexact H. eexact EVAL. eauto.
intro EVALSTORE.
- elim (make_cast_correct _ _ _ _ _ _ _ _ _ _ _ _ EVALSTORE H1 H5).
- intros tv [EVALCAST [VALINJ VALNORM]].
- exists tm4; exists tv. intuition.
+ exists tm4.
+ split. apply exec_Sexpr with tv. auto.
+ split. auto.
unfold storev in STORE; destruct tvaddr; try discriminate.
- generalize (store_inv _ _ _ _ _ _ STORE). simpl. tauto.
+ exploit store_inv; eauto. simpl. tauto.
Qed.
(** Correctness of the variable accessors [var_get], [var_set]
@@ -1009,7 +1112,7 @@ Lemma var_get_correct:
eval_var_ref prog e id b chunk ->
load chunk m b 0 = Some v ->
exists tv,
- eval_expr tge (Vptr sp Int.zero) le te tm a te tm tv /\
+ eval_expr tge (Vptr sp Int.zero) le te tm a E0 te tm tv /\
val_inject f v tv.
Proof.
unfold var_get; intros.
@@ -1025,8 +1128,8 @@ Proof.
inversion H2; [subst|congruence].
assert (b0 = b). congruence. subst b0.
assert (chunk0 = chunk). congruence. subst chunk0.
- assert (loadv chunk m (Vptr b Int.zero) = Some v). assumption.
- generalize (loadv_inject _ _ _ _ _ _ _ H1 H9 H7).
+ exploit loadv_inject; eauto.
+ unfold loadv. eexact H3.
intros [tv [LOAD INJ]].
exists tv; split.
eapply make_load_correct; eauto. eapply make_stackaddr_correct; eauto.
@@ -1052,7 +1155,7 @@ Lemma var_addr_correct:
var_addr cenv id = Some a ->
eval_var_addr prog e id b ->
exists tv,
- eval_expr tge (Vptr sp Int.zero) le te tm a te tm tv /\
+ eval_expr tge (Vptr sp Int.zero) le te tm a E0 te tm tv /\
val_inject f (Vptr b Int.zero) tv.
Proof.
unfold var_addr; intros.
@@ -1086,65 +1189,61 @@ Proof.
Qed.
Lemma var_set_correct:
- forall cenv id rhs a f e te2 sp lo hi m2 cs tm2 le te1 tm1 vrhs b chunk v1 v2 m3,
+ forall cenv id rhs a f e te2 sp lo hi m2 cs tm2 te1 tm1 tv b chunk v m3 t,
var_set cenv id rhs = Some a ->
match_callstack f (mkframe cenv e te2 sp lo hi :: cs) m2.(nextblock) tm2.(nextblock) m2 ->
- eval_expr tge (Vptr sp Int.zero) le te1 tm1 rhs te2 tm2 vrhs ->
- val_inject f v1 vrhs ->
+ eval_expr tge (Vptr sp Int.zero) nil te1 tm1 rhs t te2 tm2 tv ->
+ val_inject f v tv ->
mem_inject f m2 tm2 ->
eval_var_ref prog e id b chunk ->
- cast chunk v1 = Some v2 ->
- store chunk m2 b 0 v2 = Some m3 ->
- exists te3, exists tm3, exists tv,
- eval_expr tge (Vptr sp Int.zero) le te1 tm1 a te3 tm3 tv /\
- val_inject f v2 tv /\
+ store chunk m2 b 0 v = Some m3 ->
+ exists te3, exists tm3,
+ exec_stmt tge (Vptr sp Int.zero) te1 tm1 a t te3 tm3 Out_normal /\
mem_inject f m3 tm3 /\
match_callstack f (mkframe cenv e te3 sp lo hi :: cs) m3.(nextblock) tm3.(nextblock) m3.
Proof.
unfold var_set; intros.
assert (NEXTBLOCK: nextblock m3 = nextblock m2).
- generalize (store_inv _ _ _ _ _ _ H6). simpl. tauto.
+ exploit store_inv; eauto. simpl; tauto.
inversion H0. subst f0 cenv0 e0 te sp0 lo0 hi0 cs0 bound tbound m.
- assert (match_var f id e m2 te2 sp cenv!!id). inversion H20; auto.
- inversion H7; subst; rewrite <- H8 in H; inversion H; subst; clear H.
+ assert (match_var f id e m2 te2 sp cenv!!id). inversion H19; auto.
+ inversion H6; subst; rewrite <- H7 in H; inversion H; subst; clear H.
(* var_local *)
inversion H4; [subst|congruence].
assert (b0 = b). congruence. subst b0.
assert (chunk0 = chunk). congruence. subst chunk0.
- elim (make_cast_correct _ _ _ _ _ _ _ _ _ _ _ _ H1 H5 H2).
- intros tv [EVAL [INJ NORM]].
- exists (PTree.set id tv te2); exists tm2; exists tv.
- split. eapply eval_Eassign. auto.
- split. auto.
+ exploit make_cast_correct; eauto.
+ intros [tv' [EVAL INJ]].
+ exists (PTree.set id tv' te2); exists tm2.
+ split. eapply exec_Sexpr. eapply eval_Eassign. eauto.
split. eapply store_unmapped_inject; eauto.
rewrite NEXTBLOCK. eapply match_callstack_store_local; eauto.
(* var_stack_scalar *)
inversion H4; [subst|congruence].
assert (b0 = b). congruence. subst b0.
assert (chunk0 = chunk). congruence. subst chunk0.
- assert (storev chunk m2 (Vptr b Int.zero) v2 = Some m3). assumption.
- generalize (make_stackaddr_correct sp le te1 tm1 ofs). intro EVALSTACKADDR.
- generalize (make_store_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- EVALSTACKADDR H1 H5 H11 H3 H10 H2).
- intros [tm3 [tv [EVAL [MEMINJ [VALINJ TNEXTBLOCK]]]]].
- exists te2; exists tm3; exists tv.
- split. auto. split. auto. split. auto.
+ assert (storev chunk m2 (Vptr b Int.zero) v = Some m3). assumption.
+ exploit make_store_correct.
+ eapply make_stackaddr_correct.
+ eauto. eauto. eauto. eauto. eauto.
+ rewrite E0_left. intros [tm3 [EVAL [MEMINJ TNEXTBLOCK]]].
+ exists te2; exists tm3.
+ split. auto. split. auto.
rewrite NEXTBLOCK; rewrite TNEXTBLOCK.
eapply match_callstack_mapped; eauto.
- inversion H10; congruence.
+ inversion H9; congruence.
(* var_global_scalar *)
inversion H4; [congruence|subst].
assert (chunk0 = chunk). congruence. subst chunk0.
- assert (storev chunk m2 (Vptr b Int.zero) v2 = Some m3). assumption.
+ assert (storev chunk m2 (Vptr b Int.zero) v = Some m3). assumption.
assert (match_globalenvs f). eapply match_callstack_match_globalenvs; eauto.
- inversion H14. destruct (mg_symbols0 _ _ H11) as [A B].
- assert (val_inject f (Vptr b Int.zero) (Vptr b Int.zero)). econstructor; eauto.
- generalize (make_globaladdr_correct sp le te1 tm1 id b B). intro EVALGLOBALADDR.
- generalize (make_store_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- EVALGLOBALADDR H1 H5 H13 H3 H15 H2).
- intros [tm3 [tv [EVAL [MEMINJ [VALINJ TNEXTBLOCK]]]]].
- exists te2; exists tm3; exists tv.
- split. auto. split. auto. split. auto.
+ inversion H13. destruct (mg_symbols0 _ _ H10) as [A B].
+ exploit make_store_correct.
+ eapply make_globaladdr_correct; eauto.
+ eauto. eauto. eauto. eauto. eauto.
+ rewrite E0_left. intros [tm3 [EVAL [MEMINJ TNEXTBLOCK]]].
+ exists te2; exists tm3.
+ split. auto. split. auto.
rewrite NEXTBLOCK; rewrite TNEXTBLOCK.
eapply match_callstack_mapped; eauto. congruence.
Qed.
@@ -1228,36 +1327,26 @@ Proof.
omega. omega. omega. omega. unfold sizeof; rewrite LV. omega.
intros. left. generalize (BOUND _ _ H5). omega.
elim H3; intros MINJ1 INCR1; clear H3.
- assert (MATCH1: match_callstack f1
- (mkframe cenv1 (PTree.set id (b1, lv) e) te sp lo (nextblock m1) :: cs)
- (nextblock m1) (nextblock tm) m1).
- unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto.
- assert (SZ1POS: 0 <= sz1). rewrite <- H1. omega.
- assert (BOUND1: forall b delta, f1 b = Some(sp, delta) ->
- high_bound m1 b + delta <= sz1).
+ exploit IHalloc_variables; eauto.
+ unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto.
+ rewrite <- H1. omega.
intros until delta; unfold f1, extend_inject, eq_block.
rewrite (high_bound_alloc _ _ b _ _ _ H).
case (zeq b b1); intros.
inversion H3. unfold sizeof; rewrite LV. omega.
generalize (BOUND _ _ H3). omega.
- generalize (IHalloc_variables _ _ _ ASVS MATCH1 MINJ1 SZ1POS BOUND1 DEFINED1).
intros [f' [INCR2 [MINJ2 MATCH2]]].
exists f'; intuition. eapply inject_incr_trans; eauto.
(* 1.2 info = Var_local chunk *)
intro EQ; injection EQ; intros; clear EQ. subst sz1.
- generalize (alloc_unmapped_inject _ _ _ _ _ _ _ MINJ H).
+ exploit alloc_unmapped_inject; eauto.
set (f1 := extend_inject b1 None f). intros [MINJ1 INCR1].
- assert (MATCH1: match_callstack f1
- (mkframe cenv1 (PTree.set id (b1, lv) e) te sp lo (nextblock m1) :: cs)
- (nextblock m1) (nextblock tm) m1).
- unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto.
- assert (BOUND1: forall b delta, f1 b = Some(sp, delta) ->
- high_bound m1 b + delta <= sz).
+ exploit IHalloc_variables; eauto.
+ unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto.
intros until delta; unfold f1, extend_inject, eq_block.
rewrite (high_bound_alloc _ _ b _ _ _ H).
case (zeq b b1); intros. discriminate.
eapply BOUND; eauto.
- generalize (IHalloc_variables _ _ _ ASVS MATCH1 MINJ1 SZPOS BOUND1 DEFINED1).
intros [f' [INCR2 [MINJ2 MATCH2]]].
exists f'; intuition. eapply inject_incr_trans; eauto.
(* 2. lv = LVarray dim, info = Var_stack_array *)
@@ -1273,20 +1362,15 @@ Proof.
unfold f1; eapply alloc_mapped_inject; eauto.
omega. omega. omega. omega. unfold sizeof; rewrite LV. omega.
intros. left. generalize (BOUND _ _ H8). omega.
- elim H6; intros MINJ1 INCR1; clear H6.
- assert (MATCH1: match_callstack f1
- (mkframe cenv1 (PTree.set id (b1, lv) e) te sp lo (nextblock m1) :: cs)
- (nextblock m1) (nextblock tm) m1).
+ destruct H6 as [MINJ1 INCR1].
+ exploit IHalloc_variables; eauto.
unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto.
- assert (SZ1POS: 0 <= sz1). rewrite <- H1. omega.
- assert (BOUND1: forall b delta, f1 b = Some(sp, delta) ->
- high_bound m1 b + delta <= sz1).
- intros until delta; unfold f1, extend_inject, eq_block.
- rewrite (high_bound_alloc _ _ b _ _ _ H).
- case (zeq b b1); intros.
- inversion H6. unfold sizeof; rewrite LV. omega.
- generalize (BOUND _ _ H6). omega.
- generalize (IHalloc_variables _ _ _ ASVS MATCH1 MINJ1 SZ1POS BOUND1 DEFINED1).
+ rewrite <- H1. omega.
+ intros until delta; unfold f1, extend_inject, eq_block.
+ rewrite (high_bound_alloc _ _ b _ _ _ H).
+ case (zeq b b1); intros.
+ inversion H6. unfold sizeof; rewrite LV. omega.
+ generalize (BOUND _ _ H6). omega.
intros [f' [INCR2 [MINJ2 MATCH2]]].
exists f'; intuition. eapply inject_incr_trans; eauto.
Qed.
@@ -1370,15 +1454,15 @@ Proof.
(* me_inj *)
intros until lv2. unfold Csharpminor.empty_env; rewrite PTree.gempty; congruence.
(* me_inv *)
- intros. elim (mi_mappedblocks _ _ _ H4 _ _ _ H5); intros.
- elim (fresh_block_alloc _ _ _ _ _ H2 H6).
+ intros. exploit mi_mappedblocks; eauto. intros [A B].
+ elim (fresh_block_alloc _ _ _ _ _ H2 A).
(* me_incr *)
- intros. elim (mi_mappedblocks _ _ _ H4 _ _ _ H5); intros.
+ intros. exploit mi_mappedblocks; eauto. intros [A B].
rewrite SP; auto.
rewrite SP; auto.
eapply alloc_right_inject; eauto.
omega.
- intros. elim (mi_mappedblocks _ _ _ H4 _ _ _ H5); intros.
+ intros. exploit mi_mappedblocks; eauto. intros [A B].
unfold block in SP; omegaContradiction.
(* defined *)
intros. unfold te. apply set_locals_params_defined.
@@ -1421,7 +1505,7 @@ Proof.
unfold block; rewrite H2; omega.
elim H4; intro. left; congruence. right; auto.
elim H3; intro. subst b b1.
- generalize (alloc_variables_nextblock_incr _ _ _ _ _ _ H0).
+ generalize (alloc_variables_nextblock_incr _ _ _ _ _ _ H0).
rewrite H2. omega.
generalize (B H4). rewrite H2. omega.
Qed.
@@ -1465,7 +1549,7 @@ Lemma store_parameters_correct:
exists te2, exists tm2,
exec_stmt tge (Vptr sp Int.zero)
te1 tm1 (store_parameters cenv params)
- te2 tm2 Out_normal
+ E0 te2 tm2 Out_normal
/\ mem_inject f m2 tm2
/\ match_callstack f (mkframe cenv e te2 sp lo hi :: cs) m2.(nextblock) tm2.(nextblock) m2.
Proof.
@@ -1476,59 +1560,54 @@ Proof.
intros until tm1. intros VVM NOREPET MINJ MATCH. simpl.
inversion VVM. subst f0 id0 chunk0 vars v vals te.
inversion MATCH. subst f0 cenv0 e0 te sp0 lo0 hi0 cs0 bound tbound m0.
- inversion H19.
+ inversion H18.
inversion NOREPET. subst hd tl.
assert (NEXT: nextblock m1 = nextblock m).
- generalize (store_inv _ _ _ _ _ _ H1). simpl; tauto.
- generalize (me_vars0 id). intro. inversion H3; subst.
+ exploit store_inv; eauto. simpl; tauto.
+ generalize (me_vars0 id). intro. inversion H2; subst.
(* cenv!!id = Var_local chunk *)
assert (b0 = b). congruence. subst b0.
assert (chunk0 = chunk). congruence. subst chunk0.
assert (v' = tv). congruence. subst v'.
- assert (eval_expr tge (Vptr sp Int.zero) nil te1 tm1 (Evar id) te1 tm1 tv).
- constructor. auto.
- generalize (make_cast_correct _ _ _ _ _ _ _ _ _ _ _ _
- H15 H0 H11).
- intros [tv' [EVAL1 [VINJ1 VNORM]]].
+ exploit make_cast_correct.
+ apply eval_Evar with (id := id). eauto.
+ eexact H10.
+ intros [tv' [EVAL1 VINJ1]].
set (te2 := PTree.set id tv' te1).
assert (VVM2: vars_vals_match f params vl te2).
apply vars_vals_match_extensional with te1; auto.
intros. unfold te2; apply PTree.gso. red; intro; subst id0.
- elim H5. change id with (fst (id, lv)). apply List.in_map; auto.
- generalize (store_unmapped_inject _ _ _ _ _ _ _ _ MINJ H1 H9); intro MINJ2.
- generalize (match_callstack_store_local _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- H VINJ1 VNORM H1 MATCH);
+ elim H4. change id with (fst (id, lv)). apply List.in_map; auto.
+ exploit store_unmapped_inject; eauto. intro MINJ2.
+ exploit match_callstack_store_local; eauto.
fold te2; rewrite <- NEXT; intro MATCH2.
- destruct (IHbind_parameters _ _ _ _ _ _ _ _ VVM2 H6 MINJ2 MATCH2)
- as [te3 [tm3 [EXEC3 [MINJ3 MATCH3]]]].
+ exploit IHbind_parameters; eauto.
+ intros [te3 [tm3 [EXEC3 [MINJ3 MATCH3]]]].
exists te3; exists tm3.
(* execution *)
- split. apply exec_Sseq_continue with te2 tm1.
- econstructor. unfold te2. constructor. assumption.
- assumption.
+ split. apply exec_Sseq_continue with E0 te2 tm1 E0.
+ econstructor. unfold te2. constructor. eassumption.
+ assumption. traceEq.
(* meminj & match_callstack *)
tauto.
(* cenv!!id = Var_stack_scalar *)
assert (b0 = b). congruence. subst b0.
assert (chunk0 = chunk). congruence. subst chunk0.
- pose (EVAL1 := make_stackaddr_correct sp nil te1 tm1 ofs).
- assert (EVAL2: eval_expr tge (Vptr sp Int.zero) nil te1 tm1 (Evar id) te1 tm1 tv).
- constructor. auto.
- destruct (make_store_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- (Vptr b Int.zero) _
- EVAL1 EVAL2 H0 H1 MINJ H8 H11)
- as [tm2 [tv' [EVAL3 [MINJ2 [VINJ NEXT1]]]]].
- assert (f b <> None). inversion H8. congruence.
- generalize (match_callstack_mapped _ _ _ _ _ MATCH _ _ _ _ _ H9 H1).
+ exploit make_store_correct.
+ eapply make_stackaddr_correct.
+ apply eval_Evar with (id := id).
+ eauto. 2:eauto. 2:eauto. unfold storev; eexact H0. eauto.
+ intros [tm2 [EVAL3 [MINJ2 NEXT1]]].
+ exploit match_callstack_mapped.
+ eexact MATCH. 2:eauto. inversion H7. congruence.
rewrite <- NEXT; rewrite <- NEXT1; intro MATCH2.
- destruct (IHbind_parameters _ _ _ _ _ _ _ _
- H12 H6 MINJ2 MATCH2) as [te3 [tm3 [EVAL4 [MINJ3 MATCH3]]]].
+ exploit IHbind_parameters; eauto.
+ intros [te3 [tm3 [EVAL4 [MINJ3 MATCH3]]]].
exists te3; exists tm3.
(* execution *)
- split. apply exec_Sseq_continue with te1 tm2.
- econstructor. eauto.
- assumption.
+ split. apply exec_Sseq_continue with (E0**E0) te1 tm2 E0.
+ auto. assumption. traceEq.
(* meminj & match_callstack *)
tauto.
@@ -1589,44 +1668,6 @@ Proof.
induction 1; simpl; eauto.
Qed.
-(****
-Lemma build_compilenv_domain:
- forall f id chunk,
- In (id, chunk) f.(Csharpminor.fn_params) ->
- (fst (build_compilenv gce f))!id <> None.
-Proof.
- assert (forall atk id lv cenv_sz id0,
- let cenv_sz' := assign_variable atk (id, lv) cenv_sz in
- (fst cenv_sz')!id <> None
- /\ ((fst cenv_sz)!id0 <> None -> (fst cenv_sz')!id0 <> None)).
- intros. unfold cenv_sz'. destruct cenv_sz as [cenv sz].
- unfold assign_variable. destruct lv.
- case (Identset.mem id atk); simpl. split. rewrite PTree.gss. congruence.
- rewrite PTree.gsspec. case (peq id0 id); intros. congruence. auto.
- split. rewrite PTree.gss. congruence.
- rewrite PTree.gsspec. case (peq id0 id); intros. congruence. auto.
- simpl. split. rewrite PTree.gss. congruence.
- rewrite PTree.gsspec. case (peq id0 id); intros. congruence. auto.
-
- assert (forall atk id_lv_list cenv_sz id lv,
- In (id, lv) id_lv_list \/ (fst cenv_sz)!id <> None ->
- (fst (assign_variables atk id_lv_list cenv_sz))!id <> None).
- induction id_lv_list; simpl; intros.
- tauto.
- apply IHid_lv_list with lv.
- destruct a as [id0 lv0].
- generalize (H atk id0 lv0 cenv_sz id).
- simpl. intro. intuition. injection H0; intros; subst id0 lv0. intuition.
-
- intros. unfold build_compilenv. apply H0 with (Vscalar chunk).
- left. unfold fn_variables. apply List.in_or_app. left.
- set (g := fun (id_chunk : ident * memory_chunk) => (fst id_chunk, Vscalar (snd id_chunk))).
- change positive with ident.
- change (id, Vscalar chunk) with (g (id, chunk)).
- apply List.in_map. auto.
-Qed.
-****)
-
(** The final result in this section: the behaviour of function entry
in the generated Cminor code (allocate stack data block and store
parameters whose address is taken) simulates what happens at function
@@ -1649,7 +1690,7 @@ Lemma function_entry_ok:
exists f2, exists te2, exists tm2,
exec_stmt tge (Vptr sp Int.zero)
te tm1 (store_parameters cenv fn.(Csharpminor.fn_params))
- te2 tm2 Out_normal
+ E0 te2 tm2 Out_normal
/\ mem_inject f2 m2 tm2
/\ inject_incr f f2
/\ match_callstack f2
@@ -1658,24 +1699,21 @@ Lemma function_entry_ok:
/\ (forall b, m.(nextblock) <= b < m1.(nextblock) <-> In b lb).
Proof.
intros.
- generalize (bind_parameters_length _ _ _ _ _ H0); intro LEN1.
- destruct (match_callstack_alloc_variables _ _ _ _ _ _ _ _ _ _ _ _ tvargs
- H2 H3 H H4 H1 H6)
- as [f1 [INCR1 [MINJ1 MATCH1]]].
- fold te in MATCH1.
- assert (VLI: val_list_inject f1 vargs tvargs).
- eapply val_list_inject_incr; eauto.
- generalize (vars_vals_match_holds _ _ _ _ LEN1 VLI _
- (list_norepet_append_commut _ _ H7)).
- fold te. intro VVM.
- assert (NOREPET: list_norepet (List.map (@fst ident memory_chunk) fn.(Csharpminor.fn_params))).
- unfold fn_params_names in H7.
- eapply list_norepet_append_left; eauto.
- destruct (store_parameters_correct _ _ _ _ _ H0 _ _ _ _ _ _ _ _
- VVM NOREPET MINJ1 MATCH1)
- as [te2 [tm2 [EXEC [MINJ2 MATCH2]]]].
+ exploit bind_parameters_length; eauto. intro LEN1.
+ exploit match_callstack_alloc_variables; eauto.
+ intros [f1 [INCR1 [MINJ1 MATCH1]]].
+ exploit vars_vals_match_holds.
+ eauto. apply val_list_inject_incr with f. eauto. eauto.
+ apply list_norepet_append_commut.
+ unfold fn_vars_names in H7. eexact H7.
+ intro VVM.
+ exploit store_parameters_correct.
+ eauto. eauto.
+ unfold fn_params_names in H7. eapply list_norepet_append_left; eauto.
+ eexact MINJ1. eauto.
+ intros [te2 [tm2 [EXEC [MINJ2 MATCH2]]]].
exists f1; exists te2; exists tm2.
- split. auto. split. auto. split. auto. split. auto.
+ split; auto. split; auto. split; auto. split; auto.
intros; eapply alloc_variables_list_block; eauto.
Qed.
@@ -1745,7 +1783,7 @@ Ltac monadInv H :=
hypotheses in the proof of simulation. *)
Definition eval_expr_prop
- (le: Csharpminor.letenv) (e: Csharpminor.env) (m1: mem) (a: Csharpminor.expr) (m2: mem) (v: val) : Prop :=
+ (le: Csharpminor.letenv) (e: Csharpminor.env) (m1: mem) (a: Csharpminor.expr) (t: trace) (m2: mem) (v: val) : Prop :=
forall cenv ta f1 tle te1 tm1 sp lo hi cs
(TR: transl_expr cenv a = Some ta)
(LINJ: val_list_inject f1 le tle)
@@ -1754,7 +1792,7 @@ Definition eval_expr_prop
(mkframe cenv e te1 sp lo hi :: cs)
m1.(nextblock) tm1.(nextblock) m1),
exists f2, exists te2, exists tm2, exists tv,
- eval_expr tge (Vptr sp Int.zero) tle te1 tm1 ta te2 tm2 tv
+ eval_expr tge (Vptr sp Int.zero) tle te1 tm1 ta t te2 tm2 tv
/\ val_inject f2 v tv
/\ mem_inject f2 m2 tm2
/\ inject_incr f1 f2
@@ -1763,7 +1801,7 @@ Definition eval_expr_prop
m2.(nextblock) tm2.(nextblock) m2.
Definition eval_exprlist_prop
- (le: Csharpminor.letenv) (e: Csharpminor.env) (m1: mem) (al: Csharpminor.exprlist) (m2: mem) (vl: list val) : Prop :=
+ (le: Csharpminor.letenv) (e: Csharpminor.env) (m1: mem) (al: Csharpminor.exprlist) (t: trace) (m2: mem) (vl: list val) : Prop :=
forall cenv tal f1 tle te1 tm1 sp lo hi cs
(TR: transl_exprlist cenv al = Some tal)
(LINJ: val_list_inject f1 le tle)
@@ -1772,7 +1810,7 @@ Definition eval_exprlist_prop
(mkframe cenv e te1 sp lo hi :: cs)
m1.(nextblock) tm1.(nextblock) m1),
exists f2, exists te2, exists tm2, exists tvl,
- eval_exprlist tge (Vptr sp Int.zero) tle te1 tm1 tal te2 tm2 tvl
+ eval_exprlist tge (Vptr sp Int.zero) tle te1 tm1 tal t te2 tm2 tvl
/\ val_list_inject f2 vl tvl
/\ mem_inject f2 m2 tm2
/\ inject_incr f1 f2
@@ -1781,14 +1819,14 @@ Definition eval_exprlist_prop
m2.(nextblock) tm2.(nextblock) m2.
Definition eval_funcall_prop
- (m1: mem) (fn: Csharpminor.function) (args: list val) (m2: mem) (res: val) : Prop :=
+ (m1: mem) (fn: Csharpminor.fundef) (args: list val) (t: trace) (m2: mem) (res: val) : Prop :=
forall tfn f1 tm1 cs targs
- (TR: transl_function gce fn = Some tfn)
+ (TR: transl_fundef gce fn = Some tfn)
(MINJ: mem_inject f1 m1 tm1)
(MATCH: match_callstack f1 cs m1.(nextblock) tm1.(nextblock) m1)
(ARGSINJ: val_list_inject f1 args targs),
exists f2, exists tm2, exists tres,
- eval_funcall tge tm1 tfn targs tm2 tres
+ eval_funcall tge tm1 tfn targs t tm2 tres
/\ val_inject f2 res tres
/\ mem_inject f2 m2 tm2
/\ inject_incr f1 f2
@@ -1807,7 +1845,7 @@ Inductive outcome_inject (f: meminj) : Csharpminor.outcome -> outcome -> Prop :=
outcome_inject f (Csharpminor.Out_return (Some v1)) (Out_return (Some v2)).
Definition exec_stmt_prop
- (e: Csharpminor.env) (m1: mem) (s: Csharpminor.stmt) (m2: mem) (out: Csharpminor.outcome): Prop :=
+ (e: Csharpminor.env) (m1: mem) (s: Csharpminor.stmt) (t: trace) (m2: mem) (out: Csharpminor.outcome): Prop :=
forall cenv ts f1 te1 tm1 sp lo hi cs
(TR: transl_stmt cenv s = Some ts)
(MINJ: mem_inject f1 m1 tm1)
@@ -1815,7 +1853,7 @@ Definition exec_stmt_prop
(mkframe cenv e te1 sp lo hi :: cs)
m1.(nextblock) tm1.(nextblock) m1),
exists f2, exists te2, exists tm2, exists tout,
- exec_stmt tge (Vptr sp Int.zero) te1 tm1 ts te2 tm2 tout
+ exec_stmt tge (Vptr sp Int.zero) te1 tm1 ts t te2 tm2 tout
/\ outcome_inject f2 out tout
/\ mem_inject f2 m2 tm2
/\ inject_incr f1 f2
@@ -1823,14 +1861,6 @@ Definition exec_stmt_prop
(mkframe cenv e te2 sp lo hi :: cs)
m2.(nextblock) tm2.(nextblock) m2.
-(*
-Check (eval_funcall_ind4 prog
- eval_expr_prop
- eval_exprlist_prop
- eval_funcall_prop
- exec_stmt_prop).
-*)
-
(** There are as many cases in the inductive proof as there are evaluation
rules in the Csharpminor semantics. We treat each case as a separate
lemma. *)
@@ -1841,33 +1871,12 @@ Lemma transl_expr_Evar_correct:
(b : block) (chunk : memory_chunk) (v : val),
eval_var_ref prog e id b chunk ->
load chunk m b 0 = Some v ->
- eval_expr_prop le e m (Csharpminor.Evar id) m v.
+ eval_expr_prop le e m (Csharpminor.Evar id) E0 m v.
Proof.
intros; red; intros. unfold transl_expr in TR.
- generalize (var_get_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ tle
- TR MATCH MINJ H H0).
+ exploit var_get_correct; eauto.
intros [tv [EVAL VINJ]].
- exists f1; exists te1; exists tm1; exists tv; intuition.
-Qed.
-
-Lemma transl_expr_Eassign_correct:
- forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (id : positive) (a : Csharpminor.expr) (m1 : mem) (b : block)
- (chunk : memory_chunk) (v1 v2 : val) (m2 : mem),
- Csharpminor.eval_expr prog le e m a m1 v1 ->
- eval_expr_prop le e m a m1 v1 ->
- eval_var_ref prog e id b chunk ->
- cast chunk v1 = Some v2 ->
- store chunk m1 b 0 v2 = Some m2 ->
- eval_expr_prop le e m (Csharpminor.Eassign id a) m2 v2.
-Proof.
- intros; red; intros. monadInv TR; intro EQ0.
- generalize (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH).
- intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR12 MATCH1]]]]]]]].
- generalize (var_set_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- EQ0 MATCH1 EVAL1 VINJ1 MINJ1 H1 H2 H3).
- intros [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 MATCH2]]]]]].
- exists f2; exists te3; exists tm3; exists tv2. tauto.
+ exists f1; exists te1; exists tm1; exists tv; intuition eauto.
Qed.
Lemma transl_expr_Eaddrof_correct:
@@ -1875,199 +1884,162 @@ Lemma transl_expr_Eaddrof_correct:
(e : Csharpminor.env) (m : mem) (id : positive)
(b : block),
eval_var_addr prog e id b ->
- eval_expr_prop le e m (Eaddrof id) m (Vptr b Int.zero).
+ eval_expr_prop le e m (Eaddrof id) E0 m (Vptr b Int.zero).
Proof.
intros; red; intros. simpl in TR.
- generalize (var_addr_correct _ _ _ _ _ _ _ _ _ _ _ _ _ tle
- MATCH TR H).
+ exploit var_addr_correct; eauto.
intros [tv [EVAL VINJ]].
- exists f1; exists te1; exists tm1; exists tv. intuition.
+ exists f1; exists te1; exists tm1; exists tv. intuition eauto.
Qed.
Lemma transl_expr_Eop_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (op : Csharpminor.operation) (al : Csharpminor.exprlist) (m1 : mem)
- (vl : list val) (v : val),
- Csharpminor.eval_exprlist prog le e m al m1 vl ->
- eval_exprlist_prop le e m al m1 vl ->
+ (op : Csharpminor.operation) (al : Csharpminor.exprlist)
+ (t: trace) (m1 : mem) (vl : list val) (v : val),
+ Csharpminor.eval_exprlist prog le e m al t m1 vl ->
+ eval_exprlist_prop le e m al t m1 vl ->
Csharpminor.eval_operation op vl m1 = Some v ->
- eval_expr_prop le e m (Csharpminor.Eop op al) m1 v.
+ eval_expr_prop le e m (Csharpminor.Eop op al) t m1 v.
Proof.
intros; red; intros. monadInv TR; intro EQ0.
- generalize (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH).
+ exploit H0; eauto.
intros [f2 [te2 [tm2 [tvl [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
- generalize (make_op_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _
- EQ0 H1 EVAL1 VINJ1 MINJ1).
+ exploit make_op_correct; eauto.
intros [tv [EVAL2 VINJ2]].
exists f2; exists te2; exists tm2; exists tv. intuition.
Qed.
Lemma transl_expr_Eload_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (chunk : memory_chunk) (a : Csharpminor.expr) (m1 : mem)
+ (chunk : memory_chunk) (a : Csharpminor.expr) (t: trace) (m1 : mem)
(v1 v : val),
- Csharpminor.eval_expr prog le e m a m1 v1 ->
- eval_expr_prop le e m a m1 v1 ->
+ Csharpminor.eval_expr prog le e m a t m1 v1 ->
+ eval_expr_prop le e m a t m1 v1 ->
loadv chunk m1 v1 = Some v ->
- eval_expr_prop le e m (Csharpminor.Eload chunk a) m1 v.
+ eval_expr_prop le e m (Csharpminor.Eload chunk a) t m1 v.
Proof.
intros; red; intros.
monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]].
- destruct (loadv_inject _ _ _ _ _ _ _ MINJ2 H1 VINJ1)
- as [tv [TLOAD VINJ]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]].
+ exploit loadv_inject; eauto.
+ intros [tv [TLOAD VINJ]].
exists f2; exists te2; exists tm2; exists tv.
intuition.
subst ta. eapply make_load_correct; eauto.
Qed.
-Lemma transl_expr_Estore_correct:
- forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (chunk : memory_chunk) (a b : Csharpminor.expr) (m1 : mem)
- (v1 : val) (m2 : mem) (v2 : val) (m3 : mem) (v3 : val),
- Csharpminor.eval_expr prog le e m a m1 v1 ->
- eval_expr_prop le e m a m1 v1 ->
- Csharpminor.eval_expr prog le e m1 b m2 v2 ->
- eval_expr_prop le e m1 b m2 v2 ->
- cast chunk v2 = Some v3 ->
- storev chunk m2 v1 v3 = Some m3 ->
- eval_expr_prop le e m (Csharpminor.Estore chunk a b) m3 v3.
-Proof.
- intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
- assert (LINJ2: val_list_inject f2 le tle). eapply val_list_inject_incr; eauto.
- destruct (H2 _ _ _ _ _ _ _ _ _ _ EQ0 LINJ2 MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
- assert (VINJ1': val_inject f3 v1 tv1). eapply val_inject_incr; eauto.
- destruct (make_store_correct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- EVAL1 EVAL2 H3 H4 MINJ3 VINJ1' VINJ2)
- as [tm4 [tv [EVAL [MINJ4 [VINJ4 NEXTBLOCK]]]]].
- exists f3; exists te3; exists tm4; exists tv.
- rewrite <- H6. intuition.
- eapply inject_incr_trans; eauto.
- assert (val_inject f3 v1 tv1). eapply val_inject_incr; eauto.
- unfold storev in H4; destruct v1; try discriminate.
- inversion H5.
- rewrite NEXTBLOCK. replace (nextblock m3) with (nextblock m2).
- eapply match_callstack_mapped; eauto. congruence.
- generalize (store_inv _ _ _ _ _ _ H4). simpl; symmetry; tauto.
-Qed.
-
-Lemma sig_transl_function:
- forall f tf, transl_function gce f = Some tf -> tf.(fn_sig) = f.(Csharpminor.fn_sig).
-Proof.
- intros f tf. unfold transl_function.
- destruct (build_compilenv gce f).
- case (zle z Int.max_signed); intros.
- monadInv H. subst tf; reflexivity.
- congruence.
-Qed.
-
Lemma transl_expr_Ecall_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
(sig : signature) (a : Csharpminor.expr) (bl : Csharpminor.exprlist)
- (m1 m2 m3 : mem) (vf : val) (vargs : list val) (vres : val)
- (f : Csharpminor.function),
- Csharpminor.eval_expr prog le e m a m1 vf ->
- eval_expr_prop le e m a m1 vf ->
- Csharpminor.eval_exprlist prog le e m1 bl m2 vargs ->
- eval_exprlist_prop le e m1 bl m2 vargs ->
+ (t1: trace) (m1: mem) (t2: trace) (m2: mem) (t3: trace) (m3: mem)
+ (vf : val) (vargs : list val) (vres : val)
+ (f : Csharpminor.fundef) (t: trace),
+ Csharpminor.eval_expr prog le e m a t1 m1 vf ->
+ eval_expr_prop le e m a t1 m1 vf ->
+ Csharpminor.eval_exprlist prog le e m1 bl t2 m2 vargs ->
+ eval_exprlist_prop le e m1 bl t2 m2 vargs ->
Genv.find_funct ge vf = Some f ->
- Csharpminor.fn_sig f = sig ->
- Csharpminor.eval_funcall prog m2 f vargs m3 vres ->
- eval_funcall_prop m2 f vargs m3 vres ->
- eval_expr_prop le e m (Csharpminor.Ecall sig a bl) m3 vres.
+ Csharpminor.funsig f = sig ->
+ Csharpminor.eval_funcall prog m2 f vargs t3 m3 vres ->
+ eval_funcall_prop m2 f vargs t3 m3 vres ->
+ t = t1 ** t2 ** t3 ->
+ eval_expr_prop le e m (Csharpminor.Ecall sig a bl) t m3 vres.
Proof.
intros;red;intros. monadInv TR. subst ta.
- generalize (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH).
+ exploit H0; eauto.
intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
- assert (LINJ1: val_list_inject f2 le tle). eapply val_list_inject_incr; eauto.
- generalize (H2 _ _ _ _ _ _ _ _ _ _ EQ0 LINJ1 MINJ1 MATCH1).
+ exploit H2.
+ eauto. eapply val_list_inject_incr; eauto. eauto. eauto.
intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
assert (tv1 = vf).
elim (Genv.find_funct_inv H3). intros bf VF. rewrite VF in H3.
rewrite Genv.find_funct_find_funct_ptr in H3.
generalize (Genv.find_funct_ptr_inv H3). intro.
assert (match_globalenvs f2). eapply match_callstack_match_globalenvs; eauto.
- generalize (mg_functions _ H8 _ H7). intro.
+ generalize (mg_functions _ H9 _ H8). intro.
rewrite VF in VINJ1. inversion VINJ1. subst vf.
decEq. congruence.
subst ofs2. replace x with 0. reflexivity. congruence.
subst tv1. elim (functions_translated _ _ H3). intros tf [FIND TRF].
- generalize (H6 _ _ _ _ _ TRF MINJ2 MATCH2 VINJ2).
+ exploit H6; eauto.
intros [f4 [tm4 [tres [EVAL3 [VINJ3 [MINJ3 [INCR3 MATCH3]]]]]]].
exists f4; exists te3; exists tm4; exists tres. intuition.
- eapply eval_Ecall; eauto. rewrite <- H4. apply sig_transl_function; auto.
+ eapply eval_Ecall; eauto.
+ rewrite <- H4. apply sig_preserved; auto.
apply inject_incr_trans with f2; auto.
apply inject_incr_trans with f3; auto.
Qed.
Lemma transl_expr_Econdition_true_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (a b c : Csharpminor.expr) (m1 : mem) (v1 : val) (m2 : mem)
- (v2 : val),
- Csharpminor.eval_expr prog le e m a m1 v1 ->
- eval_expr_prop le e m a m1 v1 ->
+ (a b c : Csharpminor.expr) (t1: trace) (m1 : mem) (v1 : val)
+ (t2: trace) (m2 : mem) (v2 : val) (t: trace),
+ Csharpminor.eval_expr prog le e m a t1 m1 v1 ->
+ eval_expr_prop le e m a t1 m1 v1 ->
Val.is_true v1 ->
- Csharpminor.eval_expr prog le e m1 b m2 v2 ->
- eval_expr_prop le e m1 b m2 v2 ->
- eval_expr_prop le e m (Csharpminor.Econdition a b c) m2 v2.
+ Csharpminor.eval_expr prog le e m1 b t2 m2 v2 ->
+ eval_expr_prop le e m1 b t2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr_prop le e m (Csharpminor.Econdition a b c) t m2 v2.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
- assert (LINJ1: val_list_inject f2 le tle). eapply val_list_inject_incr; eauto.
- destruct (H3 _ _ _ _ _ _ _ _ _ _ EQ0 LINJ1 MINJ1 MATCH1)
- as [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
+ exploit H3.
+ eauto. eapply val_list_inject_incr; eauto. eauto. eauto.
+ intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f3; exists te3; exists tm3; exists tv2.
intuition.
- rewrite <- H5. eapply eval_conditionalexpr_true; eauto.
+ rewrite <- H6. subst t; eapply eval_conditionalexpr_true; eauto.
inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto.
eapply inject_incr_trans; eauto.
Qed.
Lemma transl_expr_Econdition_false_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (a b c : Csharpminor.expr) (m1 : mem) (v1 : val) (m2 : mem)
- (v2 : val),
- Csharpminor.eval_expr prog le e m a m1 v1 ->
- eval_expr_prop le e m a m1 v1 ->
+ (a b c : Csharpminor.expr) (t1: trace) (m1 : mem) (v1 : val)
+ (t2: trace) (m2 : mem) (v2 : val) (t: trace),
+ Csharpminor.eval_expr prog le e m a t1 m1 v1 ->
+ eval_expr_prop le e m a t1 m1 v1 ->
Val.is_false v1 ->
- Csharpminor.eval_expr prog le e m1 c m2 v2 ->
- eval_expr_prop le e m1 c m2 v2 ->
- eval_expr_prop le e m (Csharpminor.Econdition a b c) m2 v2.
+ Csharpminor.eval_expr prog le e m1 c t2 m2 v2 ->
+ eval_expr_prop le e m1 c t2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr_prop le e m (Csharpminor.Econdition a b c) t m2 v2.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
- assert (LINJ1: val_list_inject f2 le tle). eapply val_list_inject_incr; eauto.
- destruct (H3 _ _ _ _ _ _ _ _ _ _ EQ1 LINJ1 MINJ1 MATCH1)
- as [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
+ exploit H3.
+ eauto. eapply val_list_inject_incr; eauto. eauto. eauto.
+ intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f3; exists te3; exists tm3; exists tv2.
intuition.
- rewrite <- H5. eapply eval_conditionalexpr_false; eauto.
+ rewrite <- H6. subst t; eapply eval_conditionalexpr_false; eauto.
inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto.
eapply inject_incr_trans; eauto.
Qed.
Lemma transl_expr_Elet_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (a b : Csharpminor.expr) (m1 : mem) (v1 : val) (m2 : mem) (v2 : val),
- Csharpminor.eval_expr prog le e m a m1 v1 ->
- eval_expr_prop le e m a m1 v1 ->
- Csharpminor.eval_expr prog (v1 :: le) e m1 b m2 v2 ->
- eval_expr_prop (v1 :: le) e m1 b m2 v2 ->
- eval_expr_prop le e m (Csharpminor.Elet a b) m2 v2.
+ (a b : Csharpminor.expr) (t1: trace) (m1 : mem) (v1 : val)
+ (t2: trace) (m2 : mem) (v2 : val) (t: trace),
+ Csharpminor.eval_expr prog le e m a t1 m1 v1 ->
+ eval_expr_prop le e m a t1 m1 v1 ->
+ Csharpminor.eval_expr prog (v1 :: le) e m1 b t2 m2 v2 ->
+ eval_expr_prop (v1 :: le) e m1 b t2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr_prop le e m (Csharpminor.Elet a b) t m2 v2.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
- assert (LINJ1: val_list_inject f2 (v1 :: le) (tv1 :: tle)).
- constructor. auto. eapply val_list_inject_incr; eauto.
- destruct (H2 _ _ _ _ _ _ _ _ _ _ EQ0 LINJ1 MINJ1 MATCH1)
- as [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]].
+ exploit H2.
+ eauto.
+ constructor. eauto. eapply val_list_inject_incr; eauto.
+ eauto. eauto.
+ intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f3; exists te3; exists tm3; exists tv2.
intuition.
subst ta; eapply eval_Elet; eauto.
@@ -2089,19 +2061,46 @@ Lemma transl_expr_Eletvar_correct:
forall (le : list val) (e : Csharpminor.env) (m : mem) (n : nat)
(v : val),
nth_error le n = Some v ->
- eval_expr_prop le e m (Csharpminor.Eletvar n) m v.
+ eval_expr_prop le e m (Csharpminor.Eletvar n) E0 m v.
Proof.
intros; red; intros. monadInv TR.
- destruct (val_list_inject_nth _ _ _ LINJ _ _ H)
- as [tv [A B]].
+ exploit val_list_inject_nth; eauto. intros [tv [A B]].
exists f1; exists te1; exists tm1; exists tv.
intuition.
subst ta. eapply eval_Eletvar; auto.
Qed.
+Lemma transl_expr_Ealloc_correct:
+ forall (le: list val) (e: Csharpminor.env) (m1: mem) (a: Csharpminor.expr)
+ (t: trace) (m2: mem) (n: int) (m3: mem) (b: block),
+ Csharpminor.eval_expr prog le e m1 a t m2 (Vint n) ->
+ eval_expr_prop le e m1 a t m2 (Vint n) ->
+ Mem.alloc m2 0 (Int.signed n) = (m3, b) ->
+ eval_expr_prop le e m1 (Csharpminor.Ealloc a) t m3 (Vptr b Int.zero).
+Proof.
+ intros; red; intros. monadInv TR.
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ inversion VINJ1. subst tv1 i.
+ caseEq (alloc tm2 0 (Int.signed n)). intros tm3 tb TALLOC.
+ assert (LB: Int.min_signed <= 0). compute. congruence.
+ assert (HB: Int.signed n <= Int.max_signed).
+ generalize (Int.signed_range n); omega.
+ exploit alloc_parallel_inject; eauto.
+ intros [MINJ3 INCR3].
+ exists (extend_inject b (Some (tb, 0)) f2);
+ exists te2; exists tm3; exists (Vptr tb Int.zero).
+ split. subst ta; econstructor; eauto.
+ split. econstructor. unfold extend_inject, eq_block. rewrite zeq_true. reflexivity.
+ reflexivity.
+ split. assumption.
+ split. eapply inject_incr_trans; eauto.
+ eapply match_callstack_alloc; eauto.
+Qed.
+
Lemma transl_exprlist_Enil_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem),
- eval_exprlist_prop le e m Csharpminor.Enil m nil.
+ eval_exprlist_prop le e m Csharpminor.Enil E0 m nil.
Proof.
intros; red; intros. monadInv TR.
exists f1; exists te1; exists tm1; exists (@nil val).
@@ -2110,50 +2109,55 @@ Qed.
Lemma transl_exprlist_Econs_correct:
forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem)
- (a : Csharpminor.expr) (bl : Csharpminor.exprlist) (m1 : mem)
- (v : val) (m2 : mem) (vl : list val),
- Csharpminor.eval_expr prog le e m a m1 v ->
- eval_expr_prop le e m a m1 v ->
- Csharpminor.eval_exprlist prog le e m1 bl m2 vl ->
- eval_exprlist_prop le e m1 bl m2 vl ->
- eval_exprlist_prop le e m (Csharpminor.Econs a bl) m2 (v :: vl).
+ (a : Csharpminor.expr) (bl : Csharpminor.exprlist)
+ (t1: trace) (m1 : mem) (v : val)
+ (t2: trace) (m2 : mem) (vl : list val) (t: trace),
+ Csharpminor.eval_expr prog le e m a t1 m1 v ->
+ eval_expr_prop le e m a t1 m1 v ->
+ Csharpminor.eval_exprlist prog le e m1 bl t2 m2 vl ->
+ eval_exprlist_prop le e m1 bl t2 m2 vl ->
+ t = t1 ** t2 ->
+ eval_exprlist_prop le e m (Csharpminor.Econs a bl) t m2 (v :: vl).
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ LINJ MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
- assert (LINJ2: val_list_inject f2 le tle). eapply val_list_inject_incr; eauto.
- destruct (H2 _ _ _ _ _ _ _ _ _ _ EQ0 LINJ2 MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
- assert (VINJ1': val_inject f3 v tv1). eapply val_inject_incr; eauto.
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H2.
+ eauto. eapply val_list_inject_incr; eauto. eauto. eauto.
+ intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
exists f3; exists te3; exists tm3; exists (tv1 :: tv2).
intuition. subst tal; econstructor; eauto.
+ constructor. eapply val_inject_incr; eauto. auto.
eapply inject_incr_trans; eauto.
Qed.
-Lemma transl_funcall_correct:
+Lemma transl_funcall_internal_correct:
forall (m : mem) (f : Csharpminor.function) (vargs : list val)
- (e : Csharpminor.env) (m1 : mem) (lb : list block) (m2 m3 : mem)
- (out : Csharpminor.outcome) (vres : val),
+ (e : Csharpminor.env) (m1 : mem) (lb : list block) (m2: mem)
+ (t: trace) (m3 : mem) (out : Csharpminor.outcome) (vres : val),
list_norepet (fn_params_names f ++ fn_vars_names f) ->
alloc_variables empty_env m (fn_variables f) e m1 lb ->
bind_parameters e m1 (Csharpminor.fn_params f) vargs m2 ->
- Csharpminor.exec_stmt prog e m2 (Csharpminor.fn_body f) m3 out ->
- exec_stmt_prop e m2 (Csharpminor.fn_body f) m3 out ->
+ Csharpminor.exec_stmt prog e m2 (Csharpminor.fn_body f) t m3 out ->
+ exec_stmt_prop e m2 (Csharpminor.fn_body f) t m3 out ->
Csharpminor.outcome_result_value out (sig_res (Csharpminor.fn_sig f)) vres ->
- eval_funcall_prop m f vargs (free_list m3 lb) vres.
+ eval_funcall_prop m (Internal f) vargs t (free_list m3 lb) vres.
Proof.
intros; red. intros tfn f1 tm; intros.
- unfold transl_function in TR.
+ generalize TR; clear TR.
+ unfold transl_fundef, transf_partial_fundef.
+ caseEq (transl_function gce f); try congruence.
+ intros tf TR EQ. inversion EQ; clear EQ; subst tfn.
+ unfold transl_function in TR.
caseEq (build_compilenv gce f); intros cenv stacksize CENV.
rewrite CENV in TR.
destruct (zle stacksize Int.max_signed); try discriminate.
monadInv TR. clear TR.
caseEq (alloc tm 0 stacksize). intros tm1 sp ALLOC.
- destruct (function_entry_ok _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- H0 H1 MATCH CENV z ALLOC ARGSINJ MINJ H)
- as [f2 [te2 [tm2 [STOREPARAM [MINJ2 [INCR12 [MATCH2 BLOCKS]]]]]]].
- destruct (H3 _ _ _ _ _ _ _ _ _ EQ MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tout [EXECBODY [OUTINJ [MINJ3 [INCR23 MATCH3]]]]]]]].
+ exploit function_entry_ok; eauto.
+ intros [f2 [te2 [tm2 [STOREPARAM [MINJ2 [INCR12 [MATCH2 BLOCKS]]]]]]].
+ red in H3; exploit H3; eauto.
+ intros [f3 [te3 [tm3 [tout [EXECBODY [OUTINJ [MINJ3 [INCR23 MATCH3]]]]]]]].
assert (exists tvres,
outcome_result_value tout f.(Csharpminor.fn_sig).(sig_res) tvres /\
val_inject f3 vres tvres).
@@ -2167,13 +2171,14 @@ Proof.
destruct (sig_res (Csharpminor.fn_sig f)); intro.
exists v2; split. auto. subst vres; auto.
contradiction.
- elim H5; clear H5; intros tvres [TOUT VINJRES].
+ destruct H5 as [tvres [TOUT VINJRES]].
exists f3; exists (Mem.free tm3 sp); exists tvres.
(* execution *)
split. rewrite <- H6; econstructor; simpl; eauto.
- apply exec_Sseq_continue with te2 tm2.
+ apply exec_Sseq_continue with E0 te2 tm2 t.
exact STOREPARAM.
eexact EXECBODY.
+ traceEq.
(* val_inject *)
split. assumption.
(* mem_inject *)
@@ -2190,9 +2195,22 @@ Proof.
intros. elim (BLOCKS b); intros B1 B2. generalize (B2 H7). omega.
Qed.
+Lemma transl_funcall_external_correct:
+ forall (m : mem) (ef : external_function) (vargs : list val)
+ (t : trace) (vres : val),
+ event_match ef vargs t vres ->
+ eval_funcall_prop m (External ef) vargs t m vres.
+Proof.
+ intros; red; intros.
+ simpl in TR. inversion TR; clear TR; subst tfn.
+ exploit event_match_inject; eauto. intros [A B].
+ exists f1; exists tm1; exists vres; intuition.
+ constructor; auto.
+Qed.
+
Lemma transl_stmt_Sskip_correct:
forall (e : Csharpminor.env) (m : mem),
- exec_stmt_prop e m Csharpminor.Sskip m Csharpminor.Out_normal.
+ exec_stmt_prop e m Csharpminor.Sskip E0 m Csharpminor.Out_normal.
Proof.
intros; red; intros. monadInv TR.
exists f1; exists te1; exists tm1; exists Out_normal.
@@ -2201,33 +2219,90 @@ Qed.
Lemma transl_stmt_Sexpr_correct:
forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr)
- (m1 : mem) (v : val),
- Csharpminor.eval_expr prog nil e m a m1 v ->
- eval_expr_prop nil e m a m1 v ->
- exec_stmt_prop e m (Csharpminor.Sexpr a) m1 Csharpminor.Out_normal.
+ (t: trace) (m1 : mem) (v : val),
+ Csharpminor.eval_expr prog nil e m a t m1 v ->
+ eval_expr_prop nil e m a t m1 v ->
+ exec_stmt_prop e m (Csharpminor.Sexpr a) t m1 Csharpminor.Out_normal.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ (val_nil_inject f1) MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f2; exists te2; exists tm2; exists Out_normal.
intuition. subst ts. econstructor; eauto.
constructor.
Qed.
+Lemma transl_stmt_Sassign_correct:
+ forall (e : Csharpminor.env) (m : mem)
+ (id : ident) (a : Csharpminor.expr) (t: trace) (m1 : mem) (b : block)
+ (chunk : memory_chunk) (v : val) (m2 : mem),
+ Csharpminor.eval_expr prog nil e m a t m1 v ->
+ eval_expr_prop nil e m a t m1 v ->
+ eval_var_ref prog e id b chunk ->
+ store chunk m1 b 0 v = Some m2 ->
+ exec_stmt_prop e m (Csharpminor.Sassign id a) t m2 Csharpminor.Out_normal.
+Proof.
+ intros; red; intros. monadInv TR; intro EQ0.
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR12 MATCH1]]]]]]]].
+ exploit var_set_correct; eauto.
+ intros [te3 [tm3 [EVAL2 [MINJ2 MATCH2]]]].
+ exists f2; exists te3; exists tm3; exists Out_normal.
+ intuition. constructor.
+Qed.
+
+Lemma transl_stmt_Sstore_correct:
+ forall (e : Csharpminor.env) (m : mem)
+ (chunk : memory_chunk) (a b : Csharpminor.expr) (t1: trace) (m1 : mem)
+ (v1 : val) (t2: trace) (m2 : mem) (v2 : val)
+ (t3: trace) (m3 : mem),
+ Csharpminor.eval_expr prog nil e m a t1 m1 v1 ->
+ eval_expr_prop nil e m a t1 m1 v1 ->
+ Csharpminor.eval_expr prog nil e m1 b t2 m2 v2 ->
+ eval_expr_prop nil e m1 b t2 m2 v2 ->
+ storev chunk m2 v1 v2 = Some m3 ->
+ t3 = t1 ** t2 ->
+ exec_stmt_prop e m (Csharpminor.Sstore chunk a b) t3 m3 Csharpminor.Out_normal.
+Proof.
+ intros; red; intros. monadInv TR.
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H2.
+ eauto.
+ eapply val_list_inject_incr; eauto.
+ eauto. eauto.
+ intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
+ exploit make_store_correct.
+ eexact EVAL1. eexact EVAL2. eauto. eauto.
+ eapply val_inject_incr; eauto. eauto.
+ intros [tm4 [EVAL [MINJ4 NEXTBLOCK]]].
+ exists f3; exists te3; exists tm4; exists Out_normal.
+ rewrite <- H6. subst t3. intuition.
+ constructor.
+ eapply inject_incr_trans; eauto.
+ assert (val_inject f3 v1 tv1). eapply val_inject_incr; eauto.
+ unfold storev in H3; destruct v1; try discriminate.
+ inversion H4.
+ rewrite NEXTBLOCK. replace (nextblock m3) with (nextblock m2).
+ eapply match_callstack_mapped; eauto. congruence.
+ exploit store_inv; eauto. simpl; symmetry; tauto.
+Qed.
+
Lemma transl_stmt_Sseq_continue_correct:
forall (e : Csharpminor.env) (m : mem) (s1 s2 : Csharpminor.stmt)
- (m1 m2 : mem) (out : Csharpminor.outcome),
- Csharpminor.exec_stmt prog e m s1 m1 Csharpminor.Out_normal ->
- exec_stmt_prop e m s1 m1 Csharpminor.Out_normal ->
- Csharpminor.exec_stmt prog e m1 s2 m2 out ->
- exec_stmt_prop e m1 s2 m2 out ->
- exec_stmt_prop e m (Csharpminor.Sseq s1 s2) m2 out.
+ (t1 t2: trace) (m1 m2 : mem) (t: trace) (out : Csharpminor.outcome),
+ Csharpminor.exec_stmt prog e m s1 t1 m1 Csharpminor.Out_normal ->
+ exec_stmt_prop e m s1 t1 m1 Csharpminor.Out_normal ->
+ Csharpminor.exec_stmt prog e m1 s2 t2 m2 out ->
+ exec_stmt_prop e m1 s2 t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt_prop e m (Csharpminor.Sseq s1 s2) t m2 out.
Proof.
intros; red; intros; monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ EQ MINJ MATCH)
- as [f2 [te2 [tm2 [tout1 [EXEC1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
- destruct (H2 _ _ _ _ _ _ _ _ _ EQ0 MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tout2 [EXEC2 [OINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tout1 [EXEC1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H2; eauto.
+ intros [f3 [te3 [tm3 [tout2 [EXEC2 [OINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
exists f3; exists te3; exists tm3; exists tout2.
intuition. subst ts; eapply exec_Sseq_continue; eauto.
inversion OINJ1. subst tout1. auto.
@@ -2236,15 +2311,15 @@ Qed.
Lemma transl_stmt_Sseq_stop_correct:
forall (e : Csharpminor.env) (m : mem) (s1 s2 : Csharpminor.stmt)
- (m1 : mem) (out : Csharpminor.outcome),
- Csharpminor.exec_stmt prog e m s1 m1 out ->
- exec_stmt_prop e m s1 m1 out ->
+ (t1: trace) (m1 : mem) (out : Csharpminor.outcome),
+ Csharpminor.exec_stmt prog e m s1 t1 m1 out ->
+ exec_stmt_prop e m s1 t1 m1 out ->
out <> Csharpminor.Out_normal ->
- exec_stmt_prop e m (Csharpminor.Sseq s1 s2) m1 out.
+ exec_stmt_prop e m (Csharpminor.Sseq s1 s2) t1 m1 out.
Proof.
intros; red; intros; monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ EQ MINJ MATCH)
- as [f2 [te2 [tm2 [tout1 [EXEC1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tout1 [EXEC1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f2; exists te2; exists tm2; exists tout1.
intuition. subst ts; eapply exec_Sseq_stop; eauto.
inversion OINJ1; subst out tout1; congruence.
@@ -2252,64 +2327,70 @@ Qed.
Lemma transl_stmt_Sifthenelse_true_correct:
forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr)
- (sl1 sl2 : Csharpminor.stmt) (m1 : mem) (v1 : val) (m2 : mem)
- (out : Csharpminor.outcome),
- Csharpminor.eval_expr prog nil e m a m1 v1 ->
- eval_expr_prop nil e m a m1 v1 ->
+ (sl1 sl2 : Csharpminor.stmt)
+ (t1: trace) (m1 : mem) (v1 : val) (t2: trace) (m2 : mem)
+ (out : Csharpminor.outcome) (t: trace),
+ Csharpminor.eval_expr prog nil e m a t1 m1 v1 ->
+ eval_expr_prop nil e m a t1 m1 v1 ->
Val.is_true v1 ->
- Csharpminor.exec_stmt prog e m1 sl1 m2 out ->
- exec_stmt_prop e m1 sl1 m2 out ->
- exec_stmt_prop e m (Csharpminor.Sifthenelse a sl1 sl2) m2 out.
+ Csharpminor.exec_stmt prog e m1 sl1 t2 m2 out ->
+ exec_stmt_prop e m1 sl1 t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt_prop e m (Csharpminor.Sifthenelse a sl1 sl2) t m2 out.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ (val_nil_inject f1) MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
- destruct (H3 _ _ _ _ _ _ _ _ _ EQ0 MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tout [EVAL2 [OINJ [MINJ3 [INCR3 MATCH3]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H3; eauto.
+ intros [f3 [te3 [tm3 [tout [EVAL2 [OINJ [MINJ3 [INCR3 MATCH3]]]]]]]].
exists f3; exists te3; exists tm3; exists tout.
intuition.
- subst ts. eapply exec_ifthenelse_true; eauto.
+ subst ts t. eapply exec_ifthenelse_true; eauto.
inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto.
eapply inject_incr_trans; eauto.
Qed.
Lemma transl_stmt_Sifthenelse_false_correct:
forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr)
- (sl1 sl2 : Csharpminor.stmt) (m1 : mem) (v1 : val) (m2 : mem)
- (out : Csharpminor.outcome),
- Csharpminor.eval_expr prog nil e m a m1 v1 ->
- eval_expr_prop nil e m a m1 v1 ->
+ (sl1 sl2 : Csharpminor.stmt)
+ (t1: trace) (m1 : mem) (v1 : val) (t2: trace) (m2 : mem)
+ (out : Csharpminor.outcome) (t: trace),
+ Csharpminor.eval_expr prog nil e m a t1 m1 v1 ->
+ eval_expr_prop nil e m a t1 m1 v1 ->
Val.is_false v1 ->
- Csharpminor.exec_stmt prog e m1 sl2 m2 out ->
- exec_stmt_prop e m1 sl2 m2 out ->
- exec_stmt_prop e m (Csharpminor.Sifthenelse a sl1 sl2) m2 out.
+ Csharpminor.exec_stmt prog e m1 sl2 t2 m2 out ->
+ exec_stmt_prop e m1 sl2 t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt_prop e m (Csharpminor.Sifthenelse a sl1 sl2) t m2 out.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ (val_nil_inject f1) MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
- destruct (H3 _ _ _ _ _ _ _ _ _ EQ1 MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tout [EVAL2 [OINJ [MINJ3 [INCR3 MATCH3]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H3; eauto.
+ intros [f3 [te3 [tm3 [tout [EVAL2 [OINJ [MINJ3 [INCR3 MATCH3]]]]]]]].
exists f3; exists te3; exists tm3; exists tout.
intuition.
- subst ts. eapply exec_ifthenelse_false; eauto.
+ subst ts t. eapply exec_ifthenelse_false; eauto.
inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto.
eapply inject_incr_trans; eauto.
Qed.
Lemma transl_stmt_Sloop_loop_correct:
forall (e : Csharpminor.env) (m : mem) (sl : Csharpminor.stmt)
- (m1 m2 : mem) (out : Csharpminor.outcome),
- Csharpminor.exec_stmt prog e m sl m1 Csharpminor.Out_normal ->
- exec_stmt_prop e m sl m1 Csharpminor.Out_normal ->
- Csharpminor.exec_stmt prog e m1 (Csharpminor.Sloop sl) m2 out ->
- exec_stmt_prop e m1 (Csharpminor.Sloop sl) m2 out ->
- exec_stmt_prop e m (Csharpminor.Sloop sl) m2 out.
+ (t1: trace) (m1: mem) (t2: trace) (m2 : mem)
+ (out : Csharpminor.outcome) (t: trace),
+ Csharpminor.exec_stmt prog e m sl t1 m1 Csharpminor.Out_normal ->
+ exec_stmt_prop e m sl t1 m1 Csharpminor.Out_normal ->
+ Csharpminor.exec_stmt prog e m1 (Csharpminor.Sloop sl) t2 m2 out ->
+ exec_stmt_prop e m1 (Csharpminor.Sloop sl) t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt_prop e m (Csharpminor.Sloop sl) t m2 out.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ EQ MINJ MATCH)
- as [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
- destruct (H2 _ _ _ _ _ _ _ _ _ TR MINJ2 MATCH2)
- as [f3 [te3 [tm3 [tout2 [EVAL2 [OINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H2; eauto.
+ intros [f3 [te3 [tm3 [tout2 [EVAL2 [OINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]].
exists f3; exists te3; exists tm3; exists tout2.
intuition.
subst ts. eapply exec_Sloop_loop; eauto.
@@ -2317,18 +2398,17 @@ Proof.
eapply inject_incr_trans; eauto.
Qed.
-
Lemma transl_stmt_Sloop_exit_correct:
forall (e : Csharpminor.env) (m : mem) (sl : Csharpminor.stmt)
- (m1 : mem) (out : Csharpminor.outcome),
- Csharpminor.exec_stmt prog e m sl m1 out ->
- exec_stmt_prop e m sl m1 out ->
+ (t1: trace) (m1 : mem) (out : Csharpminor.outcome),
+ Csharpminor.exec_stmt prog e m sl t1 m1 out ->
+ exec_stmt_prop e m sl t1 m1 out ->
out <> Csharpminor.Out_normal ->
- exec_stmt_prop e m (Csharpminor.Sloop sl) m1 out.
+ exec_stmt_prop e m (Csharpminor.Sloop sl) t1 m1 out.
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ EQ MINJ MATCH)
- as [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f2; exists te2; exists tm2; exists tout1.
intuition. subst ts; eapply exec_Sloop_stop; eauto.
inversion OINJ1; subst out tout1; congruence.
@@ -2336,15 +2416,15 @@ Qed.
Lemma transl_stmt_Sblock_correct:
forall (e : Csharpminor.env) (m : mem) (sl : Csharpminor.stmt)
- (m1 : mem) (out : Csharpminor.outcome),
- Csharpminor.exec_stmt prog e m sl m1 out ->
- exec_stmt_prop e m sl m1 out ->
- exec_stmt_prop e m (Csharpminor.Sblock sl) m1
+ (t1: trace) (m1 : mem) (out : Csharpminor.outcome),
+ Csharpminor.exec_stmt prog e m sl t1 m1 out ->
+ exec_stmt_prop e m sl t1 m1 out ->
+ exec_stmt_prop e m (Csharpminor.Sblock sl) t1 m1
(Csharpminor.outcome_block out).
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ EQ MINJ MATCH)
- as [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]].
exists f2; exists te2; exists tm2; exists (outcome_block tout1).
intuition. subst ts; eapply exec_Sblock; eauto.
inversion OINJ1; subst out tout1; simpl.
@@ -2356,7 +2436,7 @@ Qed.
Lemma transl_stmt_Sexit_correct:
forall (e : Csharpminor.env) (m : mem) (n : nat),
- exec_stmt_prop e m (Csharpminor.Sexit n) m (Csharpminor.Out_exit n).
+ exec_stmt_prop e m (Csharpminor.Sexit n) E0 m (Csharpminor.Out_exit n).
Proof.
intros; red; intros. monadInv TR.
exists f1; exists te1; exists tm1; exists (Out_exit n).
@@ -2366,15 +2446,15 @@ Qed.
Lemma transl_stmt_Sswitch_correct:
forall (e : Csharpminor.env) (m : mem)
(a : Csharpminor.expr) (cases : list (int * nat)) (default : nat)
- (m1 : mem) (n : int),
- Csharpminor.eval_expr prog nil e m a m1 (Vint n) ->
- eval_expr_prop nil e m a m1 (Vint n) ->
- exec_stmt_prop e m (Csharpminor.Sswitch a cases default) m1
+ (t1 : trace) (m1 : mem) (n : int),
+ Csharpminor.eval_expr prog nil e m a t1 m1 (Vint n) ->
+ eval_expr_prop nil e m a t1 m1 (Vint n) ->
+ exec_stmt_prop e m (Csharpminor.Sswitch a cases default) t1 m1
(Csharpminor.Out_exit (Csharpminor.switch_target n default cases)).
Proof.
intros; red; intros. monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ (val_nil_inject f1) MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]].
exists f2; exists te2; exists tm2;
exists (Out_exit (switch_target n default cases)). intuition.
subst ts. constructor. inversion VINJ1. subst tv1. assumption.
@@ -2383,7 +2463,7 @@ Qed.
Lemma transl_stmt_Sreturn_none_correct:
forall (e : Csharpminor.env) (m : mem),
- exec_stmt_prop e m (Csharpminor.Sreturn None) m
+ exec_stmt_prop e m (Csharpminor.Sreturn None) E0 m
(Csharpminor.Out_return None).
Proof.
intros; red; intros. monadInv TR.
@@ -2393,15 +2473,15 @@ Qed.
Lemma transl_stmt_Sreturn_some_correct:
forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr)
- (m1 : mem) (v : val),
- Csharpminor.eval_expr prog nil e m a m1 v ->
- eval_expr_prop nil e m a m1 v ->
- exec_stmt_prop e m (Csharpminor.Sreturn (Some a)) m1
+ (t1: trace) (m1 : mem) (v : val),
+ Csharpminor.eval_expr prog nil e m a t1 m1 v ->
+ eval_expr_prop nil e m a t1 m1 v ->
+ exec_stmt_prop e m (Csharpminor.Sreturn (Some a)) t1 m1
(Csharpminor.Out_return (Some v)).
Proof.
intros; red; intros; monadInv TR.
- destruct (H0 _ _ _ _ _ _ _ _ _ _ EQ (val_nil_inject f1) MINJ MATCH)
- as [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]].
+ exploit H0; eauto.
+ intros [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]].
exists f2; exists te2; exists tm2; exists (Out_return (Some tv1)).
intuition. subst ts; econstructor; eauto. constructor; auto.
Qed.
@@ -2410,9 +2490,9 @@ Qed.
evaluation derivation, using the lemmas above for each case. *)
Lemma transl_function_correct:
- forall m1 f vargs m2 vres,
- Csharpminor.eval_funcall prog m1 f vargs m2 vres ->
- eval_funcall_prop m1 f vargs m2 vres.
+ forall m1 f vargs t m2 vres,
+ Csharpminor.eval_funcall prog m1 f vargs t m2 vres ->
+ eval_funcall_prop m1 f vargs t m2 vres.
Proof
(eval_funcall_ind4 prog
eval_expr_prop
@@ -2421,21 +2501,23 @@ Proof
exec_stmt_prop
transl_expr_Evar_correct
- transl_expr_Eassign_correct
transl_expr_Eaddrof_correct
transl_expr_Eop_correct
transl_expr_Eload_correct
- transl_expr_Estore_correct
transl_expr_Ecall_correct
transl_expr_Econdition_true_correct
transl_expr_Econdition_false_correct
transl_expr_Elet_correct
transl_expr_Eletvar_correct
+ transl_expr_Ealloc_correct
transl_exprlist_Enil_correct
transl_exprlist_Econs_correct
- transl_funcall_correct
+ transl_funcall_internal_correct
+ transl_funcall_external_correct
transl_stmt_Sskip_correct
transl_stmt_Sexpr_correct
+ transl_stmt_Sassign_correct
+ transl_stmt_Sstore_correct
transl_stmt_Sseq_continue_correct
transl_stmt_Sseq_stop_correct
transl_stmt_Sifthenelse_true_correct
@@ -2472,11 +2554,11 @@ Qed.
follows. *)
Theorem transl_program_correct:
- forall n,
- Csharpminor.exec_program prog (Vint n) ->
- exec_program tprog (Vint n).
+ forall t n,
+ Csharpminor.exec_program prog t (Vint n) ->
+ exec_program tprog t (Vint n).
Proof.
- intros n [b [fn [m [FINDS [FINDF [SIG EVAL]]]]]].
+ intros t n [b [fn [m [FINDS [FINDF [SIG EVAL]]]]]].
elim (function_ptr_translated _ _ FINDF). intros tfn [TFIND TR].
set (m0 := Genv.init_mem (program_of_program prog)) in *.
set (f := fun b => if zlt b m0.(nextblock) then Some(b, 0) else None).
@@ -2488,26 +2570,25 @@ Proof.
split. auto.
constructor. compute. split; congruence. left; auto.
intros; omega.
- generalize (Genv.initmem_undef (program_of_program prog) b0). fold m0. intros [lo [hi EQ]].
- rewrite EQ. simpl. constructor.
+ generalize (Genv.initmem_block_init (program_of_program prog) b0). fold m0. intros [idata EQ].
+ rewrite EQ. simpl. apply Mem.contents_init_data_inject.
destruct (zlt b1 (nextblock m0)); try discriminate.
destruct (zlt b2 (nextblock m0)); try discriminate.
left; congruence.
assert (MATCH0: match_callstack f nil m0.(nextblock) m0.(nextblock) m0).
constructor. unfold f; apply match_globalenvs_init.
fold ge in EVAL.
- destruct (transl_function_correct _ _ _ _ _ EVAL
- _ _ _ _ _ TR MINJ0 MATCH0 (val_nil_inject f))
- as [f1 [tm1 [tres [TEVAL [VINJ [MINJ1 [INCR MATCH1]]]]]]].
+ exploit transl_function_correct; eauto.
+ intros [f1 [tm1 [tres [TEVAL [VINJ [MINJ1 [INCR MATCH1]]]]]]].
exists b; exists tfn; exists tm1.
split. fold tge. rewrite <- FINDS.
replace (prog_main prog) with (AST.prog_main tprog). fold ge. apply symbols_preserved.
transitivity (AST.prog_main (program_of_program prog)).
- apply transform_partial_program_main with (transl_function gce). assumption.
+ apply transform_partial_program_main with (transl_fundef gce). assumption.
reflexivity.
split. assumption.
- split. rewrite <- SIG. apply sig_transl_function; auto.
- rewrite (Genv.init_mem_transf_partial (transl_function gce) _ TRANSL).
+ split. rewrite <- SIG. apply sig_preserved; auto.
+ rewrite (Genv.init_mem_transf_partial (transl_fundef gce) _ TRANSL).
inversion VINJ; subst tres. assumption.
Qed.
diff --git a/backend/Coloring.v b/backend/Coloring.v
index 1a34a124..0a2487cb 100644
--- a/backend/Coloring.v
+++ b/backend/Coloring.v
@@ -150,6 +150,12 @@ Definition add_edges_instr
(add_interf_op res live
(add_interf_call
(Regset.remove res live) destroyed_at_call_regs g)))
+ | Ialloc arg res s =>
+ add_pref_mreg arg loc_alloc_argument
+ (add_pref_mreg res loc_alloc_result
+ (add_interf_op res live
+ (add_interf_call
+ (Regset.remove res live) destroyed_at_call_regs g)))
| Ireturn (Some r) =>
add_pref_mreg r (loc_result sig) g
| _ => g
diff --git a/backend/Coloringproof.v b/backend/Coloringproof.v
index 39b208ec..54d24cc4 100644
--- a/backend/Coloringproof.v
+++ b/backend/Coloringproof.v
@@ -332,6 +332,10 @@ Proof.
eapply graph_incl_trans; [idtac|apply add_pref_mreg_incl].
eapply graph_incl_trans; [idtac|apply add_interf_op_incl].
apply add_interf_call_incl.
+ eapply graph_incl_trans; [idtac|apply add_pref_mreg_incl].
+ eapply graph_incl_trans; [idtac|apply add_pref_mreg_incl].
+ eapply graph_incl_trans; [idtac|apply add_interf_op_incl].
+ apply add_interf_call_incl.
destruct o.
apply add_pref_mreg_incl.
apply graph_incl_refl.
@@ -370,6 +374,15 @@ Definition correct_interf_instr
/\ (forall r,
Regset.mem r live = true ->
r <> res -> interfere r res g)
+ | Ialloc arg res s =>
+ (forall r mr,
+ Regset.mem r live = true ->
+ In mr destroyed_at_call_regs ->
+ r <> res ->
+ interfere_mreg r mr g)
+ /\ (forall r,
+ Regset.mem r live = true ->
+ r <> res -> interfere r res g)
| _ =>
True
end.
@@ -389,6 +402,9 @@ Proof.
intros [A B]. split.
intros. eapply interfere_mreg_incl; eauto.
intros. eapply interfere_incl; eauto.
+ intros [A B]. split.
+ intros. eapply interfere_mreg_incl; eauto.
+ intros. eapply interfere_incl; eauto.
Qed.
Lemma add_edges_instr_correct:
@@ -417,6 +433,19 @@ Proof.
eapply graph_incl_trans; [idtac|apply add_prefs_call_incl].
apply add_pref_mreg_incl.
apply add_interf_op_correct; auto.
+
+ split; intros.
+ apply interfere_mreg_incl with
+ (add_interf_call (Regset.remove r0 live) destroyed_at_call_regs g).
+ eapply graph_incl_trans; [idtac|apply add_pref_mreg_incl].
+ eapply graph_incl_trans; [idtac|apply add_pref_mreg_incl].
+ apply add_interf_op_incl.
+ apply add_interf_call_correct; auto.
+ rewrite Regset.mem_remove_other; auto.
+
+ eapply interfere_incl.
+ eapply graph_incl_trans; apply add_pref_mreg_incl.
+ apply add_interf_op_correct; auto.
Qed.
Lemma add_edges_instrs_incl_aux:
@@ -541,7 +570,7 @@ Proof.
then false else true)).
red. unfold OrderedRegReg.eq. unfold OrderedReg.eq.
intros x y [EQ1 EQ2]. rewrite EQ1; rewrite EQ2; auto.
- generalize (SetRegReg.for_all_2 H1 H H0).
+ generalize (SetRegReg.for_all_2 _ _ H1 H _ H0).
simpl. case (Loc.eq (coloring r1) (coloring r2)); intro.
intro; discriminate. auto.
Qed.
@@ -558,7 +587,7 @@ Proof.
then false else true)).
red. unfold OrderedRegMreg.eq. unfold OrderedReg.eq.
intros x y [EQ1 EQ2]. rewrite EQ1; rewrite EQ2; auto.
- generalize (SetRegMreg.for_all_2 H1 H H0).
+ generalize (SetRegMreg.for_all_2 _ _ H1 H _ H0).
simpl. case (Loc.eq (coloring r1) (R mr2)); intro.
intro; discriminate. auto.
Qed.
@@ -702,6 +731,14 @@ Definition correct_alloc_instr
/\ (forall r,
Regset.mem r live!!pc = true ->
r <> res -> alloc r <> alloc res)
+ | Ialloc arg res s =>
+ (forall r,
+ Regset.mem r live!!pc = true ->
+ r <> res ->
+ ~(In (alloc r) Conventions.destroyed_at_call))
+ /\ (forall r,
+ Regset.mem r live!!pc = true ->
+ r <> res -> alloc r <> alloc res)
| _ =>
True
end.
@@ -780,6 +817,14 @@ Proof.
generalize (A r0 mr H IN H0). intro.
generalize (ALL2 _ _ H2). contradiction.
auto.
+ intros [A B]. split.
+ intros; red; intros.
+ unfold destroyed_at_call in H1.
+ generalize (list_in_map_inv R _ _ H1).
+ intros [mr [EQ IN]].
+ generalize (A r1 mr H IN H0). intro.
+ generalize (ALL2 _ _ H2). contradiction.
+ auto.
Qed.
Lemma regalloc_correct_1:
diff --git a/backend/Constprop.v b/backend/Constprop.v
index b1c5a2bb..3820311c 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -195,7 +195,9 @@ Definition eval_static_operation (op: operation) (vl: list approx) :=
| Ofloatconst n, nil => F n
| Oaddrsymbol s n, nil => S s n
| Ocast8signed, I n1 :: nil => I(Int.cast8signed n)
+ | Ocast8unsigned, I n1 :: nil => I(Int.cast8unsigned n)
| Ocast16signed, I n1 :: nil => I(Int.cast16signed n)
+ | Ocast16unsigned, I n1 :: nil => I(Int.cast16unsigned n)
| Oadd, I n1 :: I n2 :: nil => I(Int.add n1 n2)
| Oadd, S s1 n1 :: I n2 :: nil => S s1 (Int.add n1 n2)
| Oaddimm n, I n1 :: nil => I (Int.add n1 n)
@@ -379,6 +381,12 @@ Inductive eval_static_operation_cases: forall (op: operation) (vl: list approx),
| eval_static_operation_case47:
forall c vl,
eval_static_operation_cases (Ocmp c) (vl)
+ | eval_static_operation_case48:
+ forall n1,
+ eval_static_operation_cases (Ocast8unsigned) (I n1 :: nil)
+ | eval_static_operation_case49:
+ forall n1,
+ eval_static_operation_cases (Ocast16unsigned) (I n1 :: nil)
| eval_static_operation_default:
forall (op: operation) (vl: list approx),
eval_static_operation_cases op vl.
@@ -475,6 +483,10 @@ Definition eval_static_operation_match (op: operation) (vl: list approx) :=
eval_static_operation_case46 n1
| Ocmp c, vl =>
eval_static_operation_case47 c vl
+ | Ocast8unsigned, I n1 :: nil =>
+ eval_static_operation_case48 n1
+ | Ocast16unsigned, I n1 :: nil =>
+ eval_static_operation_case49 n1
| op, vl =>
eval_static_operation_default op vl
end.
@@ -574,6 +586,10 @@ Definition eval_static_operation (op: operation) (vl: list approx) :=
| None => Unknown
| Some b => I(if b then Int.one else Int.zero)
end
+ | eval_static_operation_case48 n1 =>
+ I(Int.cast8unsigned n1)
+ | eval_static_operation_case49 n1 =>
+ I(Int.cast16unsigned n1)
| eval_static_operation_default op vl =>
Unknown
end.
@@ -603,6 +619,8 @@ Definition transfer (f: function) (pc: node) (before: D.t) :=
before
| Icall sig ros args res s =>
D.set res Unknown before
+ | Ialloc arg res s =>
+ D.set res Unknown before
| Icond cond args ifso ifnot =>
before
| Ireturn optarg =>
@@ -986,6 +1004,8 @@ Definition transf_instr (approx: D.t) (instr: instruction) :=
| inr s => ros
end in
Icall sig ros' args res s
+ | Ialloc arg res s =>
+ Ialloc arg res s
| Icond cond args s1 s2 =>
match eval_static_condition cond (approx_regs args approx) with
| Some b =>
@@ -1028,5 +1048,7 @@ Definition transf_function (f: function) : function :=
(transf_code_wf f approxs f.(fn_code_wf))
end.
+Definition transf_fundef := AST.transf_fundef transf_function.
+
Definition transf_program (p: program) : program :=
- transform_program transf_function p.
+ transform_program transf_fundef p.
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index 080aa74d..38ba38b8 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -6,6 +6,7 @@ Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Values.
+Require Import Events.
Require Import Mem.
Require Import Globalenvs.
Require Import Op.
@@ -144,11 +145,12 @@ Proof.
case (eval_static_operation_match op al); intros;
InvVLMA; simpl in *; FuncInv; try congruence.
- replace v with v0. auto. congruence.
-
destruct (Genv.find_symbol ge s). exists b. intuition congruence.
congruence.
+ rewrite <- H3. replace v0 with (Vint n1). reflexivity. congruence.
+ rewrite <- H3. replace v0 with (Vint n1). reflexivity. congruence.
+
exists b. split. auto. congruence.
exists b. split. auto. congruence.
exists b. split. auto. congruence.
@@ -178,12 +180,17 @@ Proof.
replace n2 with i0. destruct (Int.ltu i0 (Int.repr 32)).
injection H0; intro; subst v. simpl. congruence. discriminate. congruence.
+ rewrite <- H3. replace v0 with (Vfloat n1). reflexivity. congruence.
+
caseEq (eval_static_condition c vl0).
intros. generalize (eval_static_condition_correct _ _ _ _ H H1).
intro. rewrite H2 in H0.
destruct b; injection H0; intro; subst v; simpl; auto.
intros; simpl; auto.
+ rewrite <- H3. replace v0 with (Vint n1). reflexivity. congruence.
+ rewrite <- H3. replace v0 with (Vint n1). reflexivity. congruence.
+
auto.
Qed.
@@ -193,8 +200,8 @@ Qed.
the static approximation returned by the transfer function. *)
Lemma transfer_correct:
- forall f c sp pc rs m pc' rs' m' ra,
- exec_instr ge c sp pc rs m pc' rs' m' ->
+ forall f c sp pc rs m t pc' rs' m' ra,
+ exec_instr ge c sp pc rs m t pc' rs' m' ->
c = f.(fn_code) ->
regs_match_approx ra rs ->
regs_match_approx (transfer f pc ra) rs'.
@@ -208,6 +215,8 @@ Proof.
apply regs_match_approx_update; auto. simpl; auto.
(* Icall *)
apply regs_match_approx_update; auto. simpl; auto.
+ (* Ialloc *)
+ apply regs_match_approx_update; auto. simpl; auto.
Qed.
(** The correctness of the static analysis follows from the results
@@ -217,8 +226,8 @@ Qed.
Lemma analyze_correct_1:
forall f approxs,
analyze f = Some approxs ->
- forall c sp pc rs m pc' rs' m',
- exec_instr ge c sp pc rs m pc' rs' m' ->
+ forall c sp pc rs m t pc' rs' m',
+ exec_instr ge c sp pc rs m t pc' rs' m' ->
c = f.(fn_code) ->
regs_match_approx approxs!!pc rs ->
regs_match_approx approxs!!pc' rs'.
@@ -228,7 +237,7 @@ Proof.
eapply DS.fixpoint_solution.
unfold analyze in H. eexact H.
elim (fn_code_wf f pc); intro. auto.
- generalize (exec_instr_present _ _ _ _ _ _ _ _ _ H0).
+ generalize (exec_instr_present _ _ _ _ _ _ _ _ _ _ H0).
rewrite H1. intro. contradiction.
eapply successors_correct. rewrite <- H1. eauto.
eapply transfer_correct; eauto.
@@ -237,8 +246,8 @@ Qed.
Lemma analyze_correct_2:
forall f approxs,
analyze f = Some approxs ->
- forall c sp pc rs m pc' rs' m',
- exec_instrs ge c sp pc rs m pc' rs' m' ->
+ forall c sp pc rs m t pc' rs' m',
+ exec_instrs ge c sp pc rs m t pc' rs' m' ->
c = f.(fn_code) ->
regs_match_approx approxs!!pc rs ->
regs_match_approx approxs!!pc' rs'.
@@ -638,19 +647,28 @@ Let tge := Genv.globalenv tprog.
Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof (Genv.find_symbol_transf transf_function prog).
+Proof (Genv.find_symbol_transf transf_fundef prog).
Lemma functions_translated:
- forall (v: val) (f: RTL.function),
+ forall (v: val) (f: RTL.fundef),
Genv.find_funct ge v = Some f ->
- Genv.find_funct tge v = Some (transf_function f).
-Proof (@Genv.find_funct_transf _ _ transf_function prog).
+ Genv.find_funct tge v = Some (transf_fundef f).
+Proof (@Genv.find_funct_transf _ _ transf_fundef prog).
Lemma function_ptr_translated:
- forall (v: block) (f: RTL.function),
+ forall (v: block) (f: RTL.fundef),
Genv.find_funct_ptr ge v = Some f ->
- Genv.find_funct_ptr tge v = Some (transf_function f).
-Proof (@Genv.find_funct_ptr_transf _ _ transf_function prog).
+ Genv.find_funct_ptr tge v = Some (transf_fundef f).
+Proof (@Genv.find_funct_ptr_transf _ _ transf_fundef prog).
+
+Lemma sig_translated:
+ forall (f: RTL.fundef),
+ funsig (transf_fundef f) = funsig f.
+Proof.
+ intro. case f; intros; simpl.
+ unfold transf_function. case (analyze f0); intros; reflexivity.
+ reflexivity.
+Qed.
(** The proof of semantic preservation is a simulation argument
based on diagrams of the following form:
@@ -676,30 +694,30 @@ Proof (@Genv.find_funct_ptr_transf _ _ transf_function prog).
Definition exec_instr_prop
(c: code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
- exec_instr tge c sp pc rs m pc' rs' m' /\
+ exec_instr tge c sp pc rs m t pc' rs' m' /\
forall f approxs
(CF: c = f.(RTL.fn_code))
(ANL: analyze f = Some approxs)
(MATCH: regs_match_approx ge approxs!!pc rs),
- exec_instr tge (transf_code approxs c) sp pc rs m pc' rs' m'.
+ exec_instr tge (transf_code approxs c) sp pc rs m t pc' rs' m'.
Definition exec_instrs_prop
(c: code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
- exec_instrs tge c sp pc rs m pc' rs' m' /\
+ exec_instrs tge c sp pc rs m t pc' rs' m' /\
forall f approxs
(CF: c = f.(RTL.fn_code))
(ANL: analyze f = Some approxs)
(MATCH: regs_match_approx ge approxs!!pc rs),
- exec_instrs tge (transf_code approxs c) sp pc rs m pc' rs' m'.
+ exec_instrs tge (transf_code approxs c) sp pc rs m t pc' rs' m'.
Definition exec_function_prop
- (f: RTL.function) (args: list val) (m: mem)
+ (f: RTL.fundef) (args: list val) (m: mem) (t: trace)
(res: val) (m': mem) : Prop :=
- exec_function tge (transf_function f) args m res m'.
+ exec_function tge (transf_fundef f) args m t res m'.
Ltac TransfInstr :=
match goal with
@@ -716,9 +734,9 @@ Ltac TransfInstr :=
evaluation derivation of the original code. *)
Lemma transf_funct_correct:
- forall f args m res m',
- exec_function ge f args m res m' ->
- exec_function_prop f args m res m'.
+ forall f args m t res m',
+ exec_function ge f args m t res m' ->
+ exec_function_prop f args m t res m'.
Proof.
apply (exec_function_ind_3 ge
exec_instr_prop exec_instrs_prop exec_function_prop);
@@ -773,13 +791,13 @@ Proof.
rewrite ASR; simpl. congruence.
intro. eapply exec_Istore; eauto.
(* Icall *)
- assert (find_function tge ros rs = Some (transf_function f)).
+ assert (find_function tge ros rs = Some (transf_fundef f)).
generalize H0; unfold find_function; destruct ros.
apply functions_translated.
rewrite symbols_preserved. destruct (Genv.find_symbol ge i).
apply function_ptr_translated. congruence.
- assert (sig = fn_sig (transf_function f)).
- rewrite H1. unfold transf_function. destruct (analyze f); reflexivity.
+ assert (funsig (transf_fundef f) = sig).
+ generalize (sig_translated f). congruence.
split; [idtac| intros; TransfInstr].
eapply exec_Icall; eauto.
set (ros' :=
@@ -803,6 +821,11 @@ Proof.
generalize H4. simpl. rewrite A. rewrite B. subst i0.
rewrite Genv.find_funct_find_funct_ptr. auto.
+ (* Ialloc *)
+ split; [idtac|intros; TransfInstr].
+ eapply exec_Ialloc; eauto.
+ intros. eapply exec_Ialloc; eauto.
+
(* Icond, true *)
split; [idtac| intros; TransfInstr].
eapply exec_Icond_true; eauto.
@@ -844,37 +867,38 @@ Proof.
(* trans *)
elim H0; intros. elim H2; intros.
split.
- apply exec_trans with pc2 rs2 m2; auto.
- intros; apply exec_trans with pc2 rs2 m2.
- eapply H4; eauto. eapply H6; eauto.
- eapply analyze_correct_2; eauto.
+ apply exec_trans with t1 pc2 rs2 m2 t2; auto.
+ intros; apply exec_trans with t1 pc2 rs2 m2 t2.
+ eapply H5; eauto. eapply H7; eauto.
+ eapply analyze_correct_2; eauto. auto.
- (* function *)
+ (* internal function *)
elim H1; intros.
- unfold transf_function.
+ simpl. unfold transf_function.
caseEq (analyze f).
intros approxs ANL.
- eapply exec_funct; simpl; eauto.
+ eapply exec_funct_internal; simpl; eauto.
eapply H5. reflexivity. auto.
apply analyze_correct_3; auto.
TransfInstr; auto.
- intros. eapply exec_funct; eauto.
+ intros. eapply exec_funct_internal; eauto.
+ (* external function *)
+ unfold transf_function; simpl. apply exec_funct_external; auto.
Qed.
(** The preservation of the observable behavior of the program then
follows. *)
Theorem transf_program_correct:
- forall (r: val),
- exec_program prog r -> exec_program tprog r.
+ forall (t: trace) (r: val),
+ exec_program prog t r -> exec_program tprog t r.
Proof.
- intros r [b [f [m [SYMB [FIND [SIG EXEC]]]]]].
- red. exists b. exists (transf_function f). exists m.
+ intros t r [b [f [m [SYMB [FIND [SIG EXEC]]]]]].
+ red. exists b. exists (transf_fundef f). exists m.
split. change (prog_main tprog) with (prog_main prog).
rewrite symbols_preserved. auto.
split. apply function_ptr_translated; auto.
- split. unfold transf_function.
- rewrite <- SIG. destruct (analyze f); reflexivity.
+ split. generalize (sig_translated f). congruence.
apply transf_funct_correct.
unfold tprog, transf_program. rewrite Genv.init_mem_transf.
exact EXEC.
diff --git a/backend/Conventions.v b/backend/Conventions.v
index 99cc9338..5b4222df 100644
--- a/backend/Conventions.v
+++ b/backend/Conventions.v
@@ -19,12 +19,11 @@ Require Import Locations.
of callee- and caller-save registers.
*)
-Definition destroyed_at_call_regs :=
- R3 :: R4 :: R5 :: R6 :: R7 :: R8 :: R9 :: R10 ::
- F1 :: F2 :: F3 :: F4 :: F5 :: F6 :: F7 :: F8 :: F9 :: F10 :: nil.
+Definition int_caller_save_regs :=
+ R3 :: R4 :: R5 :: R6 :: R7 :: R8 :: R9 :: R10 :: nil.
-Definition destroyed_at_call :=
- List.map R destroyed_at_call_regs.
+Definition float_caller_save_regs :=
+ F1 :: F2 :: F3 :: F4 :: F5 :: F6 :: F7 :: F8 :: F9 :: F10 :: nil.
Definition int_callee_save_regs :=
R13 :: R14 :: R15 :: R16 :: R17 :: R18 :: R19 :: R20 :: R21 :: R22 ::
@@ -34,6 +33,12 @@ Definition float_callee_save_regs :=
F14 :: F15 :: F16 :: F17 :: F18 :: F19 :: F20 :: F21 :: F22 ::
F23 :: F24 :: F25 :: F26 :: F27 :: F28 :: F29 :: F30 :: F31 :: nil.
+Definition destroyed_at_call_regs :=
+ int_caller_save_regs ++ float_caller_save_regs.
+
+Definition destroyed_at_call :=
+ List.map R destroyed_at_call_regs.
+
Definition temporaries :=
R IT1 :: R IT2 :: R IT3 :: R FT1 :: R FT2 :: R FT3 :: nil.
@@ -298,6 +303,18 @@ Proof.
simpl; OrEq.
Qed.
+(** The result location is not a callee-save register. *)
+
+Lemma loc_result_not_callee_save:
+ forall (s: signature),
+ ~(In (loc_result s) int_callee_save_regs \/ In (loc_result s) float_callee_save_regs).
+Proof.
+ intros. generalize (loc_result_acceptable s).
+ generalize (int_callee_save_not_destroyed (loc_result s)).
+ generalize (float_callee_save_not_destroyed (loc_result s)).
+ tauto.
+Qed.
+
(** ** Location of function arguments *)
(** The PowerPC ABI states the following convention for passing arguments
@@ -316,11 +333,6 @@ Qed.
These conventions are somewhat baroque, but they are mandated by the ABI.
*)
-Definition drop1 (x: list mreg) :=
- match x with nil => nil | hd :: tl => tl end.
-Definition drop2 (x: list mreg) :=
- match x with nil => nil | hd :: nil => nil | hd1 :: hd2 :: tl => tl end.
-
Fixpoint loc_arguments_rec
(tyl: list typ) (iregl: list mreg) (fregl: list mreg)
(ofs: Z) {struct tyl} : list loc :=
@@ -331,13 +343,13 @@ Fixpoint loc_arguments_rec
| nil => S (Outgoing ofs Tint)
| ireg :: _ => R ireg
end ::
- loc_arguments_rec tys (drop1 iregl) fregl (ofs + 1)
+ loc_arguments_rec tys (list_drop1 iregl) fregl (ofs + 1)
| Tfloat :: tys =>
match fregl with
| nil => S (Outgoing ofs Tfloat)
| freg :: _ => R freg
end ::
- loc_arguments_rec tys (drop2 iregl) (drop1 fregl) (ofs + 2)
+ loc_arguments_rec tys (list_drop2 iregl) (list_drop1 fregl) (ofs + 2)
end.
Definition int_param_regs :=
@@ -374,18 +386,6 @@ Definition loc_argument_acceptable (l: loc) : Prop :=
| _ => False
end.
-Remark drop1_incl:
- forall x l, In x (drop1 l) -> In x l.
-Proof.
- intros. destruct l. elim H. simpl in H. auto with coqlib.
-Qed.
-Remark drop2_incl:
- forall x l, In x (drop2 l) -> In x l.
-Proof.
- intros. destruct l. elim H. destruct l. elim H.
- simpl in H. auto with coqlib.
-Qed.
-
Remark loc_arguments_rec_charact:
forall tyl iregl fregl ofs l,
In l (loc_arguments_rec tyl iregl fregl ofs) ->
@@ -400,12 +400,12 @@ Proof.
destruct a; elim H; intros.
destruct iregl; subst l. omega. left; auto with coqlib.
generalize (IHtyl _ _ _ _ H0).
- destruct l. intros [A|B]. left; apply drop1_incl; auto. tauto.
+ destruct l. intros [A|B]. left; apply list_drop1_incl; auto. tauto.
destruct s; try contradiction. omega.
destruct fregl; subst l. omega. right; auto with coqlib.
generalize (IHtyl _ _ _ _ H0).
- destruct l. intros [A|B]. left; apply drop2_incl; auto.
- right; apply drop1_incl; auto.
+ destruct l. intros [A|B]. left; apply list_drop2_incl; auto.
+ right; apply list_drop1_incl; auto.
destruct s; try contradiction. omega.
Qed.
@@ -425,18 +425,6 @@ Hint Resolve loc_arguments_acceptable: locs.
(** Arguments are parwise disjoint (in the sense of [Loc.norepet]). *)
-Remark drop1_norepet:
- forall l, list_norepet l -> list_norepet (drop1 l).
-Proof.
- intros. destruct l; simpl. constructor. inversion H. auto.
-Qed.
-Remark drop2_norepet:
- forall l, list_norepet l -> list_norepet (drop2 l).
-Proof.
- intros. destruct l; simpl. constructor.
- destruct l; simpl. constructor. inversion H. inversion H3. auto.
-Qed.
-
Remark loc_arguments_rec_notin_reg:
forall tyl iregl fregl ofs r,
~(In r iregl) -> ~(In r fregl) ->
@@ -446,10 +434,10 @@ Proof.
auto.
destruct a; simpl; split.
destruct iregl. auto. red; intro; subst m. apply H. auto with coqlib.
- apply IHtyl. red; intro. apply H. apply drop1_incl; auto. auto.
+ apply IHtyl. red; intro. apply H. apply list_drop1_incl; auto. auto.
destruct fregl. auto. red; intro; subst m. apply H0. auto with coqlib.
- apply IHtyl. red; intro. apply H. apply drop2_incl; auto.
- red; intro. apply H0. apply drop1_incl; auto.
+ apply IHtyl. red; intro. apply H. apply list_drop2_incl; auto.
+ red; intro. apply H0. apply list_drop1_incl; auto.
Qed.
Remark loc_arguments_rec_notin_local:
@@ -494,16 +482,16 @@ Proof.
apply loc_arguments_rec_notin_outgoing. simpl; omega.
apply loc_arguments_rec_notin_reg. simpl. inversion H. auto.
apply list_disjoint_notin with (m :: iregl); auto with coqlib.
- apply IHtyl. apply drop1_norepet; auto. auto.
- red; intros. apply H1. apply drop1_incl; auto. auto.
+ apply IHtyl. apply list_drop1_norepet; auto. auto.
+ red; intros. apply H1. apply list_drop1_incl; auto. auto.
destruct fregl.
apply loc_arguments_rec_notin_outgoing. simpl; omega.
apply loc_arguments_rec_notin_reg. simpl.
- red; intro. apply (H1 m m). apply drop2_incl; auto.
+ red; intro. apply (H1 m m). apply list_drop2_incl; auto.
auto with coqlib. auto.
simpl. inversion H0. auto.
- apply IHtyl. apply drop2_norepet; auto. apply drop1_norepet; auto.
- red; intros. apply H1. apply drop2_incl; auto. apply drop1_incl; auto.
+ apply IHtyl. apply list_drop2_norepet; auto. apply list_drop1_norepet; auto.
+ red; intros. apply H1. apply list_drop2_incl; auto. apply list_drop1_incl; auto.
intro. unfold loc_arguments. apply H.
unfold int_param_regs. NoRepet.
@@ -601,11 +589,11 @@ Proof.
destruct a; simpl; apply (f_equal2 (@cons typ)).
destruct iregl. reflexivity. simpl. apply H. auto with coqlib.
apply IHtyl.
- intros. apply H. apply drop1_incl. auto. auto.
+ intros. apply H. apply list_drop1_incl. auto. auto.
destruct fregl. reflexivity. simpl. apply H0. auto with coqlib.
apply IHtyl.
- intros. apply H. apply drop2_incl. auto.
- intros. apply H0. apply drop1_incl. auto.
+ intros. apply H. apply list_drop2_incl. auto.
+ intros. apply H0. apply list_drop1_incl. auto.
intros. unfold loc_arguments. apply H.
intro; simpl. ElimOrEq; reflexivity.
@@ -688,3 +676,30 @@ Proof.
intros; simpl. tauto.
Qed.
+(** ** Location of arguments to external functions *)
+
+Definition loc_external_arguments (s: signature) : list mreg :=
+ List.map
+ (fun l => match l with R r => r | S _ => IT1 end)
+ (loc_arguments s).
+
+Definition sig_external_ok (s: signature) : Prop :=
+ forall sl, ~In (S sl) (loc_arguments s).
+
+Lemma loc_external_arguments_loc_arguments:
+ forall s,
+ sig_external_ok s ->
+ loc_arguments s = List.map R (loc_external_arguments s).
+Proof.
+ intros. unfold loc_external_arguments.
+ rewrite list_map_compose.
+ transitivity (List.map (fun x => x) (loc_arguments s)).
+ symmetry; apply list_map_identity.
+ apply list_map_exten; intros.
+ destruct x. auto. elim (H _ H0).
+Qed.
+
+(** ** Location of argument and result of dynamic allocation *)
+
+Definition loc_alloc_argument := R3.
+Definition loc_alloc_result := R3.
diff --git a/backend/Csharpminor.v b/backend/Csharpminor.v
index 49fd3df3..246ebf53 100644
--- a/backend/Csharpminor.v
+++ b/backend/Csharpminor.v
@@ -7,13 +7,14 @@ Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
(** Abstract syntax *)
-(** Cminor is a low-level imperative language structured in expressions,
+(** Csharpminor is a low-level imperative language structured in expressions,
statements, functions and programs. Expressions include
- reading and writing local variables, reading and writing store locations,
+ reading global or local variables, reading store locations,
arithmetic operations, function calls, and conditional expressions
(similar to [e1 ? e2 : e3] in C). The [Elet] and [Eletvar] constructs
enable sharing the computations of subexpressions. De Bruijn notation
@@ -67,20 +68,20 @@ Inductive operation : Set :=
Inductive expr : Set :=
| Evar : ident -> expr (**r reading a scalar variable *)
| Eaddrof : ident -> expr (**r taking the address of a variable *)
- | Eassign : ident -> expr -> expr (**r assignment to a scalar variable *)
| Eop : operation -> exprlist -> expr (**r arithmetic operation *)
| Eload : memory_chunk -> expr -> expr (**r memory read *)
- | Estore : memory_chunk -> expr -> expr -> expr (**r memory write *)
| Ecall : signature -> expr -> exprlist -> expr (**r function call *)
| Econdition : expr -> expr -> expr -> expr (**r conditional expression *)
| Elet : expr -> expr -> expr (**r let binding *)
| Eletvar : nat -> expr (**r reference to a let-bound variable *)
+ | Ealloc : expr -> expr (**r memory allocation *)
with exprlist : Set :=
| Enil: exprlist
| Econs: expr -> exprlist -> exprlist.
-(** Statements include expression evaluation, an if/then/else conditional,
+(** Statements include expression evaluation, variable assignment,
+ memory stores, an if/then/else conditional,
infinite loops, blocks and early block exits, and early function returns.
[Sexit n] terminates prematurely the execution of the [n+1] enclosing
[Sblock] statements. *)
@@ -88,6 +89,8 @@ with exprlist : Set :=
Inductive stmt : Set :=
| Sskip: stmt
| Sexpr: expr -> stmt
+ | Sassign : ident -> expr -> stmt
+ | Sstore : memory_chunk -> expr -> expr -> stmt
| Sseq: stmt -> stmt -> stmt
| Sifthenelse: expr -> stmt -> stmt -> stmt
| Sloop: stmt -> stmt
@@ -117,12 +120,20 @@ Record function : Set := mkfunction {
fn_body: stmt
}.
+Definition fundef := AST.fundef function.
+
Record program : Set := mkprogram {
- prog_funct: list (ident * function);
+ prog_funct: list (ident * fundef);
prog_main: ident;
- prog_vars: list (ident * var_kind)
+ prog_vars: list (ident * var_kind * list init_data)
}.
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => f.(fn_sig)
+ | External ef => ef.(ef_sig)
+ end.
+
(** * Operational semantics *)
(** The operational semantics for Csharpminor is given in big-step operational
@@ -164,7 +175,7 @@ Fixpoint switch_target (n: int) (dfl: nat) (cases: list (int * nat))
- [env]: local environments, map local variables to memory blocks + var info;
- [lenv]: let environments, map de Bruijn indices to values.
*)
-Definition genv := Genv.t function.
+Definition genv := Genv.t fundef.
Definition gvarenv := PTree.t var_kind.
Definition env := PTree.t (block * var_kind).
Definition empty_env : env := PTree.empty (block * var_kind).
@@ -176,11 +187,11 @@ Definition sizeof (lv: var_kind) : Z :=
| Varray sz => Zmax 0 sz
end.
-Definition program_of_program (p: program): AST.program function :=
+Definition program_of_program (p: program): AST.program fundef :=
AST.mkprogram
p.(prog_funct)
p.(prog_main)
- (List.map (fun id_vi => (fst id_vi, sizeof (snd id_vi))) p.(prog_vars)).
+ (List.map (fun x => match x with (id, k, init) => (id, init) end) p.(prog_vars)).
Definition fn_variables (f: function) :=
List.map
@@ -195,7 +206,7 @@ Definition fn_vars_names (f: function) :=
Definition global_var_env (p: program): gvarenv :=
List.fold_left
- (fun gve id_vi => PTree.set (fst id_vi) (snd id_vi) gve)
+ (fun gve x => match x with (id, k, init) => PTree.set id k gve end)
p.(prog_vars) (PTree.empty var_kind).
(** Evaluation of operator applications. *)
@@ -270,22 +281,6 @@ Definition eval_operation (op: operation) (vl: list val) (m: mem): option val :=
| _, _ => None
end.
-(** ``Casting'' a value to a memory chunk. The value is truncated and
- zero- or sign-extended as dictated by the memory chunk. *)
-
-Definition cast (chunk: memory_chunk) (v: val) : option val :=
- match chunk, v with
- | Mint8signed, Vint n => Some (Vint (Int.cast8signed n))
- | Mint8unsigned, Vint n => Some (Vint (Int.cast8unsigned n))
- | Mint16signed, Vint n => Some (Vint (Int.cast16signed n))
- | Mint16unsigned, Vint n => Some (Vint (Int.cast16unsigned n))
- | Mint32, Vint n => Some(Vint n)
- | Mint32, Vptr b ofs => Some(Vptr b ofs)
- | Mfloat32, Vfloat f => Some(Vfloat(Float.singleoffloat f))
- | Mfloat64, Vfloat f => Some(Vfloat f)
- | _, _ => None
- end.
-
(** Allocation of local variables at function entry. Each variable is
bound to the reference to a fresh block of the appropriate size. *)
@@ -312,10 +307,9 @@ Inductive bind_parameters: env ->
forall e m,
bind_parameters e m nil nil m
| bind_parameters_cons:
- forall e m id chunk params v1 v2 vl b m1 m2,
+ forall e m id chunk params v1 vl b m1 m2,
PTree.get id e = Some (b, Vscalar chunk) ->
- cast chunk v1 = Some v2 ->
- Mem.store chunk m b 0 v2 = Some m1 ->
+ Mem.store chunk m b 0 v1 = Some m1 ->
bind_parameters e m1 params vl m2 ->
bind_parameters e m ((id, chunk) :: params) (v1 :: vl) m2.
@@ -364,69 +358,64 @@ Inductive eval_var_ref: env -> ident -> block -> memory_chunk -> Prop :=
Inductive eval_expr:
letenv -> env ->
- mem -> expr -> mem -> val -> Prop :=
+ mem -> expr -> trace -> mem -> val -> Prop :=
| eval_Evar:
forall le e m id b chunk v,
eval_var_ref e id b chunk ->
Mem.load chunk m b 0 = Some v ->
- eval_expr le e m (Evar id) m v
- | eval_Eassign:
- forall le e m id a m1 b chunk v1 v2 m2,
- eval_expr le e m a m1 v1 ->
- eval_var_ref e id b chunk ->
- cast chunk v1 = Some v2 ->
- Mem.store chunk m1 b 0 v2 = Some m2 ->
- eval_expr le e m (Eassign id a) m2 v2
+ eval_expr le e m (Evar id) E0 m v
| eval_Eaddrof:
forall le e m id b,
eval_var_addr e id b ->
- eval_expr le e m (Eaddrof id) m (Vptr b Int.zero)
+ eval_expr le e m (Eaddrof id) E0 m (Vptr b Int.zero)
| eval_Eop:
- forall le e m op al m1 vl v,
- eval_exprlist le e m al m1 vl ->
+ forall le e m op al t m1 vl v,
+ eval_exprlist le e m al t m1 vl ->
eval_operation op vl m1 = Some v ->
- eval_expr le e m (Eop op al) m1 v
+ eval_expr le e m (Eop op al) t m1 v
| eval_Eload:
- forall le e m chunk a m1 v1 v,
- eval_expr le e m a m1 v1 ->
+ forall le e m chunk a t m1 v1 v,
+ eval_expr le e m a t m1 v1 ->
Mem.loadv chunk m1 v1 = Some v ->
- eval_expr le e m (Eload chunk a) m1 v
- | eval_Estore:
- forall le e m chunk a b m1 v1 m2 v2 m3 v3,
- eval_expr le e m a m1 v1 ->
- eval_expr le e m1 b m2 v2 ->
- cast chunk v2 = Some v3 ->
- Mem.storev chunk m2 v1 v3 = Some m3 ->
- eval_expr le e m (Estore chunk a b) m3 v3
+ eval_expr le e m (Eload chunk a) t m1 v
| eval_Ecall:
- forall le e m sig a bl m1 m2 m3 vf vargs vres f,
- eval_expr le e m a m1 vf ->
- eval_exprlist le e m1 bl m2 vargs ->
+ forall le e m sig a bl t1 m1 t2 m2 t3 m3 vf vargs vres f t,
+ eval_expr le e m a t1 m1 vf ->
+ eval_exprlist le e m1 bl t2 m2 vargs ->
Genv.find_funct ge vf = Some f ->
- f.(fn_sig) = sig ->
- eval_funcall m2 f vargs m3 vres ->
- eval_expr le e m (Ecall sig a bl) m3 vres
+ funsig f = sig ->
+ eval_funcall m2 f vargs t3 m3 vres ->
+ t = t1 ** t2 ** t3 ->
+ eval_expr le e m (Ecall sig a bl) t m3 vres
| eval_Econdition_true:
- forall le e m a b c m1 v1 m2 v2,
- eval_expr le e m a m1 v1 ->
+ forall le e m a b c t1 m1 v1 t2 m2 v2 t,
+ eval_expr le e m a t1 m1 v1 ->
Val.is_true v1 ->
- eval_expr le e m1 b m2 v2 ->
- eval_expr le e m (Econdition a b c) m2 v2
+ eval_expr le e m1 b t2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr le e m (Econdition a b c) t m2 v2
| eval_Econdition_false:
- forall le e m a b c m1 v1 m2 v2,
- eval_expr le e m a m1 v1 ->
+ forall le e m a b c t1 m1 v1 t2 m2 v2 t,
+ eval_expr le e m a t1 m1 v1 ->
Val.is_false v1 ->
- eval_expr le e m1 c m2 v2 ->
- eval_expr le e m (Econdition a b c) m2 v2
+ eval_expr le e m1 c t2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr le e m (Econdition a b c) t m2 v2
| eval_Elet:
- forall le e m a b m1 v1 m2 v2,
- eval_expr le e m a m1 v1 ->
- eval_expr (v1::le) e m1 b m2 v2 ->
- eval_expr le e m (Elet a b) m2 v2
+ forall le e m a b t1 m1 v1 t2 m2 v2 t,
+ eval_expr le e m a t1 m1 v1 ->
+ eval_expr (v1::le) e m1 b t2 m2 v2 ->
+ t = t1 ** t2 ->
+ eval_expr le e m (Elet a b) t m2 v2
| eval_Eletvar:
forall le e m n v,
nth_error le n = Some v ->
- eval_expr le e m (Eletvar n) m v
+ eval_expr le e m (Eletvar n) E0 m v
+ | eval_Ealloc:
+ forall le e m a t m1 n m2 b,
+ eval_expr le e m a t m1 (Vint n) ->
+ Mem.alloc m1 0 (Int.signed n) = (m2, b) ->
+ eval_expr le e m (Ealloc a) t m2 (Vptr b Int.zero)
(** Evaluation of a list of expressions:
[eval_exprlist prg le al m a m' vl]
@@ -437,32 +426,37 @@ Inductive eval_expr:
with eval_exprlist:
letenv -> env ->
- mem -> exprlist ->
+ mem -> exprlist -> trace ->
mem -> list val -> Prop :=
| eval_Enil:
forall le e m,
- eval_exprlist le e m Enil m nil
+ eval_exprlist le e m Enil E0 m nil
| eval_Econs:
- forall le e m a bl m1 v m2 vl,
- eval_expr le e m a m1 v ->
- eval_exprlist le e m1 bl m2 vl ->
- eval_exprlist le e m (Econs a bl) m2 (v :: vl)
+ forall le e m a bl t1 m1 v t2 m2 vl t,
+ eval_expr le e m a t1 m1 v ->
+ eval_exprlist le e m1 bl t2 m2 vl ->
+ t = t1 ** t2 ->
+ eval_exprlist le e m (Econs a bl) t m2 (v :: vl)
(** Evaluation of a function invocation: [eval_funcall prg m f args m' res]
means that the function [f], applied to the arguments [args] in
memory state [m], returns the value [res] in modified memory state [m'].
*)
with eval_funcall:
- mem -> function -> list val ->
+ mem -> fundef -> list val -> trace ->
mem -> val -> Prop :=
- | eval_funcall_intro:
- forall m f vargs e m1 lb m2 m3 out vres,
+ | eval_funcall_internal:
+ forall m f vargs e m1 lb m2 t m3 out vres,
list_norepet (fn_params_names f ++ fn_vars_names f) ->
alloc_variables empty_env m (fn_variables f) e m1 lb ->
bind_parameters e m1 f.(fn_params) vargs m2 ->
- exec_stmt e m2 f.(fn_body) m3 out ->
+ exec_stmt e m2 f.(fn_body) t m3 out ->
outcome_result_value out f.(fn_sig).(sig_res) vres ->
- eval_funcall m f vargs (Mem.free_list m3 lb) vres
+ eval_funcall m (Internal f) vargs t (Mem.free_list m3 lb) vres
+ | eval_funcall_external:
+ forall m ef vargs t vres,
+ event_match ef vargs t vres ->
+ eval_funcall m (External ef) vargs t m vres
(** Execution of a statement: [exec_stmt prg e m s m' out]
means that statement [s] executes with outcome [out].
@@ -470,66 +464,83 @@ with eval_funcall:
with exec_stmt:
env ->
- mem -> stmt ->
+ mem -> stmt -> trace ->
mem -> outcome -> Prop :=
| exec_Sskip:
forall e m,
- exec_stmt e m Sskip m Out_normal
+ exec_stmt e m Sskip E0 m Out_normal
| exec_Sexpr:
- forall e m a m1 v,
- eval_expr nil e m a m1 v ->
- exec_stmt e m (Sexpr a) m1 Out_normal
+ forall e m a t m1 v,
+ eval_expr nil e m a t m1 v ->
+ exec_stmt e m (Sexpr a) t m1 Out_normal
+ | eval_Sassign:
+ forall e m id a t m1 b chunk v m2,
+ eval_expr nil e m a t m1 v ->
+ eval_var_ref e id b chunk ->
+ Mem.store chunk m1 b 0 v = Some m2 ->
+ exec_stmt e m (Sassign id a) t m2 Out_normal
+ | eval_Sstore:
+ forall e m chunk a b t1 m1 v1 t2 m2 v2 t3 m3,
+ eval_expr nil e m a t1 m1 v1 ->
+ eval_expr nil e m1 b t2 m2 v2 ->
+ Mem.storev chunk m2 v1 v2 = Some m3 ->
+ t3 = t1 ** t2 ->
+ exec_stmt e m (Sstore chunk a b) t3 m3 Out_normal
| exec_Sseq_continue:
- forall e m s1 s2 m1 m2 out,
- exec_stmt e m s1 m1 Out_normal ->
- exec_stmt e m1 s2 m2 out ->
- exec_stmt e m (Sseq s1 s2) m2 out
+ forall e m s1 s2 t1 t2 m1 m2 t out,
+ exec_stmt e m s1 t1 m1 Out_normal ->
+ exec_stmt e m1 s2 t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt e m (Sseq s1 s2) t m2 out
| exec_Sseq_stop:
- forall e m s1 s2 m1 out,
- exec_stmt e m s1 m1 out ->
+ forall e m s1 s2 t1 m1 out,
+ exec_stmt e m s1 t1 m1 out ->
out <> Out_normal ->
- exec_stmt e m (Sseq s1 s2) m1 out
+ exec_stmt e m (Sseq s1 s2) t1 m1 out
| exec_Sifthenelse_true:
- forall e m a sl1 sl2 m1 v1 m2 out,
- eval_expr nil e m a m1 v1 ->
+ forall e m a sl1 sl2 t1 m1 v1 t2 m2 out t,
+ eval_expr nil e m a t1 m1 v1 ->
Val.is_true v1 ->
- exec_stmt e m1 sl1 m2 out ->
- exec_stmt e m (Sifthenelse a sl1 sl2) m2 out
+ exec_stmt e m1 sl1 t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt e m (Sifthenelse a sl1 sl2) t m2 out
| exec_Sifthenelse_false:
- forall e m a sl1 sl2 m1 v1 m2 out,
- eval_expr nil e m a m1 v1 ->
+ forall e m a sl1 sl2 t1 m1 v1 t2 m2 out t,
+ eval_expr nil e m a t1 m1 v1 ->
Val.is_false v1 ->
- exec_stmt e m1 sl2 m2 out ->
- exec_stmt e m (Sifthenelse a sl1 sl2) m2 out
+ exec_stmt e m1 sl2 t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt e m (Sifthenelse a sl1 sl2) t m2 out
| exec_Sloop_loop:
- forall e m sl m1 m2 out,
- exec_stmt e m sl m1 Out_normal ->
- exec_stmt e m1 (Sloop sl) m2 out ->
- exec_stmt e m (Sloop sl) m2 out
+ forall e m sl t1 m1 t2 m2 out t,
+ exec_stmt e m sl t1 m1 Out_normal ->
+ exec_stmt e m1 (Sloop sl) t2 m2 out ->
+ t = t1 ** t2 ->
+ exec_stmt e m (Sloop sl) t m2 out
| exec_Sloop_stop:
- forall e m sl m1 out,
- exec_stmt e m sl m1 out ->
+ forall e m sl t1 m1 out,
+ exec_stmt e m sl t1 m1 out ->
out <> Out_normal ->
- exec_stmt e m (Sloop sl) m1 out
+ exec_stmt e m (Sloop sl) t1 m1 out
| exec_Sblock:
- forall e m sl m1 out,
- exec_stmt e m sl m1 out ->
- exec_stmt e m (Sblock sl) m1 (outcome_block out)
+ forall e m sl t1 m1 out,
+ exec_stmt e m sl t1 m1 out ->
+ exec_stmt e m (Sblock sl) t1 m1 (outcome_block out)
| exec_Sexit:
forall e m n,
- exec_stmt e m (Sexit n) m (Out_exit n)
+ exec_stmt e m (Sexit n) E0 m (Out_exit n)
| exec_Sswitch:
- forall e m a cases default m1 n,
- eval_expr nil e m a m1 (Vint n) ->
+ forall e m a cases default t1 m1 n,
+ eval_expr nil e m a t1 m1 (Vint n) ->
exec_stmt e m (Sswitch a cases default)
- m1 (Out_exit (switch_target n default cases))
+ t1 m1 (Out_exit (switch_target n default cases))
| exec_Sreturn_none:
forall e m,
- exec_stmt e m (Sreturn None) m (Out_return None)
+ exec_stmt e m (Sreturn None) E0 m (Out_return None)
| exec_Sreturn_some:
- forall e m a m1 v,
- eval_expr nil e m a m1 v ->
- exec_stmt e m (Sreturn (Some a)) m1 (Out_return (Some v)).
+ forall e m a t1 m1 v,
+ eval_expr nil e m a t1 m1 v ->
+ exec_stmt e m (Sreturn (Some a)) t1 m1 (Out_return (Some v)).
Scheme eval_expr_ind4 := Minimality for eval_expr Sort Prop
with eval_exprlist_ind4 := Minimality for eval_exprlist Sort Prop
@@ -542,12 +553,11 @@ End RELSEM.
holds if the application of [p]'s main function to no arguments
in the initial memory state for [p] eventually returns value [r]. *)
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv (program_of_program p) in
let m0 := Genv.init_mem (program_of_program p) in
exists b, exists f, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- f.(fn_sig) = mksignature nil (Some Tint) /\
- eval_funcall p m0 f nil m r.
-
+ funsig f = mksignature nil (Some Tint) /\
+ eval_funcall p m0 f nil t m r.
diff --git a/backend/Events.v b/backend/Events.v
new file mode 100644
index 00000000..a0559fd0
--- /dev/null
+++ b/backend/Events.v
@@ -0,0 +1,103 @@
+(** Representation of (traces of) observable events. *)
+
+Require Import Coqlib.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Values.
+
+Inductive eventval: Set :=
+ | EVint: int -> eventval
+ | EVfloat: float -> eventval.
+
+Parameter trace: Set.
+Parameter E0: trace.
+Parameter Eextcall: ident -> list eventval -> eventval -> trace.
+Parameter Eapp: trace -> trace -> trace.
+
+Infix "**" := Eapp (at level 60, right associativity).
+
+Axiom E0_left: forall t, E0 ** t = t.
+Axiom E0_right: forall t, t ** E0 = t.
+Axiom Eapp_assoc: forall t1 t2 t3, (t1 ** t2) ** t3 = t1 ** (t2 ** t3).
+
+Hint Rewrite E0_left E0_right Eapp_assoc: trace_rewrite.
+
+Ltac substTraceHyp :=
+ match goal with
+ | [ H: (@eq trace ?x ?y) |- _ ] =>
+ subst x || clear H
+ end.
+
+Ltac decomposeTraceEq :=
+ match goal with
+ | [ |- (_ ** _) = (_ ** _) ] =>
+ apply (f_equal2 Eapp); auto; decomposeTraceEq
+ | _ =>
+ auto
+ end.
+
+Ltac traceEq :=
+ repeat substTraceHyp; autorewrite with trace_rewrite; decomposeTraceEq.
+
+Inductive eventval_match: eventval -> typ -> val -> Prop :=
+ | ev_match_int:
+ forall i, eventval_match (EVint i) Tint (Vint i)
+ | ev_match_float:
+ forall f, eventval_match (EVfloat f) Tfloat (Vfloat f).
+
+Inductive eventval_list_match: list eventval -> list typ -> list val -> Prop :=
+ | evl_match_nil:
+ eventval_list_match nil nil nil
+ | evl_match_cons:
+ forall ev1 evl ty1 tyl v1 vl,
+ eventval_match ev1 ty1 v1 ->
+ eventval_list_match evl tyl vl ->
+ eventval_list_match (ev1::evl) (ty1::tyl) (v1::vl).
+
+Inductive event_match:
+ external_function -> list val -> trace -> val -> Prop :=
+ event_match_intro:
+ forall ef vargs vres eargs eres,
+ eventval_list_match eargs (sig_args ef.(ef_sig)) vargs ->
+ eventval_match eres (proj_sig_res ef.(ef_sig)) vres ->
+ event_match ef vargs (Eextcall ef.(ef_id) eargs eres) vres.
+
+Require Import Mem.
+
+Section EVENT_MATCH_INJECT.
+
+Variable f: meminj.
+
+Remark eventval_match_inject:
+ forall ev ty v1, eventval_match ev ty v1 ->
+ forall v2, val_inject f v1 v2 ->
+ eventval_match ev ty v2.
+Proof.
+ induction 1; intros; inversion H; constructor.
+Qed.
+
+Remark eventval_list_match_inject:
+ forall evl tyl vl1, eventval_list_match evl tyl vl1 ->
+ forall vl2, val_list_inject f vl1 vl2 ->
+ eventval_list_match evl tyl vl2.
+Proof.
+ induction 1; intros.
+ inversion H; constructor.
+ inversion H1; constructor.
+ eapply eventval_match_inject; eauto.
+ eauto.
+Qed.
+
+Lemma event_match_inject:
+ forall ef args1 t res args2,
+ event_match ef args1 t res ->
+ val_list_inject f args1 args2 ->
+ event_match ef args2 t res /\ val_inject f res res.
+Proof.
+ intros. inversion H; subst.
+ split. constructor. eapply eventval_list_match_inject; eauto. auto.
+ inversion H2; constructor.
+Qed.
+
+End EVENT_MATCH_INJECT.
diff --git a/backend/Globalenvs.v b/backend/Globalenvs.v
index 55afc353..036fd8f6 100644
--- a/backend/Globalenvs.v
+++ b/backend/Globalenvs.v
@@ -70,15 +70,20 @@ Module Type GENV.
(forall id f, In (id, f) (prog_funct p) -> P f) ->
find_funct (globalenv p) v = Some f ->
P f.
+ Hypothesis find_funct_ptr_symbol_inversion:
+ forall (F: Set) (p: program F) (id: ident) (b: block) (f: F),
+ find_symbol (globalenv p) id = Some b ->
+ find_funct_ptr (globalenv p) b = Some f ->
+ In (id, f) p.(prog_funct).
+
Hypothesis initmem_nullptr:
forall (F: Set) (p: program F),
let m := init_mem p in
valid_block m nullptr /\
m.(blocks) nullptr = empty_block 0 0.
- Hypothesis initmem_undef:
+ Hypothesis initmem_block_init:
forall (F: Set) (p: program F) (b: block),
- exists lo, exists hi,
- (init_mem p).(blocks) b = empty_block lo hi.
+ exists id, (init_mem p).(blocks) b = block_init_data id.
Hypothesis find_funct_ptr_inv:
forall (F: Set) (p: program F) (b: block) (f: F),
find_funct_ptr (globalenv p) b = Some f -> b < 0.
@@ -206,12 +211,12 @@ Definition add_functs (init: genv) (fns: list (ident * funct)) : genv :=
List.fold_right add_funct init fns.
Definition add_globals
- (init: genv * mem) (vars: list (ident * Z)) : genv * mem :=
+ (init: genv * mem) (vars: list (ident * list init_data)) : genv * mem :=
List.fold_right
- (fun (id_sz: ident * Z) (g_st: genv * mem) =>
- let (id, sz) := id_sz in
+ (fun (id_init: ident * list init_data) (g_st: genv * mem) =>
+ let (id, init) := id_init in
let (g, st) := g_st in
- let (st', b) := Mem.alloc st 0 sz in
+ let (st', b) := Mem.alloc_init_data st init in
(add_symbol id b g, st'))
init vars.
@@ -229,7 +234,7 @@ Lemma functions_globalenv:
forall (p: program funct),
functions (globalenv p) = functions (add_functs empty p.(prog_funct)).
Proof.
- assert (forall (init: genv * mem) (vars: list (ident * Z)),
+ assert (forall init vars,
functions (fst (add_globals init vars)) = functions (fst init)).
induction vars; simpl.
reflexivity.
@@ -248,11 +253,11 @@ Lemma initmem_nullptr:
m.(blocks) nullptr = mkblock 0 0 (fun y => Undef) (undef_undef_outside 0 0).
Proof.
assert
- (forall (init: genv * mem),
+ (forall init,
let m1 := snd init in
0 < m1.(nextblock) ->
m1.(blocks) nullptr = mkblock 0 0 (fun y => Undef) (undef_undef_outside 0 0) ->
- forall (vars: list (ident * Z)),
+ forall vars,
let m2 := snd (add_globals init vars) in
0 < m2.(nextblock) /\
m2.(blocks) nullptr = mkblock 0 0 (fun y => Undef) (undef_undef_outside 0 0)).
@@ -268,23 +273,21 @@ Proof.
unfold valid_block. apply H. simpl. omega. reflexivity.
Qed.
-Lemma initmem_undef:
+Lemma initmem_block_init:
forall (p: program funct) (b: block),
- exists lo, exists hi,
- (init_mem p).(blocks) b = empty_block lo hi.
+ exists id, (init_mem p).(blocks) b = block_init_data id.
Proof.
assert (forall g0 vars g1 m b,
add_globals (g0, Mem.empty) vars = (g1, m) ->
- exists lo, exists hi,
- m.(blocks) b = empty_block lo hi).
+ exists id, m.(blocks) b = block_init_data id).
induction vars; simpl.
intros. inversion H. unfold Mem.empty; simpl.
- exists 0; exists 0. auto.
+ exists (@nil init_data). symmetry. apply Mem.block_init_data_empty.
destruct a. caseEq (add_globals (g0, Mem.empty) vars). intros g1 m1 EQ.
intros g m b EQ1. injection EQ1; intros EQ2 EQ3; clear EQ1.
rewrite <- EQ2; simpl. unfold update.
case (zeq b (nextblock m1)); intro.
- exists 0; exists z; auto.
+ exists l; auto.
eauto.
intros. caseEq (globalenv_initmem p).
intros g m EQ. unfold init_mem; rewrite EQ; simpl.
@@ -372,6 +375,59 @@ Proof.
intros. eapply find_funct_ptr_prop; eauto.
Qed.
+Lemma find_funct_ptr_symbol_inversion:
+ forall (F: Set) (p: program F) (id: ident) (b: block) (f: F),
+ find_symbol (globalenv p) id = Some b ->
+ find_funct_ptr (globalenv p) b = Some f ->
+ In (id, f) p.(prog_funct).
+Proof.
+ intros until f.
+ assert (forall fns,
+ let g := add_functs (empty F) fns in
+ PTree.get id g.(symbols) = Some b ->
+ b > g.(nextfunction)).
+ induction fns; simpl.
+ rewrite PTree.gempty. congruence.
+ rewrite PTree.gsspec. case (peq id (fst a)); intro.
+ intro EQ. inversion EQ. unfold Zpred. omega.
+ intros. generalize (IHfns H). unfold Zpred; omega.
+ assert (forall fns,
+ let g := add_functs (empty F) fns in
+ PTree.get id g.(symbols) = Some b ->
+ ZMap.get b g.(functions) = Some f ->
+ In (id, f) fns).
+ induction fns; simpl.
+ rewrite ZMap.gi. congruence.
+ set (g := add_functs (empty F) fns).
+ rewrite PTree.gsspec. rewrite ZMap.gsspec.
+ case (peq id (fst a)); intro.
+ intro EQ. inversion EQ. unfold ZIndexed.eq. rewrite zeq_true.
+ intro EQ2. left. destruct a. simpl in *. congruence.
+ intro. unfold ZIndexed.eq. rewrite zeq_false. intro. eauto.
+ generalize (H _ H0). fold g. unfold block. omega.
+ assert (forall g0 m0, b < 0 ->
+ forall vars g m,
+ @add_globals F (g0, m0) vars = (g, m) ->
+ PTree.get id g.(symbols) = Some b ->
+ PTree.get id g0.(symbols) = Some b).
+ induction vars; simpl.
+ intros. inversion H2. auto.
+ destruct a. caseEq (add_globals (g0, m0) vars).
+ intros g1 m1 EQ g m EQ1. injection EQ1; simpl; clear EQ1.
+ unfold add_symbol; intros A B. rewrite <- B. simpl.
+ rewrite PTree.gsspec. case (peq id i); intros.
+ assert (b > 0). injection H2; intros. rewrite <- H3. apply nextblock_pos.
+ omegaContradiction.
+ eauto.
+ intros.
+ generalize (find_funct_ptr_inv _ _ H3). intro.
+ pose (g := add_functs (empty F) (prog_funct p)).
+ apply H0.
+ apply H1 with Mem.empty (prog_vars p) (globalenv p) (init_mem p).
+ auto. unfold globalenv, init_mem. rewrite <- surjective_pairing.
+ reflexivity. assumption. rewrite <- functions_globalenv. assumption.
+Qed.
+
(* Global environments and program transformations. *)
Section TRANSF_PROGRAM_PARTIAL.
@@ -420,7 +476,7 @@ Proof.
Qed.
Lemma mem_add_globals_transf:
- forall (g1: genv A) (g2: genv B) (m: mem) (vars: list (ident * Z)),
+ forall (g1: genv A) (g2: genv B) (m: mem) (vars: list (ident * list init_data)),
snd (add_globals (g1, m) vars) = snd (add_globals (g2, m) vars).
Proof.
induction vars; simpl.
@@ -433,7 +489,7 @@ Qed.
Lemma symbols_add_globals_transf:
forall (g1: genv A) (g2: genv B) (m: mem),
symbols g1 = symbols g2 ->
- forall (vars: list (ident * Z)),
+ forall (vars: list (ident * list init_data)),
symbols (fst (add_globals (g1, m) vars)) =
symbols (fst (add_globals (g2, m) vars)).
Proof.
diff --git a/backend/InterfGraph.v b/backend/InterfGraph.v
index 37248f58..78112c33 100644
--- a/backend/InterfGraph.v
+++ b/backend/InterfGraph.v
@@ -1,7 +1,8 @@
(** Representation of interference graphs for register allocation. *)
Require Import Coqlib.
-Require Import FSet.
+Require Import FSets.
+Require Import FSetAVL.
Require Import Maps.
Require Import Ordered.
Require Import Registers.
@@ -39,10 +40,14 @@ Module OrderedRegReg := OrderedPair(OrderedReg)(OrderedReg).
Module OrderedMreg := OrderedIndexed(IndexedMreg).
Module OrderedRegMreg := OrderedPair(OrderedReg)(OrderedMreg).
+(*
Module SetDepRegReg := FSetAVL.Make(OrderedRegReg).
Module SetRegReg := NodepOfDep(SetDepRegReg).
Module SetDepRegMreg := FSetAVL.Make(OrderedRegMreg).
Module SetRegMreg := NodepOfDep(SetDepRegMreg).
+*)
+Module SetRegReg := FSetAVL.Make(OrderedRegReg).
+Module SetRegMreg := FSetAVL.Make(OrderedRegMreg).
Record graph: Set := mkgraph {
interf_reg_reg: SetRegReg.t;
@@ -197,12 +202,15 @@ Qed.
(** [all_interf_regs g] returns the set of pseudo-registers that
are nodes of [g]. *)
+Definition add_intf2 (r1r2: reg * reg) (u: Regset.t) : Regset.t :=
+ Regset.add (fst r1r2) (Regset.add (snd r1r2) u).
+Definition add_intf1 (r1m2: reg * mreg) (u: Regset.t) : Regset.t :=
+ Regset.add (fst r1m2) u.
+
Definition all_interf_regs (g: graph) : Regset.t :=
- SetRegReg.fold
- (fun r1r2 u => Regset.add (fst r1r2) (Regset.add (snd r1r2) u))
+ SetRegReg.fold _ add_intf2
g.(interf_reg_reg)
- (SetRegMreg.fold
- (fun r1m2 u => Regset.add (fst r1m2) u)
+ (SetRegMreg.fold _ add_intf1
g.(interf_reg_mreg)
Regset.empty).
@@ -215,53 +223,63 @@ Proof.
rewrite Regset.mem_add_other; auto.
Qed.
-Lemma all_interf_regs_correct_aux_1:
- forall l u r,
- Regset.mem r u = true ->
- Regset.mem r
- (List.fold_right
- (fun r1r2 u => Regset.add (fst r1r2) (Regset.add (snd r1r2) u))
- u l) = true.
+Lemma in_setregreg_fold:
+ forall g r1 r2 u,
+ SetRegReg.In (r1, r2) g \/ Regset.mem r1 u = true /\ Regset.mem r2 u = true ->
+ Regset.mem r1 (SetRegReg.fold _ add_intf2 g u) = true /\
+ Regset.mem r2 (SetRegReg.fold _ add_intf2 g u) = true.
Proof.
- induction l; simpl; intros.
- auto.
- apply mem_add_tail. apply mem_add_tail. auto.
+ set (add_intf2' := fun u r1r2 => add_intf2 r1r2 u).
+ assert (forall l r1 r2 u,
+ InA OrderedRegReg.eq (r1,r2) l \/ Regset.mem r1 u = true /\ Regset.mem r2 u = true ->
+ Regset.mem r1 (List.fold_left add_intf2' l u) = true /\
+ Regset.mem r2 (List.fold_left add_intf2' l u) = true).
+ induction l; intros; simpl.
+ elim H; intro. inversion H0. auto.
+ apply IHl. destruct a as [a1 a2].
+ elim H; intro. inversion H0; subst.
+ red in H2. simpl in H2. destruct H2. red in H1. red in H2. subst r1 r2.
+ right; unfold add_intf2', add_intf2; simpl; split.
+ apply Regset.mem_add_same. apply mem_add_tail. apply Regset.mem_add_same.
+ tauto.
+ right; unfold add_intf2', add_intf2; simpl; split;
+ apply mem_add_tail; apply mem_add_tail; tauto.
+
+ intros. rewrite SetRegReg.fold_1. apply H.
+ intuition. left. apply SetRegReg.elements_1. auto.
Qed.
-Lemma all_interf_regs_correct_aux_2:
- forall l u r1 r2,
- InList OrderedRegReg.eq (r1, r2) l ->
- let u' :=
- List.fold_right
- (fun r1r2 u => Regset.add (fst r1r2) (Regset.add (snd r1r2) u))
- u l in
- Regset.mem r1 u' = true /\ Regset.mem r2 u' = true.
+Lemma in_setregreg_fold':
+ forall g r u,
+ Regset.mem r u = true ->
+ Regset.mem r (SetRegReg.fold _ add_intf2 g u) = true.
Proof.
- induction l; simpl; intros.
- inversion H.
- inversion H. elim H1. simpl. unfold OrderedReg.eq.
- intros; subst r1; subst r2.
- split. apply Regset.mem_add_same.
- apply mem_add_tail. apply Regset.mem_add_same.
- generalize (IHl u r1 r2 H1). intros [A B].
- split; repeat rewrite mem_add_tail; auto.
+ intros. exploit in_setregreg_fold. eauto.
+ intros [A B]. eauto.
Qed.
-Lemma all_interf_regs_correct_aux_3:
- forall l u r1 r2,
- InList OrderedRegMreg.eq (r1, r2) l ->
- let u' :=
- List.fold_right
- (fun r1r2 u => Regset.add (fst r1r2) u)
- u l in
- Regset.mem r1 u' = true.
+Lemma in_setregmreg_fold:
+ forall g r1 mr2 u,
+ SetRegMreg.In (r1, mr2) g \/ Regset.mem r1 u = true ->
+ Regset.mem r1 (SetRegMreg.fold _ add_intf1 g u) = true.
Proof.
- induction l; simpl; intros.
- inversion H.
- inversion H. elim H1. simpl. unfold OrderedReg.eq.
- intros; subst r1.
+ set (add_intf1' := fun u r1mr2 => add_intf1 r1mr2 u).
+ assert (forall l r1 mr2 u,
+ InA OrderedRegMreg.eq (r1,mr2) l \/ Regset.mem r1 u = true ->
+ Regset.mem r1 (List.fold_left add_intf1' l u) = true).
+ induction l; intros; simpl.
+ elim H; intro. inversion H0. auto.
+ apply IHl with mr2. destruct a as [a1 a2].
+ elim H; intro. inversion H0; subst.
+ red in H2. simpl in H2. destruct H2. red in H1. red in H2. subst r1 mr2.
+ right; unfold add_intf1', add_intf1; simpl.
apply Regset.mem_add_same.
- apply mem_add_tail. apply IHl with r2. auto.
+ tauto.
+ right; unfold add_intf1', add_intf1; simpl.
+ apply mem_add_tail; auto.
+
+ intros. rewrite SetRegMreg.fold_1. apply H with mr2.
+ intuition. left. apply SetRegMreg.elements_1. auto.
Qed.
Lemma all_interf_regs_correct_1:
@@ -271,16 +289,7 @@ Lemma all_interf_regs_correct_1:
Regset.mem r2 (all_interf_regs g) = true.
Proof.
intros. unfold all_interf_regs.
- generalize (SetRegReg.fold_1
- g.(interf_reg_reg)
- (SetRegMreg.fold
- (fun (r1m2 : SetDepRegMreg.elt) (u : Regset.t) =>
- Regset.add (fst r1m2) u) (interf_reg_mreg g) Regset.empty)
- (fun (r1r2 : SetDepRegReg.elt) (u : Regset.t) =>
- Regset.add (fst r1r2) (Regset.add (snd r1r2) u))).
- intros [l [UN [INEQ EQ]]].
- rewrite EQ. apply all_interf_regs_correct_aux_2.
- elim (INEQ (r1, r2)); intros. auto.
+ apply in_setregreg_fold. tauto.
Qed.
Lemma all_interf_regs_correct_2:
@@ -289,22 +298,6 @@ Lemma all_interf_regs_correct_2:
Regset.mem r1 (all_interf_regs g) = true.
Proof.
intros. unfold all_interf_regs.
- generalize (SetRegReg.fold_1
- g.(interf_reg_reg)
- (SetRegMreg.fold
- (fun (r1m2 : SetDepRegMreg.elt) (u : Regset.t) =>
- Regset.add (fst r1m2) u) (interf_reg_mreg g) Regset.empty)
- (fun (r1r2 : SetDepRegReg.elt) (u : Regset.t) =>
- Regset.add (fst r1r2) (Regset.add (snd r1r2) u))).
- intros [l [UN [INEQ EQ]]].
- rewrite EQ. apply all_interf_regs_correct_aux_1.
- generalize (SetRegMreg.fold_1
- g.(interf_reg_mreg)
- Regset.empty
- (fun (r1r2 : SetDepRegMreg.elt) (u : Regset.t) =>
- Regset.add (fst r1r2) u)).
- change (PTree.t unit) with Regset.t.
- intros [l' [UN' [INEQ' EQ']]].
- rewrite EQ'. apply all_interf_regs_correct_aux_3 with mr2.
- elim (INEQ' (r1, mr2)); intros. auto.
+ apply in_setregreg_fold'. eapply in_setregmreg_fold. eauto.
Qed.
+
diff --git a/backend/Kildall.v b/backend/Kildall.v
index 10b2e1d9..0210b73f 100644
--- a/backend/Kildall.v
+++ b/backend/Kildall.v
@@ -1,6 +1,7 @@
(** Solvers for dataflow inequations. *)
Require Import Coqlib.
+Require Import Iteration.
Require Import Maps.
Require Import Lattice.
@@ -40,128 +41,6 @@ sets [X(n) = top] for all merge points [n], that is, program points having
several predecessors. This solver is useful when least upper bounds of
approximations do not exist or are too expensive to compute. *)
-(** * Bounded iteration *)
-
-(** The three solvers proceed iteratively, increasing the value of one of
- the unknowns [X(n)] at each iteration until a solution is reached.
- This section defines the general form of iteration used. *)
-
-Section BOUNDED_ITERATION.
-
-Variables A B: Set.
-Variable step: A -> B + A.
-
-(** The [step] parameter represents one step of the iteration. From a
- current iteration state [a: A], it either returns a value of type [B],
- meaning that iteration is over and that this [B] value is the final
- result of the iteration, or a value [a' : A] which is the next state
- of the iteration.
-
- The naive way to define the iteration is:
-<<
-Fixpoint iterate (a: A) : B :=
- match step a with
- | inl b => b
- | inr a' => iterate a'
- end.
->>
- However, this is a general recursion, not guaranteed to terminate,
- and therefore not expressible in Coq. The standard way to work around
- this difficulty is to use Noetherian recursion (Coq module [Wf]).
- This requires that we equip the type [A] with a well-founded ordering [<]
- (no infinite ascending chains) and we demand that [step] satisfies
- [step a = inr a' -> a < a']. For the types [A] that are of interest to us
- in this development, it is however very painful to define adequate
- well-founded orderings, even though we know our iterations always
- terminate.
-
- Instead, we choose to bound the number of iterations by an arbitrary
- constant. [iterate] then becomes a function that can fail,
- of type [A -> option B]. The [None] result denotes failure to reach
- a result in the number of iterations prescribed, or, in other terms,
- failure to find a solution to the dataflow problem. The compiler
- passes that exploit dataflow analysis (the [Constprop], [CSE] and
- [Allocation] passes) will, in this case, either fail ([Allocation])
- or turn off the optimization pass ([Constprop] and [CSE]).
-
- Since we know (informally) that our computations terminate, we can
- take a very large constant as the maximal number of iterations.
- Failure will therefore never happen in practice, but of
- course our proofs also cover the failure case and show that
- nothing bad happens in this hypothetical case either. *)
-
-Definition num_iterations := 1000000000000%positive.
-
-(** The simple definition of bounded iteration is:
-<<
-Fixpoint iterate (niter: nat) (a: A) {struct niter} : option B :=
- match niter with
- | O => None
- | S niter' =>
- match step a with
- | inl b => b
- | inr a' => iterate niter' a'
- end
- end.
->>
- This function is structural recursive over the parameter [niter]
- (number of iterations), represented here as a Peano integer (type [nat]).
- However, we want to use very large values of [niter]. As Peano integers,
- these values would be much too large to fit in memory. Therefore,
- we must express iteration counts as a binary integer (type [positive]).
- However, Peano induction over type [positive] is not structural recursion,
- so we cannot define [iterate] as a Coq fixpoint and must use
- Noetherian recursion instead. *)
-
-Definition iter_step (x: positive)
- (next: forall y, Plt y x -> A -> option B)
- (s: A) : option B :=
- match peq x xH with
- | left EQ => None
- | right NOTEQ =>
- match step s with
- | inl res => Some res
- | inr s' => next (Ppred x) (Ppred_Plt x NOTEQ) s'
- end
- end.
-
-Definition iterate: positive -> A -> option B :=
- Fix Plt_wf (fun _ => A -> option B) iter_step.
-
-(** We then prove the expected unrolling equations for [iterate]. *)
-
-Remark unroll_iterate:
- forall x, iterate x = iter_step x (fun y _ => iterate y).
-Proof.
- unfold iterate; apply (Fix_eq Plt_wf (fun _ => A -> option B) iter_step).
- intros. unfold iter_step. apply extensionality. intro s.
- case (peq x xH); intro. auto.
- rewrite H. auto.
-Qed.
-
-Lemma iterate_base:
- forall s, iterate 1%positive s = None.
-Proof.
- intro; rewrite unroll_iterate; unfold iter_step.
- case (peq 1 1); congruence.
-Qed.
-
-Lemma iterate_step:
- forall x s,
- iterate (Psucc x) s =
- match step s with
- | inl res => Some res
- | inr s' => iterate x s'
- end.
-Proof.
- intro; rewrite unroll_iterate; unfold iter_step; intros.
- case (peq (Psucc x) 1); intro.
- destruct x; simpl in e; discriminate.
- rewrite Ppred_succ. auto.
-Qed.
-
-End BOUNDED_ITERATION.
-
(** * Solving forward dataflow problems using Kildall's algorithm *)
(** A forward dataflow solver has the following generic interface.
@@ -316,7 +195,7 @@ Definition step (s: state) : PMap.t L.t + state :=
the start state. *)
Definition fixpoint : option (PMap.t L.t) :=
- iterate _ _ step num_iterations start_state.
+ PrimIter.iterate _ _ step start_state.
(** ** Monotonicity properties *)
@@ -361,33 +240,24 @@ Proof.
apply propagate_succ_incr. auto.
Qed.
-Lemma iterate_incr:
- forall n st res,
- iterate _ _ step n st = Some res ->
- in_incr st.(st_in) res.
-Proof.
- intro n; pattern n. apply positive_Peano_ind; intros until res.
- rewrite iterate_base. congruence.
- rewrite iterate_step. unfold step.
- destruct st.(st_wrk); intros.
- injection H0; intro; subst res.
- red; intros; apply L.ge_refl.
- apply in_incr_trans with
- (propagate_succ_list (mkstate (st_in st) l)
- (transf p (st_in st)!!p)
- (successors p)).(st_in).
- change (st_in st) with (st_in (mkstate (st_in st) l)).
- apply propagate_succ_list_incr.
- apply H. auto.
-Qed.
-
Lemma fixpoint_incr:
forall res,
fixpoint = Some res -> in_incr (start_state_in entrypoints) res.
Proof.
unfold fixpoint; intros.
change (start_state_in entrypoints) with start_state.(st_in).
- apply iterate_incr with num_iterations; auto.
+ eapply (PrimIter.iterate_prop _ _ step
+ (fun st => in_incr start_state.(st_in) st.(st_in))
+ (fun res => in_incr start_state.(st_in) res)).
+
+ intros st INCR. unfold step.
+ destruct st.(st_wrk) as [ | n rem ].
+ auto.
+ apply in_incr_trans with st.(st_in). auto.
+ change st.(st_in) with (mkstate st.(st_in) rem).(st_in).
+ apply propagate_succ_list_incr.
+
+ eauto. apply in_incr_refl.
Qed.
(** ** Correctness invariant *)
@@ -563,30 +433,8 @@ Qed.
(** ** Correctness of the solution returned by [iterate]. *)
(** As a consequence of the [good_state] invariant, the result of
- [iterate], if defined, is a solution of the dataflow inequations,
- since [st_wrk] is empty when [iterate] terminates. *)
-
-Lemma iterate_solution:
- forall niter st res n s,
- good_state st ->
- iterate _ _ step niter st = Some res ->
- Plt n topnode -> In s (successors n) ->
- L.ge res!!s (transf n res!!n).
-Proof.
- intro niter; pattern niter; apply positive_Peano_ind; intros until s.
- rewrite iterate_base. congruence.
- intro GS. rewrite iterate_step.
- unfold step; caseEq (st.(st_wrk)).
- intros. injection H1; intros; subst res.
- elim (GS n H2); intro.
- rewrite H0 in H4. elim H4.
- auto.
- intros. apply H with
- (propagate_succ_list (mkstate st.(st_in) l)
- (transf p st.(st_in)!!p) (successors p)).
- apply step_state_good; auto.
- auto. auto. auto.
-Qed.
+ [fixpoint], if defined, is a solution of the dataflow inequations,
+ since [st_wrk] is empty when the iteration terminates. *)
Theorem fixpoint_solution:
forall res n s,
@@ -594,10 +442,20 @@ Theorem fixpoint_solution:
Plt n topnode -> In s (successors n) ->
L.ge res!!s (transf n res!!n).
Proof.
- unfold fixpoint. intros.
- apply iterate_solution with num_iterations start_state.
- apply start_state_good.
- auto. auto. auto.
+ assert (forall res, fixpoint = Some res ->
+ forall n s, Plt n topnode -> In s (successors n) ->
+ L.ge res!!s (transf n res!!n)).
+ unfold fixpoint. intros res PI. pattern res.
+ eapply (PrimIter.iterate_prop _ _ step good_state).
+
+ intros st GOOD. unfold step.
+ caseEq (st.(st_wrk)).
+ intros. elim (GOOD n H0); intro.
+ rewrite H in H2. contradiction.
+ auto.
+ intros n rem WRK. apply step_state_good; auto.
+
+ eauto. apply start_state_good. eauto.
Qed.
(** As a consequence of the monotonicity property, the result of
@@ -936,8 +794,7 @@ Definition basic_block_list (bb: bbmap) : list positive :=
Definition fixpoint : option result :=
let bb := basic_block_map in
- iterate _ _ (step bb) num_iterations
- (mkstate (PMap.init L.top) (basic_block_list bb)).
+ PrimIter.iterate _ _ (step bb) (mkstate (PMap.init L.top) (basic_block_list bb)).
(** ** Properties of predecessors and multiple-predecessors nodes *)
@@ -1121,23 +978,6 @@ Proof.
right. assumption.
Qed.
-Lemma analyze_invariant:
- forall res count st,
- iterate _ _ (step basic_block_map) count st = Some res ->
- state_invariant st ->
- state_invariant (mkstate res nil).
-Proof.
- intros until count. pattern count.
- apply positive_Peano_ind; intros until st.
- rewrite iterate_base. congruence.
- rewrite iterate_step. unfold step at 1. case st; intros r w; simpl.
- case w.
- intros. replace res with r. auto. congruence.
- intros pc wl. case (plt pc topnode); intros.
- eapply H. eauto. apply propagate_successors_invariant; auto.
- eapply H. eauto. eapply discard_top_worklist_invariant; eauto.
-Qed.
-
Lemma initial_state_invariant:
state_invariant (mkstate (PMap.init L.top) (basic_block_list basic_block_map)).
Proof.
@@ -1146,6 +986,27 @@ Proof.
right. intros. repeat rewrite PMap.gi. apply L.top_ge.
Qed.
+Lemma analyze_invariant:
+ forall res,
+ fixpoint = Some res ->
+ state_invariant (mkstate res nil).
+Proof.
+ unfold fixpoint; intros. pattern res.
+ eapply (PrimIter.iterate_prop _ _ (step basic_block_map)
+ state_invariant).
+
+ intros st INV. destruct st as [stin stwrk].
+ unfold step. simpl. caseEq stwrk.
+ intro. congruence.
+
+ intros pc rem WRK. case (plt pc topnode); intro.
+ apply propagate_successors_invariant; auto. congruence.
+ eapply discard_top_worklist_invariant; eauto. congruence.
+
+ eauto. apply initial_state_invariant.
+Qed.
+
+
(** ** Correctness of the returned solution *)
Theorem fixpoint_solution:
@@ -1154,10 +1015,9 @@ Theorem fixpoint_solution:
Plt n topnode -> In s (successors n) ->
L.ge res!!s (transf n res!!n).
Proof.
- unfold fixpoint.
intros.
assert (state_invariant (mkstate res nil)).
- eapply analyze_invariant; eauto. apply initial_state_invariant.
+ eapply analyze_invariant; eauto.
elim H2; simpl; intros.
elim (H4 n H0); intros.
contradiction.
@@ -1169,10 +1029,9 @@ Theorem fixpoint_entry:
fixpoint = Some res ->
res!!entrypoint = L.top.
Proof.
- unfold fixpoint.
intros.
assert (state_invariant (mkstate res nil)).
- eapply analyze_invariant; eauto. apply initial_state_invariant.
+ eapply analyze_invariant; eauto.
elim H0; simpl; intros.
apply H1. unfold basic_block_map, is_basic_block_head.
fold predecessors.
@@ -1201,28 +1060,21 @@ Proof.
auto. apply H0.
Qed.
-Lemma analyze_P:
- forall bb res count st,
- iterate _ _ (step bb) count st = Some res ->
- Pstate st ->
- (forall pc, P res!!pc).
-Proof.
- intros until count; pattern count; apply positive_Peano_ind; intros until st.
- rewrite iterate_base. congruence.
- rewrite iterate_step; unfold step at 1; destruct st.(st_wrk).
- intros. inversion H0. apply H1.
- destruct (plt p topnode).
- intros. eapply H. eauto.
- apply propagate_successors_P. apply Ptransf. apply H1.
- red; intro; simpl. apply H1.
- intros. eauto.
-Qed.
-
Theorem fixpoint_invariant:
forall res pc, fixpoint = Some res -> P res!!pc.
Proof.
- intros. unfold fixpoint in H. eapply analyze_P; eauto.
- red; intro; simpl. rewrite PMap.gi. apply Ptop.
+ unfold fixpoint; intros. pattern res.
+ eapply (PrimIter.iterate_prop _ _ (step basic_block_map) Pstate).
+
+ intros st PS. unfold step. destruct (st.(st_wrk)).
+ apply PS.
+ assert (PS2: Pstate (mkstate st.(st_in) l)).
+ red; intro; simpl. apply PS.
+ case (plt p topnode); intro.
+ apply propagate_successors_P. auto. auto.
+ auto.
+
+ eauto. red; intro; simpl. rewrite PMap.gi. apply Ptop.
Qed.
End Solver.
diff --git a/backend/LTL.v b/backend/LTL.v
index 2c36cba9..f20ba3fc 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -9,6 +9,7 @@ Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Values.
+Require Import Events.
Require Import Mem.
Require Import Globalenvs.
Require Import Op.
@@ -39,6 +40,7 @@ Inductive block: Set :=
| Bload: memory_chunk -> addressing -> list mreg -> mreg -> block -> block
| Bstore: memory_chunk -> addressing -> list mreg -> mreg -> block -> block
| Bcall: signature -> mreg + ident -> block -> block
+ | Balloc: block -> block
| Bgoto: node -> block
| Bcond: condition -> list mreg -> node -> node -> block
| Breturn: block.
@@ -60,11 +62,19 @@ Record function: Set := mkfunction {
forall (pc: node), Plt pc (Psucc fn_entrypoint) \/ fn_code!pc = None
}.
-Definition program := AST.program function.
+Definition fundef := AST.fundef function.
+
+Definition program := AST.program fundef.
+
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => f.(fn_sig)
+ | External ef => ef.(ef_sig)
+ end.
(** * Operational semantics *)
-Definition genv := Genv.t function.
+Definition genv := Genv.t fundef.
Definition locset := Locmap.t.
Section RELSEM.
@@ -117,7 +127,7 @@ Definition return_regs (caller callee: locset) : locset :=
Variable ge: genv.
-Definition find_function (ros: mreg + ident) (rs: locset) : option function :=
+Definition find_function (ros: mreg + ident) (rs: locset) : option fundef :=
match ros with
| inl r => Genv.find_funct ge (rs (R r))
| inr symb =>
@@ -166,104 +176,119 @@ Inductive outcome: Set :=
| Return: outcome.
Inductive exec_instr: val ->
- block -> locset -> mem ->
+ block -> locset -> mem -> trace ->
block -> locset -> mem -> Prop :=
| exec_Bgetstack:
forall sp sl r b rs m,
exec_instr sp (Bgetstack sl r b) rs m
- b (Locmap.set (R r) (rs (S sl)) rs) m
+ E0 b (Locmap.set (R r) (rs (S sl)) rs) m
| exec_Bsetstack:
forall sp r sl b rs m,
exec_instr sp (Bsetstack r sl b) rs m
- b (Locmap.set (S sl) (rs (R r)) rs) m
+ E0 b (Locmap.set (S sl) (rs (R r)) rs) m
| exec_Bop:
forall sp op args res b rs m v,
eval_operation ge sp op (reglist args rs) = Some v ->
exec_instr sp (Bop op args res b) rs m
- b (Locmap.set (R res) v rs) m
+ E0 b (Locmap.set (R res) v rs) m
| exec_Bload:
forall sp chunk addr args dst b rs m a v,
eval_addressing ge sp addr (reglist args rs) = Some a ->
loadv chunk m a = Some v ->
exec_instr sp (Bload chunk addr args dst b) rs m
- b (Locmap.set (R dst) v rs) m
+ E0 b (Locmap.set (R dst) v rs) m
| exec_Bstore:
forall sp chunk addr args src b rs m m' a,
eval_addressing ge sp addr (reglist args rs) = Some a ->
storev chunk m a (rs (R src)) = Some m' ->
exec_instr sp (Bstore chunk addr args src b) rs m
- b rs m'
+ E0 b rs m'
| exec_Bcall:
- forall sp sig ros b rs m f rs' m',
+ forall sp sig ros b rs m t f rs' m',
find_function ros rs = Some f ->
- sig = f.(fn_sig) ->
- exec_function f rs m rs' m' ->
+ sig = funsig f ->
+ exec_function f rs m t rs' m' ->
exec_instr sp (Bcall sig ros b) rs m
- b (return_regs rs rs') m'
+ t b (return_regs rs rs') m'
+ | exec_Balloc:
+ forall sp b rs m sz m' blk,
+ rs (R Conventions.loc_alloc_argument) = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', blk) ->
+ exec_instr sp (Balloc b) rs m E0 b
+ (Locmap.set (R Conventions.loc_alloc_result)
+ (Vptr blk Int.zero) rs) m'
with exec_instrs: val ->
- block -> locset -> mem ->
+ block -> locset -> mem -> trace ->
block -> locset -> mem -> Prop :=
| exec_refl:
forall sp b rs m,
- exec_instrs sp b rs m b rs m
+ exec_instrs sp b rs m E0 b rs m
| exec_one:
- forall sp b1 rs1 m1 b2 rs2 m2,
- exec_instr sp b1 rs1 m1 b2 rs2 m2 ->
- exec_instrs sp b1 rs1 m1 b2 rs2 m2
+ forall sp b1 rs1 m1 t b2 rs2 m2,
+ exec_instr sp b1 rs1 m1 t b2 rs2 m2 ->
+ exec_instrs sp b1 rs1 m1 t b2 rs2 m2
| exec_trans:
- forall sp b1 rs1 m1 b2 rs2 m2 b3 rs3 m3,
- exec_instrs sp b1 rs1 m1 b2 rs2 m2 ->
- exec_instrs sp b2 rs2 m2 b3 rs3 m3 ->
- exec_instrs sp b1 rs1 m1 b3 rs3 m3
+ forall sp b1 rs1 m1 t t1 b2 rs2 m2 t2 b3 rs3 m3,
+ exec_instrs sp b1 rs1 m1 t1 b2 rs2 m2 ->
+ exec_instrs sp b2 rs2 m2 t2 b3 rs3 m3 ->
+ t = t1 ** t2 ->
+ exec_instrs sp b1 rs1 m1 t b3 rs3 m3
with exec_block: val ->
- block -> locset -> mem ->
+ block -> locset -> mem -> trace ->
outcome -> locset -> mem -> Prop :=
| exec_Bgoto:
- forall sp b rs m s rs' m',
- exec_instrs sp b rs m (Bgoto s) rs' m' ->
- exec_block sp b rs m (Cont s) rs' m'
+ forall sp b rs m t s rs' m',
+ exec_instrs sp b rs m t (Bgoto s) rs' m' ->
+ exec_block sp b rs m t (Cont s) rs' m'
| exec_Bcond_true:
- forall sp b rs m cond args ifso ifnot rs' m',
- exec_instrs sp b rs m (Bcond cond args ifso ifnot) rs' m' ->
+ forall sp b rs m t cond args ifso ifnot rs' m',
+ exec_instrs sp b rs m t (Bcond cond args ifso ifnot) rs' m' ->
eval_condition cond (reglist args rs') = Some true ->
- exec_block sp b rs m (Cont ifso) rs' m'
+ exec_block sp b rs m t (Cont ifso) rs' m'
| exec_Bcond_false:
- forall sp b rs m cond args ifso ifnot rs' m',
- exec_instrs sp b rs m (Bcond cond args ifso ifnot) rs' m' ->
+ forall sp b rs m t cond args ifso ifnot rs' m',
+ exec_instrs sp b rs m t (Bcond cond args ifso ifnot) rs' m' ->
eval_condition cond (reglist args rs') = Some false ->
- exec_block sp b rs m (Cont ifnot) rs' m'
+ exec_block sp b rs m t (Cont ifnot) rs' m'
| exec_Breturn:
- forall sp b rs m rs' m',
- exec_instrs sp b rs m Breturn rs' m' ->
- exec_block sp b rs m Return rs' m'
+ forall sp b rs m t rs' m',
+ exec_instrs sp b rs m t Breturn rs' m' ->
+ exec_block sp b rs m t Return rs' m'
with exec_blocks: code -> val ->
- node -> locset -> mem ->
+ node -> locset -> mem -> trace ->
outcome -> locset -> mem -> Prop :=
| exec_blocks_refl:
forall c sp pc rs m,
- exec_blocks c sp pc rs m (Cont pc) rs m
+ exec_blocks c sp pc rs m E0 (Cont pc) rs m
| exec_blocks_one:
- forall c sp pc b m rs out rs' m',
+ forall c sp pc b m rs t out rs' m',
c!pc = Some b ->
- exec_block sp b rs m out rs' m' ->
- exec_blocks c sp pc rs m out rs' m'
+ exec_block sp b rs m t out rs' m' ->
+ exec_blocks c sp pc rs m t out rs' m'
| exec_blocks_trans:
- forall c sp pc1 rs1 m1 pc2 rs2 m2 out rs3 m3,
- exec_blocks c sp pc1 rs1 m1 (Cont pc2) rs2 m2 ->
- exec_blocks c sp pc2 rs2 m2 out rs3 m3 ->
- exec_blocks c sp pc1 rs1 m1 out rs3 m3
+ forall c sp pc1 rs1 m1 t t1 pc2 rs2 m2 t2 out rs3 m3,
+ exec_blocks c sp pc1 rs1 m1 t1 (Cont pc2) rs2 m2 ->
+ exec_blocks c sp pc2 rs2 m2 t2 out rs3 m3 ->
+ t = t1 ** t2 ->
+ exec_blocks c sp pc1 rs1 m1 t out rs3 m3
-with exec_function: function -> locset -> mem ->
+with exec_function: fundef -> locset -> mem -> trace ->
locset -> mem -> Prop :=
- | exec_funct:
- forall f rs m m1 stk rs2 m2,
+ | exec_funct_internal:
+ forall f rs m m1 stk t rs2 m2,
alloc m 0 f.(fn_stacksize) = (m1, stk) ->
exec_blocks f.(fn_code) (Vptr stk Int.zero)
- f.(fn_entrypoint) (call_regs rs) m1 Return rs2 m2 ->
- exec_function f rs m rs2 (free m2 stk).
+ f.(fn_entrypoint) (call_regs rs) m1 t Return rs2 m2 ->
+ exec_function (Internal f) rs m t rs2 (free m2 stk)
+ | exec_funct_external:
+ forall ef args res rs1 rs2 m t,
+ event_match ef args t res ->
+ args = List.map rs1 (Conventions.loc_arguments ef.(ef_sig)) ->
+ rs2 = Locmap.set (R (Conventions.loc_result ef.(ef_sig))) res rs1 ->
+ exec_function (External ef) rs1 m t rs2 m.
End RELSEM.
@@ -272,15 +297,15 @@ End RELSEM.
main function, to be found in the machine register dictated
by the calling conventions. *)
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
exists b, exists f, exists rs, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- f.(fn_sig) = mksignature nil (Some Tint) /\
- exec_function ge f (Locmap.init Vundef) m0 rs m /\
- rs (R (Conventions.loc_result f.(fn_sig))) = r.
+ funsig f = mksignature nil (Some Tint) /\
+ exec_function ge f (Locmap.init Vundef) m0 t rs m /\
+ rs (R (Conventions.loc_result (funsig f))) = r.
(** We remark that the [exec_blocks] relation is stable if
the control-flow graph is extended by adding new basic blocks
@@ -293,9 +318,9 @@ Variable c1 c2: code.
Hypothesis EXT: forall pc, c2!pc = c1!pc \/ c1!pc = None.
Lemma exec_blocks_extends:
- forall sp pc1 rs1 m1 out rs2 m2,
- exec_blocks ge c1 sp pc1 rs1 m1 out rs2 m2 ->
- exec_blocks ge c2 sp pc1 rs1 m1 out rs2 m2.
+ forall sp pc1 rs1 m1 t out rs2 m2,
+ exec_blocks ge c1 sp pc1 rs1 m1 t out rs2 m2 ->
+ exec_blocks ge c2 sp pc1 rs1 m1 t out rs2 m2.
Proof.
induction 1.
apply exec_blocks_refl.
@@ -319,6 +344,7 @@ Fixpoint successors_aux (b: block) : list node :=
| Bload chunk addr args dst b => successors_aux b
| Bstore chunk addr args src b => successors_aux b
| Bcall sig ros b => successors_aux b
+ | Balloc b => successors_aux b
| Bgoto n => n :: nil
| Bcond cond args ifso ifnot => ifso :: ifnot :: nil
| Breturn => nil
@@ -331,8 +357,8 @@ Definition successors (f: function) (pc: node) : list node :=
end.
Lemma successors_aux_invariant:
- forall ge sp b rs m b' rs' m',
- exec_instrs ge sp b rs m b' rs' m' ->
+ forall ge sp b rs m t b' rs' m',
+ exec_instrs ge sp b rs m t b' rs' m' ->
successors_aux b = successors_aux b'.
Proof.
induction 1; simpl; intros.
@@ -342,16 +368,16 @@ Proof.
Qed.
Lemma successors_correct:
- forall ge f sp pc rs m pc' rs' m' b,
+ forall ge f sp pc rs m t pc' rs' m' b,
f.(fn_code)!pc = Some b ->
- exec_block ge sp b rs m (Cont pc') rs' m' ->
+ exec_block ge sp b rs m t (Cont pc') rs' m' ->
In pc' (successors f pc).
Proof.
intros. unfold successors. rewrite H. inversion H0.
- rewrite (successors_aux_invariant _ _ _ _ _ _ _ _ H6); simpl.
+ rewrite (successors_aux_invariant _ _ _ _ _ _ _ _ _ H7); simpl.
tauto.
- rewrite (successors_aux_invariant _ _ _ _ _ _ _ _ H2); simpl.
+ rewrite (successors_aux_invariant _ _ _ _ _ _ _ _ _ H2); simpl.
tauto.
- rewrite (successors_aux_invariant _ _ _ _ _ _ _ _ H2); simpl.
+ rewrite (successors_aux_invariant _ _ _ _ _ _ _ _ _ H2); simpl.
tauto.
Qed.
diff --git a/backend/LTLtyping.v b/backend/LTLtyping.v
index 3f13ac3c..34508140 100644
--- a/backend/LTLtyping.v
+++ b/backend/LTLtyping.v
@@ -73,6 +73,10 @@ Inductive wt_block : block -> Prop :=
match ros with inl r => mreg_type r = Tint | _ => True end ->
wt_block b ->
wt_block (Bcall sig ros b)
+ | wt_Balloc:
+ forall b,
+ wt_block b ->
+ wt_block (Balloc b)
| wt_Bgoto:
forall lbl,
wt_block (Bgoto lbl)
@@ -88,6 +92,14 @@ End WT_BLOCK.
Definition wt_function (f: function) : Prop :=
forall pc b, f.(fn_code)!pc = Some b -> wt_block f b.
+Inductive wt_fundef: fundef -> Prop :=
+ | wt_fundef_external: forall ef,
+ Conventions.sig_external_ok ef.(ef_sig) ->
+ wt_fundef (External ef)
+ | wt_function_internal: forall f,
+ wt_function f ->
+ wt_fundef (Internal f).
+
Definition wt_program (p: program) : Prop :=
- forall i f, In (i, f) (prog_funct p) -> wt_function f.
+ forall i f, In (i, f) (prog_funct p) -> wt_fundef f.
diff --git a/backend/Linear.v b/backend/Linear.v
index f4ed0454..2520f5bf 100644
--- a/backend/Linear.v
+++ b/backend/Linear.v
@@ -11,6 +11,7 @@ Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -37,6 +38,7 @@ Inductive instruction: Set :=
| Lload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lcall: signature -> mreg + ident -> instruction
+ | Lalloc: instruction
| Llabel: label -> instruction
| Lgoto: label -> instruction
| Lcond: condition -> list mreg -> label -> instruction
@@ -50,9 +52,17 @@ Record function: Set := mkfunction {
fn_code: code
}.
-Definition program := AST.program function.
+Definition fundef := AST.fundef function.
-Definition genv := Genv.t function.
+Definition program := AST.program fundef.
+
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => f.(fn_sig)
+ | External ef => ef.(ef_sig)
+ end.
+
+Definition genv := Genv.t fundef.
Definition locset := Locmap.t.
(** * Operational semantics *)
@@ -88,7 +98,7 @@ Section RELSEM.
Variable ge: genv.
-Definition find_function (ros: mreg + ident) (rs: locset) : option function :=
+Definition find_function (ros: mreg + ident) (rs: locset) : option fundef :=
match ros with
| inl r => Genv.find_funct ge (rs (R r))
| inr symb =>
@@ -106,84 +116,99 @@ Definition find_function (ros: mreg + ident) (rs: locset) : option function :=
[ls'] and [m'] are the final location state and memory state. *)
Inductive exec_instr: function -> val ->
- code -> locset -> mem ->
+ code -> locset -> mem -> trace ->
code -> locset -> mem -> Prop :=
| exec_Lgetstack:
forall f sp sl r b rs m,
exec_instr f sp (Lgetstack sl r :: b) rs m
- b (Locmap.set (R r) (rs (S sl)) rs) m
+ E0 b (Locmap.set (R r) (rs (S sl)) rs) m
| exec_Lsetstack:
forall f sp r sl b rs m,
exec_instr f sp (Lsetstack r sl :: b) rs m
- b (Locmap.set (S sl) (rs (R r)) rs) m
+ E0 b (Locmap.set (S sl) (rs (R r)) rs) m
| exec_Lop:
forall f sp op args res b rs m v,
eval_operation ge sp op (reglist args rs) = Some v ->
exec_instr f sp (Lop op args res :: b) rs m
- b (Locmap.set (R res) v rs) m
+ E0 b (Locmap.set (R res) v rs) m
| exec_Lload:
forall f sp chunk addr args dst b rs m a v,
eval_addressing ge sp addr (reglist args rs) = Some a ->
loadv chunk m a = Some v ->
exec_instr f sp (Lload chunk addr args dst :: b) rs m
- b (Locmap.set (R dst) v rs) m
+ E0 b (Locmap.set (R dst) v rs) m
| exec_Lstore:
forall f sp chunk addr args src b rs m m' a,
eval_addressing ge sp addr (reglist args rs) = Some a ->
storev chunk m a (rs (R src)) = Some m' ->
exec_instr f sp (Lstore chunk addr args src :: b) rs m
- b rs m'
+ E0 b rs m'
| exec_Lcall:
- forall f sp sig ros b rs m f' rs' m',
+ forall f sp sig ros b rs m t f' rs' m',
find_function ros rs = Some f' ->
- sig = f'.(fn_sig) ->
- exec_function f' rs m rs' m' ->
+ sig = funsig f' ->
+ exec_function f' rs m t rs' m' ->
exec_instr f sp (Lcall sig ros :: b) rs m
- b (return_regs rs rs') m'
+ t b (return_regs rs rs') m'
+ | exec_Lalloc:
+ forall f sp b rs m sz m' blk,
+ rs (R Conventions.loc_alloc_argument) = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', blk) ->
+ exec_instr f sp (Lalloc :: b) rs m
+ E0 b (Locmap.set (R Conventions.loc_alloc_result)
+ (Vptr blk Int.zero) rs) m'
| exec_Llabel:
forall f sp lbl b rs m,
exec_instr f sp (Llabel lbl :: b) rs m
- b rs m
+ E0 b rs m
| exec_Lgoto:
forall f sp lbl b rs m b',
find_label lbl f.(fn_code) = Some b' ->
exec_instr f sp (Lgoto lbl :: b) rs m
- b' rs m
+ E0 b' rs m
| exec_Lcond_true:
forall f sp cond args lbl b rs m b',
eval_condition cond (reglist args rs) = Some true ->
find_label lbl f.(fn_code) = Some b' ->
exec_instr f sp (Lcond cond args lbl :: b) rs m
- b' rs m
+ E0 b' rs m
| exec_Lcond_false:
forall f sp cond args lbl b rs m,
eval_condition cond (reglist args rs) = Some false ->
exec_instr f sp (Lcond cond args lbl :: b) rs m
- b rs m
+ E0 b rs m
with exec_instrs: function -> val ->
- code -> locset -> mem ->
+ code -> locset -> mem -> trace ->
code -> locset -> mem -> Prop :=
| exec_refl:
forall f sp b rs m,
- exec_instrs f sp b rs m b rs m
+ exec_instrs f sp b rs m E0 b rs m
| exec_one:
- forall f sp b1 rs1 m1 b2 rs2 m2,
- exec_instr f sp b1 rs1 m1 b2 rs2 m2 ->
- exec_instrs f sp b1 rs1 m1 b2 rs2 m2
+ forall f sp b1 rs1 m1 t b2 rs2 m2,
+ exec_instr f sp b1 rs1 m1 t b2 rs2 m2 ->
+ exec_instrs f sp b1 rs1 m1 t b2 rs2 m2
| exec_trans:
- forall f sp b1 rs1 m1 b2 rs2 m2 b3 rs3 m3,
- exec_instrs f sp b1 rs1 m1 b2 rs2 m2 ->
- exec_instrs f sp b2 rs2 m2 b3 rs3 m3 ->
- exec_instrs f sp b1 rs1 m1 b3 rs3 m3
-
-with exec_function: function -> locset -> mem -> locset -> mem -> Prop :=
- | exec_funct:
- forall f rs m m1 stk rs2 m2 b,
+ forall f sp b1 rs1 m1 t1 b2 rs2 m2 t2 b3 rs3 m3 t,
+ exec_instrs f sp b1 rs1 m1 t1 b2 rs2 m2 ->
+ exec_instrs f sp b2 rs2 m2 t2 b3 rs3 m3 ->
+ t = t1 ** t2 ->
+ exec_instrs f sp b1 rs1 m1 t b3 rs3 m3
+
+with exec_function: fundef -> locset -> mem -> trace -> locset -> mem -> Prop :=
+ | exec_funct_internal:
+ forall f rs m t m1 stk rs2 m2 b,
alloc m 0 f.(fn_stacksize) = (m1, stk) ->
exec_instrs f (Vptr stk Int.zero)
- f.(fn_code) (call_regs rs) m1 (Lreturn :: b) rs2 m2 ->
- exec_function f rs m rs2 (free m2 stk).
+ f.(fn_code) (call_regs rs) m1
+ t (Lreturn :: b) rs2 m2 ->
+ exec_function (Internal f) rs m t rs2 (free m2 stk)
+ | exec_funct_external:
+ forall ef args res rs1 rs2 m t,
+ event_match ef args t res ->
+ args = List.map rs1 (Conventions.loc_arguments ef.(ef_sig)) ->
+ rs2 = Locmap.set (R (Conventions.loc_result ef.(ef_sig))) res rs1 ->
+ exec_function (External ef) rs1 m t rs2 m.
Scheme exec_instr_ind3 := Minimality for exec_instr Sort Prop
with exec_instrs_ind3 := Minimality for exec_instrs Sort Prop
@@ -191,13 +216,13 @@ Scheme exec_instr_ind3 := Minimality for exec_instr Sort Prop
End RELSEM.
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
exists b, exists f, exists rs, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- f.(fn_sig) = mksignature nil (Some Tint) /\
- exec_function ge f (Locmap.init Vundef) m0 rs m /\
- rs (R (Conventions.loc_result f.(fn_sig))) = r.
+ funsig f = mksignature nil (Some Tint) /\
+ exec_function ge f (Locmap.init Vundef) m0 t rs m /\
+ rs (R (Conventions.loc_result (funsig f))) = r.
diff --git a/backend/Linearize.v b/backend/Linearize.v
index af70b0fd..f5b2a9e2 100644
--- a/backend/Linearize.v
+++ b/backend/Linearize.v
@@ -148,6 +148,8 @@ Fixpoint linearize_block (b: block) (k: code) {struct b} : code :=
Lstore chunk addr args src :: linearize_block b k
| Bcall sig ros b =>
Lcall sig ros :: linearize_block b k
+ | Balloc b =>
+ Lalloc :: linearize_block b k
| Bgoto s =>
Lgoto s :: k
| Bcond cond args s1 s2 =>
@@ -208,5 +210,8 @@ Definition cleanup_function (f: Linear.function) : Linear.function :=
Definition transf_function (f: LTL.function) : Linear.function :=
cleanup_function (linearize_function f).
+Definition transf_fundef (f: LTL.fundef) : Linear.fundef :=
+ AST.transf_fundef transf_function f.
+
Definition transf_program (p: LTL.program) : Linear.program :=
- transform_program transf_function p.
+ transform_program transf_fundef p.
diff --git a/backend/Linearizeproof.v b/backend/Linearizeproof.v
index b80acb4d..22bf19c0 100644
--- a/backend/Linearizeproof.v
+++ b/backend/Linearizeproof.v
@@ -6,6 +6,7 @@ Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -25,19 +26,25 @@ Let tge := Genv.globalenv tprog.
Lemma functions_translated:
forall v f,
Genv.find_funct ge v = Some f ->
- Genv.find_funct tge v = Some (transf_function f).
-Proof (@Genv.find_funct_transf _ _ transf_function prog).
+ Genv.find_funct tge v = Some (transf_fundef f).
+Proof (@Genv.find_funct_transf _ _ transf_fundef prog).
Lemma function_ptr_translated:
forall v f,
Genv.find_funct_ptr ge v = Some f ->
- Genv.find_funct_ptr tge v = Some (transf_function f).
-Proof (@Genv.find_funct_ptr_transf _ _ transf_function prog).
+ Genv.find_funct_ptr tge v = Some (transf_fundef f).
+Proof (@Genv.find_funct_ptr_transf _ _ transf_fundef prog).
Lemma symbols_preserved:
forall id,
Genv.find_symbol tge id = Genv.find_symbol ge id.
-Proof (@Genv.find_symbol_transf _ _ transf_function prog).
+Proof (@Genv.find_symbol_transf _ _ transf_fundef prog).
+
+Lemma sig_preserved:
+ forall f, funsig (transf_fundef f) = LTL.funsig f.
+Proof.
+ destruct f; reflexivity.
+Qed.
(** * Correctness of reachability analysis *)
@@ -80,9 +87,9 @@ Qed.
[pc] is reachable, then [pc'] is reachable. *)
Lemma reachable_correct_1:
- forall f sp pc rs m pc' rs' m' b,
+ forall f sp pc rs m t pc' rs' m' b,
f.(LTL.fn_code)!pc = Some b ->
- exec_block ge sp b rs m (Cont pc') rs' m' ->
+ exec_block ge sp b rs m t (Cont pc') rs' m' ->
(reachable f)!!pc = true ->
(reachable f)!!pc' = true.
Proof.
@@ -92,8 +99,8 @@ Proof.
Qed.
Lemma reachable_correct_2:
- forall c sp pc rs m out rs' m',
- exec_blocks ge c sp pc rs m out rs' m' ->
+ forall c sp pc rs m t out rs' m',
+ exec_blocks ge c sp pc rs m t out rs' m' ->
forall f pc',
c = f.(LTL.fn_code) ->
out = Cont pc' ->
@@ -236,18 +243,19 @@ Lemma starts_with_correct:
starts_with lbl c1 = true ->
find_label lbl c2 = Some c3 ->
exec_instrs tge f sp (cleanup_code c1) ls m
- (cleanup_code c3) ls m.
+ E0 (cleanup_code c3) ls m.
Proof.
induction c1.
simpl; intros; discriminate.
simpl starts_with. destruct a; try (intros; discriminate).
- intros. apply exec_trans with (cleanup_code c1) ls m.
+ intros. apply exec_trans with E0 (cleanup_code c1) ls m E0.
simpl.
constructor. constructor.
destruct (peq lbl l).
subst l. replace c3 with c1. constructor.
apply find_label_unique with lbl c2; auto.
apply IHc1 with c2; auto. eapply is_tail_cons_left; eauto.
+ traceEq.
Qed.
(** Code cleanup preserves the labeling of the code. *)
@@ -273,13 +281,13 @@ Qed.
or one transitions in the cleaned-up code. *)
Lemma cleanup_code_correct_1:
- forall f sp c1 ls1 m1 c2 ls2 m2,
- exec_instr tge f sp c1 ls1 m1 c2 ls2 m2 ->
+ forall f sp c1 ls1 m1 t c2 ls2 m2,
+ exec_instr tge f sp c1 ls1 m1 t c2 ls2 m2 ->
is_tail c1 f.(fn_code) ->
unique_labels f.(fn_code) ->
exec_instrs tge (cleanup_function f) sp
(cleanup_code c1) ls1 m1
- (cleanup_code c2) ls2 m2.
+ t (cleanup_code c2) ls2 m2.
Proof.
induction 1; simpl; intros;
try (apply exec_one; econstructor; eauto; fail).
@@ -310,8 +318,8 @@ Proof.
Qed.
Lemma is_tail_exec_instr:
- forall f sp c1 ls1 m1 c2 ls2 m2,
- exec_instr tge f sp c1 ls1 m1 c2 ls2 m2 ->
+ forall f sp c1 ls1 m1 t c2 ls2 m2,
+ exec_instr tge f sp c1 ls1 m1 t c2 ls2 m2 ->
is_tail c1 f.(fn_code) -> is_tail c2 f.(fn_code).
Proof.
induction 1; intros;
@@ -321,8 +329,8 @@ Proof.
Qed.
Lemma is_tail_exec_instrs:
- forall f sp c1 ls1 m1 c2 ls2 m2,
- exec_instrs tge f sp c1 ls1 m1 c2 ls2 m2 ->
+ forall f sp c1 ls1 m1 t c2 ls2 m2,
+ exec_instrs tge f sp c1 ls1 m1 t c2 ls2 m2 ->
is_tail c1 f.(fn_code) -> is_tail c2 f.(fn_code).
Proof.
induction 1; intros.
@@ -335,31 +343,32 @@ Qed.
to zero, one or several transitions in the cleaned-up code. *)
Lemma cleanup_code_correct_2:
- forall f sp c1 ls1 m1 c2 ls2 m2,
- exec_instrs tge f sp c1 ls1 m1 c2 ls2 m2 ->
+ forall f sp c1 ls1 m1 t c2 ls2 m2,
+ exec_instrs tge f sp c1 ls1 m1 t c2 ls2 m2 ->
is_tail c1 f.(fn_code) ->
unique_labels f.(fn_code) ->
exec_instrs tge (cleanup_function f) sp
(cleanup_code c1) ls1 m1
- (cleanup_code c2) ls2 m2.
+ t (cleanup_code c2) ls2 m2.
Proof.
induction 1; simpl; intros.
apply exec_refl.
eapply cleanup_code_correct_1; eauto.
- apply exec_trans with (cleanup_code b2) rs2 m2.
+ apply exec_trans with t1 (cleanup_code b2) rs2 m2 t2.
auto.
apply IHexec_instrs2; auto.
eapply is_tail_exec_instrs; eauto.
+ auto.
Qed.
Lemma cleanup_function_correct:
- forall f ls1 m1 ls2 m2,
- exec_function tge f ls1 m1 ls2 m2 ->
+ forall f ls1 m1 t ls2 m2,
+ exec_function tge (Internal f) ls1 m1 t ls2 m2 ->
unique_labels f.(fn_code) ->
- exec_function tge (cleanup_function f) ls1 m1 ls2 m2.
+ exec_function tge (Internal (cleanup_function f)) ls1 m1 t ls2 m2.
Proof.
- induction 1; intro.
- generalize (cleanup_code_correct_2 _ _ _ _ _ _ _ _ H0 (is_tail_refl _) H1).
+ intros. inversion H; subst.
+ generalize (cleanup_code_correct_2 _ _ _ _ _ _ _ _ _ H3 (is_tail_refl _) H0).
simpl. intro.
econstructor; eauto.
Qed.
@@ -479,8 +488,8 @@ Definition valid_outcome (f: LTL.function) (out: outcome) :=
(** Auxiliary lemma used to establish the [valid_outcome] property. *)
Lemma exec_blocks_valid_outcome:
- forall c sp pc ls1 m1 out ls2 m2,
- exec_blocks ge c sp pc ls1 m1 out ls2 m2 ->
+ forall c sp pc ls1 m1 t out ls2 m2,
+ exec_blocks ge c sp pc ls1 m1 t out ls2 m2 ->
forall f,
c = f.(LTL.fn_code) ->
(reachable f)!!pc = true ->
@@ -514,34 +523,34 @@ Inductive cont_for_outcome: LTL.function -> outcome -> code -> Prop :=
Definition exec_instr_prop
(sp: val) (b1: block) (ls1: locset) (m1: mem)
- (b2: block) (ls2: locset) (m2: mem) : Prop :=
+ (t: trace) (b2: block) (ls2: locset) (m2: mem) : Prop :=
forall f k,
exec_instr tge (linearize_function f) sp
(linearize_block b1 k) ls1 m1
- (linearize_block b2 k) ls2 m2.
+ t (linearize_block b2 k) ls2 m2.
Definition exec_instrs_prop
(sp: val) (b1: block) (ls1: locset) (m1: mem)
- (b2: block) (ls2: locset) (m2: mem) : Prop :=
+ (t: trace) (b2: block) (ls2: locset) (m2: mem) : Prop :=
forall f k,
exec_instrs tge (linearize_function f) sp
(linearize_block b1 k) ls1 m1
- (linearize_block b2 k) ls2 m2.
+ t (linearize_block b2 k) ls2 m2.
Definition exec_block_prop
(sp: val) (b: block) (ls1: locset) (m1: mem)
- (out: outcome) (ls2: locset) (m2: mem): Prop :=
+ (t: trace) (out: outcome) (ls2: locset) (m2: mem): Prop :=
forall f k,
valid_outcome f out ->
exists k',
exec_instrs tge (linearize_function f) sp
(linearize_block b k) ls1 m1
- k' ls2 m2
+ t k' ls2 m2
/\ cont_for_outcome f out k'.
Definition exec_blocks_prop
(c: LTL.code) (sp: val) (pc: node) (ls1: locset) (m1: mem)
- (out: outcome) (ls2: locset) (m2: mem): Prop :=
+ (t: trace) (out: outcome) (ls2: locset) (m2: mem): Prop :=
forall f k,
c = f.(LTL.fn_code) ->
(reachable f)!!pc = true ->
@@ -550,12 +559,13 @@ Definition exec_blocks_prop
exists k',
exec_instrs tge (linearize_function f) sp
k ls1 m1
- k' ls2 m2
+ t k' ls2 m2
/\ cont_for_outcome f out k'.
Definition exec_function_prop
- (f: LTL.function) (ls1: locset) (m1: mem) (ls2: locset) (m2: mem): Prop :=
- exec_function tge (transf_function f) ls1 m1 ls2 m2.
+ (f: LTL.fundef) (ls1: locset) (m1: mem) (t: trace)
+ (ls2: locset) (m2: mem): Prop :=
+ exec_function tge (transf_fundef f) ls1 m1 t ls2 m2.
Scheme exec_instr_ind5 := Minimality for LTL.exec_instr Sort Prop
with exec_instrs_ind5 := Minimality for LTL.exec_instrs Sort Prop
@@ -567,9 +577,9 @@ Scheme exec_instr_ind5 := Minimality for LTL.exec_instr Sort Prop
derivation. *)
Lemma transf_function_correct:
- forall f ls1 m1 ls2 m2,
- LTL.exec_function ge f ls1 m1 ls2 m2 ->
- exec_function_prop f ls1 m1 ls2 m2.
+ forall f ls1 m1 t ls2 m2,
+ LTL.exec_function ge f ls1 m1 t ls2 m2 ->
+ exec_function_prop f ls1 m1 t ls2 m2.
Proof.
apply (exec_function_ind5 ge
exec_instr_prop
@@ -596,27 +606,29 @@ Proof.
exact symbols_preserved.
auto.
(* call *)
- apply exec_Lcall with (transf_function f).
+ apply exec_Lcall with (transf_fundef f).
generalize H. destruct ros; simpl.
intro; apply functions_translated; auto.
rewrite symbols_preserved. destruct (Genv.find_symbol ge i).
intro; apply function_ptr_translated; auto. congruence.
- rewrite H0; reflexivity.
+ generalize (sig_preserved f). congruence.
apply H2.
+ (* alloc *)
+ eapply exec_Lalloc; eauto.
(* instr_refl *)
apply exec_refl.
(* instr_one *)
apply exec_one. apply H0.
(* instr_trans *)
- apply exec_trans with (linearize_block b2 k) rs2 m2.
- apply H0. apply H2.
+ apply exec_trans with t1 (linearize_block b2 k) rs2 m2 t2.
+ apply H0. apply H2. auto.
(* goto *)
elim H1. intros REACH [b2 AT2].
generalize (H0 f k). simpl. intro.
elim (find_label_lin f s b2 AT2 REACH). intros k2 FIND.
exists (linearize_block b2 k2).
split.
- eapply exec_trans. eexact H2. constructor. constructor. auto.
+ eapply exec_trans. eexact H2. constructor. constructor. auto. traceEq.
constructor. auto.
(* cond, true *)
elim H2. intros REACH [b2 AT2].
@@ -628,10 +640,10 @@ Proof.
eapply exec_trans. eexact H3.
eapply exec_trans. apply exec_one. apply exec_Lcond_false.
change false with (negb true). apply eval_negate_condition. auto.
- apply exec_one. apply exec_Lgoto. auto.
+ apply exec_one. apply exec_Lgoto. auto. reflexivity. traceEq.
eapply exec_trans. eexact H3.
apply exec_one. apply exec_Lcond_true.
- auto. auto.
+ auto. auto. traceEq.
constructor; auto.
(* cond, false *)
elim H2. intros REACH [b2 AT2].
@@ -643,10 +655,10 @@ Proof.
eapply exec_trans. eexact H3.
apply exec_one. apply exec_Lcond_true.
change true with (negb false). apply eval_negate_condition. auto.
- auto.
+ auto. traceEq.
eapply exec_trans. eexact H3.
eapply exec_trans. apply exec_one. apply exec_Lcond_false. auto.
- apply exec_one. apply exec_Lgoto. auto.
+ apply exec_one. apply exec_Lgoto. auto. reflexivity. traceEq.
constructor; auto.
(* return *)
exists (Lreturn :: k). split.
@@ -663,11 +675,11 @@ Proof.
eapply reachable_correct_2. eexact H. auto. auto. auto.
assert (valid_outcome f (Cont pc2)).
eapply exec_blocks_valid_outcome; eauto.
- generalize (H0 f k H3 H4 H5 H8). intros [k1 [EX1 CFO2]].
+ generalize (H0 f k H4 H5 H6 H9). intros [k1 [EX1 CFO2]].
inversion CFO2.
- generalize (H2 f k1 H3 H7 H11 H6). intros [k2 [EX2 CFO3]].
+ generalize (H2 f k1 H4 H8 H12 H7). intros [k2 [EX2 CFO3]].
exists k2. split. eapply exec_trans; eauto. auto.
- (* function -- TA-DA! *)
+ (* internal function -- TA-DA! *)
assert (REACH0: (reachable f)!!(fn_entrypoint f) = true).
apply reachable_entrypoint.
assert (VO2: valid_outcome f Return). simpl; auto.
@@ -687,25 +699,27 @@ Proof.
inversion CO. subst k'.
unfold transf_function. apply cleanup_function_correct.
econstructor. eauto. rewrite EQ. eapply exec_trans.
- apply exec_one. constructor. eauto.
+ apply exec_one. constructor. eauto. traceEq.
apply unique_labels_lin_function.
+ (* external function *)
+ econstructor; eauto.
Qed.
End LINEARIZATION.
Theorem transf_program_correct:
- forall (p: LTL.program) (r: val),
- LTL.exec_program p r ->
- Linear.exec_program (transf_program p) r.
+ forall (p: LTL.program) (t: trace) (r: val),
+ LTL.exec_program p t r ->
+ Linear.exec_program (transf_program p) t r.
Proof.
- intros p r [b [f [ls [m [FIND1 [FIND2 [SIG [EX RES]]]]]]]].
- red. exists b; exists (transf_function f); exists ls; exists m.
+ intros p t r [b [f [ls [m [FIND1 [FIND2 [SIG [EX RES]]]]]]]].
+ red. exists b; exists (transf_fundef f); exists ls; exists m.
split. change (prog_main (transf_program p)) with (prog_main p).
rewrite symbols_preserved; auto.
split. apply function_ptr_translated. auto.
- split. auto.
+ split. generalize (sig_preserved f); congruence.
split. apply transf_function_correct.
unfold transf_program. rewrite Genv.init_mem_transf. auto.
- exact RES.
+ rewrite sig_preserved. exact RES.
Qed.
diff --git a/backend/Linearizetyping.v b/backend/Linearizetyping.v
index 6cebca8d..66926e9a 100644
--- a/backend/Linearizetyping.v
+++ b/backend/Linearizetyping.v
@@ -274,6 +274,8 @@ Proof.
(* call *)
constructor; auto.
eapply size_arguments_bound; eauto.
+ (* alloc *)
+ constructor.
(* goto *)
constructor.
(* cond *)
@@ -327,14 +329,24 @@ Proof.
apply cleanup_function_conservation_2; auto.
Qed.
+Lemma wt_transf_fundef:
+ forall f,
+ LTLtyping.wt_fundef f ->
+ wt_fundef (transf_fundef f).
+Proof.
+ induction 1; simpl.
+ constructor; assumption.
+ constructor; apply wt_transf_function; assumption.
+Qed.
+
Lemma program_typing_preserved:
forall (p: LTL.program),
LTLtyping.wt_program p ->
Lineartyping.wt_program (transf_program p).
Proof.
intros; red; intros.
- generalize (transform_program_function transf_function p i f H0).
+ generalize (transform_program_function transf_fundef p i f H0).
intros [f0 [IN TR]].
- subst f. apply wt_transf_function; auto.
+ subst f. apply wt_transf_fundef; auto.
apply (H i f0 IN).
Qed.
diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v
index 0b13b40a..bf41908b 100644
--- a/backend/Lineartyping.v
+++ b/backend/Lineartyping.v
@@ -38,6 +38,7 @@ Definition regs_of_instr (i: instruction) : list mreg :=
| Lstore chunk addr args src => src :: args
| Lcall sig (inl fn) => fn :: nil
| Lcall sig (inr _) => nil
+ | Lalloc => nil
| Llabel lbl => nil
| Lgoto lbl => nil
| Lcond cond args lbl => args
@@ -231,6 +232,8 @@ Inductive wt_instr : instruction -> Prop :=
size_arguments sig <= bound_outgoing b ->
match ros with inl r => mreg_type r = Tint | _ => True end ->
wt_instr (Lcall sig ros)
+ | wt_Lalloc:
+ wt_instr Lalloc
| wt_Llabel:
forall lbl,
wt_instr (Llabel lbl)
@@ -249,6 +252,14 @@ End WT_INSTR.
Definition wt_function (f: function) : Prop :=
forall instr, In instr f.(fn_code) -> wt_instr f instr.
+Inductive wt_fundef: fundef -> Prop :=
+ | wt_fundef_external: forall ef,
+ Conventions.sig_external_ok ef.(ef_sig) ->
+ wt_fundef (External ef)
+ | wt_function_internal: forall f,
+ wt_function f ->
+ wt_fundef (Internal f).
+
Definition wt_program (p: program) : Prop :=
- forall i f, In (i, f) (prog_funct p) -> wt_function f.
+ forall i f, In (i, f) (prog_funct p) -> wt_fundef f.
diff --git a/backend/Mach.v b/backend/Mach.v
index f9537985..1a9a94ae 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -10,9 +10,11 @@ Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
+Require Conventions.
(** * Abstract syntax *)
@@ -43,6 +45,7 @@ Inductive instruction: Set :=
| Mload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mcall: signature -> mreg + ident -> instruction
+ | Malloc: instruction
| Mlabel: label -> instruction
| Mgoto: label -> instruction
| Mcond: condition -> list mreg -> label -> instruction
@@ -56,9 +59,17 @@ Record function: Set := mkfunction
fn_stacksize: Z;
fn_framesize: Z }.
-Definition program := AST.program function.
+Definition fundef := AST.fundef function.
-Definition genv := Genv.t function.
+Definition program := AST.program fundef.
+
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => f.(fn_sig)
+ | External ef => ef.(ef_sig)
+ end.
+
+Definition genv := Genv.t fundef.
(** * Dynamic semantics *)
@@ -122,7 +133,7 @@ Section RELSEM.
Variable ge: genv.
-Definition find_function (ros: mreg + ident) (rs: regset) : option function :=
+Definition find_function (ros: mreg + ident) (rs: regset) : option fundef :=
match ros with
| inl r => Genv.find_funct ge (rs r)
| inr symb =>
@@ -141,95 +152,104 @@ Definition find_function (ros: mreg + ident) (rs: regset) : option function :=
Inductive exec_instr:
function -> val ->
- code -> regset -> mem ->
+ code -> regset -> mem -> trace ->
code -> regset -> mem -> Prop :=
| exec_Mlabel:
forall f sp lbl c rs m,
exec_instr f sp
(Mlabel lbl :: c) rs m
- c rs m
+ E0 c rs m
| exec_Mgetstack:
forall f sp ofs ty dst c rs m v,
load_stack m sp ty ofs = Some v ->
exec_instr f sp
(Mgetstack ofs ty dst :: c) rs m
- c (rs#dst <- v) m
+ E0 c (rs#dst <- v) m
| exec_Msetstack:
forall f sp src ofs ty c rs m m',
store_stack m sp ty ofs (rs src) = Some m' ->
exec_instr f sp
(Msetstack src ofs ty :: c) rs m
- c rs m'
+ E0 c rs m'
| exec_Mgetparam:
forall f sp parent ofs ty dst c rs m v,
load_stack m sp Tint (Int.repr 0) = Some parent ->
load_stack m parent ty ofs = Some v ->
exec_instr f sp
(Mgetparam ofs ty dst :: c) rs m
- c (rs#dst <- v) m
+ E0 c (rs#dst <- v) m
| exec_Mop:
forall f sp op args res c rs m v,
eval_operation ge sp op rs##args = Some v ->
exec_instr f sp
(Mop op args res :: c) rs m
- c (rs#res <- v) m
+ E0 c (rs#res <- v) m
| exec_Mload:
forall f sp chunk addr args dst c rs m a v,
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
exec_instr f sp
(Mload chunk addr args dst :: c) rs m
- c (rs#dst <- v) m
+ E0 c (rs#dst <- v) m
| exec_Mstore:
forall f sp chunk addr args src c rs m m' a,
eval_addressing ge sp addr rs##args = Some a ->
Mem.storev chunk m a (rs src) = Some m' ->
exec_instr f sp
(Mstore chunk addr args src :: c) rs m
- c rs m'
+ E0 c rs m'
| exec_Mcall:
- forall f sp sig ros c rs m f' rs' m',
+ forall f sp sig ros c rs m f' t rs' m',
find_function ros rs = Some f' ->
- exec_function f' sp rs m rs' m' ->
+ exec_function f' sp rs m t rs' m' ->
exec_instr f sp
(Mcall sig ros :: c) rs m
- c rs' m'
+ t c rs' m'
+ | exec_Malloc:
+ forall f sp c rs m sz m' blk,
+ rs (Conventions.loc_alloc_argument) = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', blk) ->
+ exec_instr f sp
+ (Malloc :: c) rs m
+ E0 c (rs#Conventions.loc_alloc_result <-
+ (Vptr blk Int.zero)) m'
| exec_Mgoto:
forall f sp lbl c rs m c',
find_label lbl f.(fn_code) = Some c' ->
exec_instr f sp
(Mgoto lbl :: c) rs m
- c' rs m
+ E0 c' rs m
| exec_Mcond_true:
forall f sp cond args lbl c rs m c',
eval_condition cond rs##args = Some true ->
find_label lbl f.(fn_code) = Some c' ->
exec_instr f sp
(Mcond cond args lbl :: c) rs m
- c' rs m
+ E0 c' rs m
| exec_Mcond_false:
forall f sp cond args lbl c rs m,
eval_condition cond rs##args = Some false ->
exec_instr f sp
(Mcond cond args lbl :: c) rs m
- c rs m
+ E0 c rs m
with exec_instrs:
function -> val ->
- code -> regset -> mem ->
+ code -> regset -> mem -> trace ->
code -> regset -> mem -> Prop :=
| exec_refl:
forall f sp c rs m,
- exec_instrs f sp c rs m c rs m
+ exec_instrs f sp c rs m E0 c rs m
| exec_one:
- forall f sp c rs m c' rs' m',
- exec_instr f sp c rs m c' rs' m' ->
- exec_instrs f sp c rs m c' rs' m'
+ forall f sp c rs m t c' rs' m',
+ exec_instr f sp c rs m t c' rs' m' ->
+ exec_instrs f sp c rs m t c' rs' m'
| exec_trans:
- forall f sp c1 rs1 m1 c2 rs2 m2 c3 rs3 m3,
- exec_instrs f sp c1 rs1 m1 c2 rs2 m2 ->
- exec_instrs f sp c2 rs2 m2 c3 rs3 m3 ->
- exec_instrs f sp c1 rs1 m1 c3 rs3 m3
+ forall f sp c1 rs1 m1 t1 c2 rs2 m2 t2 c3 rs3 m3 t3,
+ exec_instrs f sp c1 rs1 m1 t1 c2 rs2 m2 ->
+ exec_instrs f sp c2 rs2 m2 t2 c3 rs3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_instrs f sp c1 rs1 m1 t3 c3 rs3 m3
(** In addition to reserving the word at offset 0 in the activation
record for maintaining the linking of activation records,
@@ -252,9 +272,9 @@ with exec_instrs:
with exec_function_body:
function -> val -> val ->
- regset -> mem -> regset -> mem -> Prop :=
+ regset -> mem -> trace -> regset -> mem -> Prop :=
| exec_funct_body:
- forall f parent ra rs m rs' m1 m2 m3 m4 stk c,
+ forall f parent ra rs m t rs' m1 m2 m3 m4 stk c,
Mem.alloc m (- f.(fn_framesize))
(align_16_top (- f.(fn_framesize)) f.(fn_stacksize))
= (m1, stk) ->
@@ -263,19 +283,25 @@ with exec_function_body:
store_stack m2 sp Tint (Int.repr 4) ra = Some m3 ->
exec_instrs f sp
f.(fn_code) rs m3
- (Mreturn :: c) rs' m4 ->
+ t (Mreturn :: c) rs' m4 ->
load_stack m4 sp Tint (Int.repr 0) = Some parent ->
load_stack m4 sp Tint (Int.repr 4) = Some ra ->
- exec_function_body f parent ra rs m rs' (Mem.free m4 stk)
+ exec_function_body f parent ra rs m t rs' (Mem.free m4 stk)
with exec_function:
- function -> val -> regset -> mem -> regset -> mem -> Prop :=
- | exec_funct:
- forall f parent rs m rs' m',
+ fundef -> val -> regset -> mem -> trace -> regset -> mem -> Prop :=
+ | exec_funct_internal:
+ forall f parent rs m t rs' m',
(forall ra,
Val.has_type ra Tint ->
- exec_function_body f parent ra rs m rs' m') ->
- exec_function f parent rs m rs' m'.
+ exec_function_body f parent ra rs m t rs' m') ->
+ exec_function (Internal f) parent rs m t rs' m'
+ | exec_funct_external:
+ forall ef parent args res rs1 rs2 m t,
+ event_match ef args t res ->
+ args = rs1##(Conventions.loc_external_arguments ef.(ef_sig)) ->
+ rs2 = (rs1#(Conventions.loc_result ef.(ef_sig)) <- res) ->
+ exec_function (External ef) parent rs1 m t rs2 m.
Scheme exec_instr_ind4 := Minimality for exec_instr Sort Prop
with exec_instrs_ind4 := Minimality for exec_instrs Sort Prop
@@ -284,12 +310,13 @@ Scheme exec_instr_ind4 := Minimality for exec_instr Sort Prop
End RELSEM.
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
exists b, exists f, exists rs, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- exec_function ge f (Vptr Mem.nullptr Int.zero) (Regmap.init Vundef) m0 rs m /\
+ exec_function ge f (Vptr Mem.nullptr Int.zero) (Regmap.init Vundef) m0
+ t rs m /\
rs R3 = r.
diff --git a/backend/Machabstr.v b/backend/Machabstr.v
index 25458dcc..8d5d72a9 100644
--- a/backend/Machabstr.v
+++ b/backend/Machabstr.v
@@ -7,6 +7,7 @@ Require Import Mem.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -99,117 +100,132 @@ Variable ge: genv.
Inductive exec_instr:
function -> val -> frame ->
- code -> regset -> frame -> mem ->
+ code -> regset -> frame -> mem -> trace ->
code -> regset -> frame -> mem -> Prop :=
| exec_Mlabel:
forall f sp parent lbl c rs fr m,
exec_instr f sp parent
(Mlabel lbl :: c) rs fr m
- c rs fr m
+ E0 c rs fr m
| exec_Mgetstack:
forall f sp parent ofs ty dst c rs fr m v,
get_slot fr ty (Int.signed ofs) v ->
exec_instr f sp parent
(Mgetstack ofs ty dst :: c) rs fr m
- c (rs#dst <- v) fr m
+ E0 c (rs#dst <- v) fr m
| exec_Msetstack:
forall f sp parent src ofs ty c rs fr m fr',
set_slot fr ty (Int.signed ofs) (rs src) fr' ->
exec_instr f sp parent
(Msetstack src ofs ty :: c) rs fr m
- c rs fr' m
+ E0 c rs fr' m
| exec_Mgetparam:
forall f sp parent ofs ty dst c rs fr m v,
get_slot parent ty (Int.signed ofs) v ->
exec_instr f sp parent
(Mgetparam ofs ty dst :: c) rs fr m
- c (rs#dst <- v) fr m
+ E0 c (rs#dst <- v) fr m
| exec_Mop:
forall f sp parent op args res c rs fr m v,
eval_operation ge sp op rs##args = Some v ->
exec_instr f sp parent
(Mop op args res :: c) rs fr m
- c (rs#res <- v) fr m
+ E0 c (rs#res <- v) fr m
| exec_Mload:
forall f sp parent chunk addr args dst c rs fr m a v,
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
exec_instr f sp parent
(Mload chunk addr args dst :: c) rs fr m
- c (rs#dst <- v) fr m
+ E0 c (rs#dst <- v) fr m
| exec_Mstore:
forall f sp parent chunk addr args src c rs fr m m' a,
eval_addressing ge sp addr rs##args = Some a ->
Mem.storev chunk m a (rs src) = Some m' ->
exec_instr f sp parent
(Mstore chunk addr args src :: c) rs fr m
- c rs fr m'
+ E0 c rs fr m'
| exec_Mcall:
- forall f sp parent sig ros c rs fr m f' rs' m',
+ forall f sp parent sig ros c rs fr m t f' rs' m',
find_function ge ros rs = Some f' ->
- exec_function f' fr rs m rs' m' ->
+ exec_function f' fr rs m t rs' m' ->
exec_instr f sp parent
(Mcall sig ros :: c) rs fr m
- c rs' fr m'
+ t c rs' fr m'
+ | exec_Malloc:
+ forall f sp parent c rs fr m sz m' blk,
+ rs (Conventions.loc_alloc_argument) = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', blk) ->
+ exec_instr f sp parent
+ (Malloc :: c) rs fr m
+ E0 c (rs#Conventions.loc_alloc_result <-
+ (Vptr blk Int.zero)) fr m'
| exec_Mgoto:
forall f sp parent lbl c rs fr m c',
find_label lbl f.(fn_code) = Some c' ->
exec_instr f sp parent
(Mgoto lbl :: c) rs fr m
- c' rs fr m
+ E0 c' rs fr m
| exec_Mcond_true:
forall f sp parent cond args lbl c rs fr m c',
eval_condition cond rs##args = Some true ->
find_label lbl f.(fn_code) = Some c' ->
exec_instr f sp parent
(Mcond cond args lbl :: c) rs fr m
- c' rs fr m
+ E0 c' rs fr m
| exec_Mcond_false:
forall f sp parent cond args lbl c rs fr m,
eval_condition cond rs##args = Some false ->
exec_instr f sp parent
(Mcond cond args lbl :: c) rs fr m
- c rs fr m
+ E0 c rs fr m
with exec_instrs:
function -> val -> frame ->
- code -> regset -> frame -> mem ->
+ code -> regset -> frame -> mem -> trace ->
code -> regset -> frame -> mem -> Prop :=
| exec_refl:
forall f sp parent c rs fr m,
- exec_instrs f sp parent c rs fr m c rs fr m
+ exec_instrs f sp parent c rs fr m E0 c rs fr m
| exec_one:
- forall f sp parent c rs fr m c' rs' fr' m',
- exec_instr f sp parent c rs fr m c' rs' fr' m' ->
- exec_instrs f sp parent c rs fr m c' rs' fr' m'
+ forall f sp parent c rs fr m t c' rs' fr' m',
+ exec_instr f sp parent c rs fr m t c' rs' fr' m' ->
+ exec_instrs f sp parent c rs fr m t c' rs' fr' m'
| exec_trans:
- forall f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 c3 rs3 fr3 m3,
- exec_instrs f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- exec_instrs f sp parent c2 rs2 fr2 m2 c3 rs3 fr3 m3 ->
- exec_instrs f sp parent c1 rs1 fr1 m1 c3 rs3 fr3 m3
+ forall f sp parent c1 rs1 fr1 m1 t1 c2 rs2 fr2 m2 t2 c3 rs3 fr3 m3 t3,
+ exec_instrs f sp parent c1 rs1 fr1 m1 t1 c2 rs2 fr2 m2 ->
+ exec_instrs f sp parent c2 rs2 fr2 m2 t2 c3 rs3 fr3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_instrs f sp parent c1 rs1 fr1 m1 t3 c3 rs3 fr3 m3
with exec_function_body:
function -> frame -> val -> val ->
- regset -> mem -> regset -> mem -> Prop :=
+ regset -> mem -> trace -> regset -> mem -> Prop :=
| exec_funct_body:
- forall f parent link ra rs m rs' m1 m2 stk fr1 fr2 fr3 c,
+ forall f parent link ra rs m t rs' m1 m2 stk fr1 fr2 fr3 c,
Mem.alloc m 0 f.(fn_stacksize) = (m1, stk) ->
set_slot (init_frame f) Tint 0 link fr1 ->
set_slot fr1 Tint 4 ra fr2 ->
exec_instrs f (Vptr stk (Int.repr (-f.(fn_framesize)))) parent
f.(fn_code) rs fr2 m1
- (Mreturn :: c) rs' fr3 m2 ->
- exec_function_body f parent link ra rs m rs' (Mem.free m2 stk)
+ t (Mreturn :: c) rs' fr3 m2 ->
+ exec_function_body f parent link ra rs m t rs' (Mem.free m2 stk)
with exec_function:
- function -> frame -> regset -> mem -> regset -> mem -> Prop :=
- | exec_funct:
- forall f parent rs m rs' m',
+ fundef -> frame -> regset -> mem -> trace -> regset -> mem -> Prop :=
+ | exec_funct_internal:
+ forall f parent rs m t rs' m',
(forall link ra,
Val.has_type link Tint ->
Val.has_type ra Tint ->
- exec_function_body f parent link ra rs m rs' m') ->
- exec_function f parent rs m rs' m'.
+ exec_function_body f parent link ra rs m t rs' m') ->
+ exec_function (Internal f) parent rs m t rs' m'
+ | exec_funct_external:
+ forall ef parent args res rs1 rs2 m t,
+ event_match ef args t res ->
+ args = rs1##(Conventions.loc_external_arguments ef.(ef_sig)) ->
+ rs2 = (rs1#(Conventions.loc_result ef.(ef_sig)) <- res) ->
+ exec_function (External ef) parent rs1 m t rs2 m.
Scheme exec_instr_ind4 := Minimality for exec_instr Sort Prop
with exec_instrs_ind4 := Minimality for exec_instrs Sort Prop
@@ -222,134 +238,155 @@ Scheme exec_instr_ind4 := Minimality for exec_instr Sort Prop
by the [Scheme] command above. *)
Lemma exec_mutual_induction:
- forall (P
+ forall
+ (P
P0 : function ->
val ->
frame ->
code ->
regset ->
- frame -> mem -> code -> regset -> frame -> mem -> Prop)
+ frame ->
+ mem -> trace -> code -> regset -> frame -> mem -> Prop)
(P1 : function ->
- frame -> val -> val -> regset -> mem -> regset -> mem -> Prop)
- (P2 : function -> frame -> regset -> mem -> regset -> mem -> Prop),
+ frame ->
+ val -> val -> regset -> mem -> trace -> regset -> mem -> Prop)
+ (P2 : fundef ->
+ frame -> regset -> mem -> trace -> regset -> mem -> Prop),
(forall (f : function) (sp : val) (parent : frame) (lbl : label)
(c : list instruction) (rs : regset) (fr : frame) (m : mem),
- P f sp parent (Mlabel lbl :: c) rs fr m c rs fr m) ->
- (forall (f : function) (sp : val) (parent : frame) (ofs : int)
+ P f sp parent (Mlabel lbl :: c) rs fr m E0 c rs fr m) ->
+ (forall (f0 : function) (sp : val) (parent : frame) (ofs : int)
(ty : typ) (dst : mreg) (c : list instruction) (rs : regset)
(fr : frame) (m : mem) (v : val),
get_slot fr ty (Int.signed ofs) v ->
- P f sp parent (Mgetstack ofs ty dst :: c) rs fr m c rs # dst <- v fr
- m) ->
- (forall (f : function) (sp : val) (parent : frame) (src : mreg)
+ P f0 sp parent (Mgetstack ofs ty dst :: c) rs fr m E0 c rs # dst <- v
+ fr m) ->
+ (forall (f1 : function) (sp : val) (parent : frame) (src : mreg)
(ofs : int) (ty : typ) (c : list instruction) (rs : mreg -> val)
(fr : frame) (m : mem) (fr' : frame),
set_slot fr ty (Int.signed ofs) (rs src) fr' ->
- P f sp parent (Msetstack src ofs ty :: c) rs fr m c rs fr' m) ->
- (forall (f : function) (sp : val) (parent : frame) (ofs : int)
+ P f1 sp parent (Msetstack src ofs ty :: c) rs fr m E0 c rs fr' m) ->
+ (forall (f2 : function) (sp : val) (parent : frame) (ofs : int)
(ty : typ) (dst : mreg) (c : list instruction) (rs : regset)
(fr : frame) (m : mem) (v : val),
get_slot parent ty (Int.signed ofs) v ->
- P f sp parent (Mgetparam ofs ty dst :: c) rs fr m c rs # dst <- v fr
- m) ->
- (forall (f : function) (sp : val) (parent : frame) (op : operation)
+ P f2 sp parent (Mgetparam ofs ty dst :: c) rs fr m E0 c rs # dst <- v
+ fr m) ->
+ (forall (f3 : function) (sp : val) (parent : frame) (op : operation)
(args : list mreg) (res : mreg) (c : list instruction)
(rs : mreg -> val) (fr : frame) (m : mem) (v : val),
eval_operation ge sp op rs ## args = Some v ->
- P f sp parent (Mop op args res :: c) rs fr m c rs # res <- v fr m) ->
- (forall (f : function) (sp : val) (parent : frame)
+ P f3 sp parent (Mop op args res :: c) rs fr m E0 c rs # res <- v fr m) ->
+ (forall (f4 : function) (sp : val) (parent : frame)
(chunk : memory_chunk) (addr : addressing) (args : list mreg)
(dst : mreg) (c : list instruction) (rs : mreg -> val) (fr : frame)
(m : mem) (a v : val),
eval_addressing ge sp addr rs ## args = Some a ->
loadv chunk m a = Some v ->
- P f sp parent (Mload chunk addr args dst :: c) rs fr m c
+ P f4 sp parent (Mload chunk addr args dst :: c) rs fr m E0 c
rs # dst <- v fr m) ->
- (forall (f : function) (sp : val) (parent : frame)
+ (forall (f5 : function) (sp : val) (parent : frame)
(chunk : memory_chunk) (addr : addressing) (args : list mreg)
(src : mreg) (c : list instruction) (rs : mreg -> val) (fr : frame)
(m m' : mem) (a : val),
eval_addressing ge sp addr rs ## args = Some a ->
storev chunk m a (rs src) = Some m' ->
- P f sp parent (Mstore chunk addr args src :: c) rs fr m c rs fr m') ->
- (forall (f : function) (sp : val) (parent : frame) (sig : signature)
+ P f5 sp parent (Mstore chunk addr args src :: c) rs fr m E0 c rs fr
+ m') ->
+ (forall (f6 : function) (sp : val) (parent : frame) (sig : signature)
(ros : mreg + ident) (c : list instruction) (rs : regset)
- (fr : frame) (m : mem) (f' : function) (rs' : regset) (m' : mem),
+ (fr : frame) (m : mem) (t : trace) (f' : fundef) (rs' : regset)
+ (m' : mem),
find_function ge ros rs = Some f' ->
- exec_function f' fr rs m rs' m' ->
- P2 f' fr rs m rs' m' ->
- P f sp parent (Mcall sig ros :: c) rs fr m c rs' fr m') ->
- (forall (f : function) (sp : val) (parent : frame) (lbl : label)
- (c : list instruction) (rs : regset) (fr : frame) (m : mem)
- (c' : code),
- find_label lbl (fn_code f) = Some c' ->
- P f sp parent (Mgoto lbl :: c) rs fr m c' rs fr m) ->
- (forall (f : function) (sp : val) (parent : frame)
- (cond : condition) (args : list mreg) (lbl : label)
+ exec_function f' fr rs m t rs' m' ->
+ P2 f' fr rs m t rs' m' ->
+ P f6 sp parent (Mcall sig ros :: c) rs fr m t c rs' fr m') ->
+ (forall (f7 : function) (sp : val) (parent : frame)
(c : list instruction) (rs : mreg -> val) (fr : frame) (m : mem)
+ (sz : int) (m' : mem) (blk : block),
+ rs Conventions.loc_alloc_argument = Vint sz ->
+ alloc m 0 (Int.signed sz) = (m', blk) ->
+ P f7 sp parent (Malloc :: c) rs fr m E0 c
+ rs # Conventions.loc_alloc_result <- (Vptr blk Int.zero) fr m') ->
+ (forall (f7 : function) (sp : val) (parent : frame) (lbl : label)
+ (c : list instruction) (rs : regset) (fr : frame) (m : mem)
(c' : code),
+ find_label lbl (fn_code f7) = Some c' ->
+ P f7 sp parent (Mgoto lbl :: c) rs fr m E0 c' rs fr m) ->
+ (forall (f8 : function) (sp : val) (parent : frame) (cond : condition)
+ (args : list mreg) (lbl : label) (c : list instruction)
+ (rs : mreg -> val) (fr : frame) (m : mem) (c' : code),
eval_condition cond rs ## args = Some true ->
- find_label lbl (fn_code f) = Some c' ->
- P f sp parent (Mcond cond args lbl :: c) rs fr m c' rs fr m) ->
- (forall (f : function) (sp : val) (parent : frame)
- (cond : condition) (args : list mreg) (lbl : label)
- (c : list instruction) (rs : mreg -> val) (fr : frame) (m : mem),
+ find_label lbl (fn_code f8) = Some c' ->
+ P f8 sp parent (Mcond cond args lbl :: c) rs fr m E0 c' rs fr m) ->
+ (forall (f9 : function) (sp : val) (parent : frame) (cond : condition)
+ (args : list mreg) (lbl : label) (c : list instruction)
+ (rs : mreg -> val) (fr : frame) (m : mem),
eval_condition cond rs ## args = Some false ->
- P f sp parent (Mcond cond args lbl :: c) rs fr m c rs fr m) ->
- (forall (f : function) (sp : val) (parent : frame) (c : code)
+ P f9 sp parent (Mcond cond args lbl :: c) rs fr m E0 c rs fr m) ->
+ (forall (f10 : function) (sp : val) (parent : frame) (c : code)
(rs : regset) (fr : frame) (m : mem),
- P0 f sp parent c rs fr m c rs fr m) ->
- (forall (f : function) (sp : val) (parent : frame) (c : code)
- (rs : regset) (fr : frame) (m : mem) (c' : code) (rs' : regset)
- (fr' : frame) (m' : mem),
- exec_instr f sp parent c rs fr m c' rs' fr' m' ->
- P f sp parent c rs fr m c' rs' fr' m' ->
- P0 f sp parent c rs fr m c' rs' fr' m') ->
- (forall (f : function) (sp : val) (parent : frame) (c1 : code)
- (rs1 : regset) (fr1 : frame) (m1 : mem) (c2 : code) (rs2 : regset)
- (fr2 : frame) (m2 : mem) (c3 : code) (rs3 : regset) (fr3 : frame)
- (m3 : mem),
- exec_instrs f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- P0 f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- exec_instrs f sp parent c2 rs2 fr2 m2 c3 rs3 fr3 m3 ->
- P0 f sp parent c2 rs2 fr2 m2 c3 rs3 fr3 m3 ->
- P0 f sp parent c1 rs1 fr1 m1 c3 rs3 fr3 m3) ->
- (forall (f : function) (parent : frame) (link ra : val) (rs : regset)
- (m : mem) (rs' : regset) (m1 m2 : mem) (stk : block)
- (fr1 fr2 fr3 : frame) (c : list instruction),
- alloc m 0 (fn_stacksize f) = (m1, stk) ->
- set_slot (init_frame f) Tint 0 link fr1 ->
+ P0 f10 sp parent c rs fr m E0 c rs fr m) ->
+ (forall (f11 : function) (sp : val) (parent : frame) (c : code)
+ (rs : regset) (fr : frame) (m : mem) (t : trace) (c' : code)
+ (rs' : regset) (fr' : frame) (m' : mem),
+ exec_instr f11 sp parent c rs fr m t c' rs' fr' m' ->
+ P f11 sp parent c rs fr m t c' rs' fr' m' ->
+ P0 f11 sp parent c rs fr m t c' rs' fr' m') ->
+ (forall (f12 : function) (sp : val) (parent : frame) (c1 : code)
+ (rs1 : regset) (fr1 : frame) (m1 : mem) (t1 : trace) (c2 : code)
+ (rs2 : regset) (fr2 : frame) (m2 : mem) (t2 : trace) (c3 : code)
+ (rs3 : regset) (fr3 : frame) (m3 : mem) (t3 : trace),
+ exec_instrs f12 sp parent c1 rs1 fr1 m1 t1 c2 rs2 fr2 m2 ->
+ P0 f12 sp parent c1 rs1 fr1 m1 t1 c2 rs2 fr2 m2 ->
+ exec_instrs f12 sp parent c2 rs2 fr2 m2 t2 c3 rs3 fr3 m3 ->
+ P0 f12 sp parent c2 rs2 fr2 m2 t2 c3 rs3 fr3 m3 ->
+ t3 = t1 ** t2 -> P0 f12 sp parent c1 rs1 fr1 m1 t3 c3 rs3 fr3 m3) ->
+ (forall (f13 : function) (parent : frame) (link ra : val)
+ (rs : regset) (m : mem) (t : trace) (rs' : regset) (m1 m2 : mem)
+ (stk : block) (fr1 fr2 fr3 : frame) (c : list instruction),
+ alloc m 0 (fn_stacksize f13) = (m1, stk) ->
+ set_slot (init_frame f13) Tint 0 link fr1 ->
set_slot fr1 Tint 4 ra fr2 ->
- exec_instrs f (Vptr stk (Int.repr (-f.(fn_framesize)))) parent (fn_code f) rs fr2 m1 (Mreturn :: c) rs' fr3
- m2 ->
- P0 f (Vptr stk (Int.repr (-f.(fn_framesize)))) parent (fn_code f) rs fr2 m1 (Mreturn :: c) rs' fr3 m2 ->
- P1 f parent link ra rs m rs' (free m2 stk)) ->
- (forall (f : function) (parent : frame) (rs : regset) (m : mem)
- (rs' : regset) (m' : mem),
+ exec_instrs f13 (Vptr stk (Int.repr (- fn_framesize f13))) parent
+ (fn_code f13) rs fr2 m1 t (Mreturn :: c) rs' fr3 m2 ->
+ P0 f13 (Vptr stk (Int.repr (- fn_framesize f13))) parent
+ (fn_code f13) rs fr2 m1 t (Mreturn :: c) rs' fr3 m2 ->
+ P1 f13 parent link ra rs m t rs' (free m2 stk)) ->
+ (forall (f14 : function) (parent : frame) (rs : regset) (m : mem)
+ (t : trace) (rs' : regset) (m' : mem),
(forall link ra : val,
Val.has_type link Tint ->
Val.has_type ra Tint ->
- exec_function_body f parent link ra rs m rs' m') ->
+ exec_function_body f14 parent link ra rs m t rs' m') ->
(forall link ra : val,
Val.has_type link Tint ->
- Val.has_type ra Tint -> P1 f parent link ra rs m rs' m') ->
- P2 f parent rs m rs' m') ->
- (forall (f15 : function) (sp : val) (f16 : frame) (c : code)
- (r : regset) (f17 : frame) (m : mem) (c0 : code) (r0 : regset)
- (f18 : frame) (m0 : mem),
- exec_instr f15 sp f16 c r f17 m c0 r0 f18 m0 ->
- P f15 sp f16 c r f17 m c0 r0 f18 m0)
- /\ (forall (f15 : function) (sp : val) (f16 : frame) (c : code)
- (r : regset) (f17 : frame) (m : mem) (c0 : code) (r0 : regset)
- (f18 : frame) (m0 : mem),
- exec_instrs f15 sp f16 c r f17 m c0 r0 f18 m0 ->
- P0 f15 sp f16 c r f17 m c0 r0 f18 m0)
- /\ (forall (f15 : function) (f16 : frame) (v1 v2 : val) (r : regset) (m : mem)
- (r0 : regset) (m0 : mem),
- exec_function_body f15 f16 v1 v2 r m r0 m0 -> P1 f15 f16 v1 v2 r m r0 m0)
- /\ (forall (f15 : function) (f16 : frame) (r : regset) (m : mem)
+ Val.has_type ra Tint -> P1 f14 parent link ra rs m t rs' m') ->
+ P2 (Internal f14) parent rs m t rs' m') ->
+ (forall (ef : external_function) (parent : frame) (args : list val)
+ (res : val) (rs1 : mreg -> val) (rs2 : RegEq.t -> val) (m : mem)
+ (t0 : trace),
+ event_match ef args t0 res ->
+ args = rs1 ## (Conventions.loc_external_arguments (ef_sig ef)) ->
+ rs2 = rs1 # (Conventions.loc_result (ef_sig ef)) <- res ->
+ P2 (External ef) parent rs1 m t0 rs2 m) ->
+ (forall (f16 : function) (v : val) (f17 : frame) (c : code)
+ (r : regset) (f18 : frame) (m : mem) (t : trace) (c0 : code)
+ (r0 : regset) (f19 : frame) (m0 : mem),
+ exec_instr f16 v f17 c r f18 m t c0 r0 f19 m0 ->
+ P f16 v f17 c r f18 m t c0 r0 f19 m0)
+ /\ (forall (f16 : function) (v : val) (f17 : frame) (c : code)
+ (r : regset) (f18 : frame) (m : mem) (t : trace) (c0 : code)
+ (r0 : regset) (f19 : frame) (m0 : mem),
+ exec_instrs f16 v f17 c r f18 m t c0 r0 f19 m0 ->
+ P0 f16 v f17 c r f18 m t c0 r0 f19 m0)
+ /\ (forall (f16 : function) (f17 : frame) (v v0 : val) (r : regset)
+ (m : mem) (t : trace) (r0 : regset) (m0 : mem),
+ exec_function_body f16 f17 v v0 r m t r0 m0 ->
+ P1 f16 f17 v v0 r m t r0 m0)
+ /\ (forall (f16 : fundef) (f17 : frame) (r : regset) (m : mem) (t : trace)
(r0 : regset) (m0 : mem),
- exec_function f15 f16 r m r0 m0 -> P2 f15 f16 r m r0 m0).
+ exec_function f16 f17 r m t r0 m0 -> P2 f16 f17 r m t r0 m0).
Proof.
intros. split. apply (exec_instr_ind4 P P0 P1 P2); assumption.
split. apply (exec_instrs_ind4 P P0 P1 P2); assumption.
@@ -359,13 +396,12 @@ Qed.
End RELSEM.
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
exists b, exists f, exists rs, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- f.(fn_sig) = mksignature nil (Some Tint) /\
- exec_function ge f empty_frame (Regmap.init Vundef) m0 rs m /\
- rs (Conventions.loc_result f.(fn_sig)) = r.
+ exec_function ge f empty_frame (Regmap.init Vundef) m0 t rs m /\
+ rs R3 = r.
diff --git a/backend/Machabstr2mach.v b/backend/Machabstr2mach.v
index 8549cefc..0513cbee 100644
--- a/backend/Machabstr2mach.v
+++ b/backend/Machabstr2mach.v
@@ -6,6 +6,7 @@ Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -757,6 +758,36 @@ Proof.
exists ms'. split. auto. eapply callstack_store_aux; eauto.
Qed.
+(** Allocations of heap blocks can be performed in parallel in
+ both semantics, preserving [callstack_invariant]. *)
+
+Lemma callstack_alloc:
+ forall cs mm ms lo hi mm' blk,
+ callstack_invariant cs mm ms ->
+ Mem.alloc mm lo hi = (mm', blk) ->
+ exists ms',
+ Mem.alloc ms lo hi = (ms', blk) /\
+ callstack_invariant cs mm' ms'.
+Proof.
+ intros. inversion H.
+ caseEq (alloc ms lo hi). intros ms' blk' ALLOC'.
+ injection H0; intros. injection ALLOC'; intros.
+ assert (blk' = blk). congruence. rewrite H5 in H3. rewrite H5.
+ exists ms'; split. auto.
+ constructor.
+ (* frame *)
+ intros; eapply frame_match_alloc; eauto.
+ (* linked *)
+ auto.
+ (* others *)
+ intros. rewrite <- H2; rewrite <- H4; simpl.
+ rewrite H1; rewrite H3. unfold update. case (zeq b blk); auto.
+ (* next *)
+ rewrite <- H2; rewrite <- H4; simpl; congruence.
+ (* dom *)
+ eapply callstack_dom_incr; eauto. rewrite <- H4; simpl. omega.
+Qed.
+
(** At function entry, a new frame is pushed on the call stack,
and memory blocks are allocated in both semantics. Moreover,
the back link to the caller's activation record is set up
@@ -905,7 +936,7 @@ Let ge := Genv.globalenv p.
Definition exec_instr_prop
(f: function) (sp: val) (parent: frame)
- (c: code) (rs: regset) (fr: frame) (mm: mem)
+ (c: code) (rs: regset) (fr: frame) (mm: mem) (t: trace)
(c': code) (rs': regset) (fr': frame) (mm': mem) : Prop :=
forall ms stk base pstk pbase cs
(WTF: wt_function f)
@@ -916,12 +947,12 @@ Definition exec_instr_prop
(CSI: callstack_invariant ((fr, stk, base) :: (parent, pstk, pbase) :: cs) mm ms)
(SP: sp = Vptr stk base),
exists ms',
- exec_instr ge f sp c rs ms c' rs' ms' /\
+ exec_instr ge f sp c rs ms t c' rs' ms' /\
callstack_invariant ((fr', stk, base) :: (parent, pstk, pbase) :: cs) mm' ms'.
Definition exec_instrs_prop
(f: function) (sp: val) (parent: frame)
- (c: code) (rs: regset) (fr: frame) (mm: mem)
+ (c: code) (rs: regset) (fr: frame) (mm: mem) (t: trace)
(c': code) (rs': regset) (fr': frame) (mm': mem) : Prop :=
forall ms stk base pstk pbase cs
(WTF: wt_function f)
@@ -932,12 +963,12 @@ Definition exec_instrs_prop
(CSI: callstack_invariant ((fr, stk, base) :: (parent, pstk, pbase) :: cs) mm ms)
(SP: sp = Vptr stk base),
exists ms',
- exec_instrs ge f sp c rs ms c' rs' ms' /\
+ exec_instrs ge f sp c rs ms t c' rs' ms' /\
callstack_invariant ((fr', stk, base) :: (parent, pstk, pbase) :: cs) mm' ms'.
Definition exec_function_body_prop
(f: function) (parent: frame) (link ra: val)
- (rs: regset) (mm: mem)
+ (rs: regset) (mm: mem) (t: trace)
(rs': regset) (mm': mem) : Prop :=
forall ms pstk pbase cs
(WTF: wt_function f)
@@ -947,26 +978,26 @@ Definition exec_function_body_prop
(LINK: link = Vptr pstk pbase)
(CSI: callstack_invariant ((parent, pstk, pbase) :: cs) mm ms),
exists ms',
- exec_function_body ge f (Vptr pstk pbase) ra rs ms rs' ms' /\
+ exec_function_body ge f (Vptr pstk pbase) ra rs ms t rs' ms' /\
callstack_invariant ((parent, pstk, pbase) :: cs) mm' ms'.
Definition exec_function_prop
- (f: function) (parent: frame)
- (rs: regset) (mm: mem)
+ (f: fundef) (parent: frame)
+ (rs: regset) (mm: mem) (t: trace)
(rs': regset) (mm': mem) : Prop :=
forall ms pstk pbase cs
- (WTF: wt_function f)
+ (WTF: wt_fundef f)
(WTRS: wt_regset rs)
(WTPA: wt_frame parent)
(CSI: callstack_invariant ((parent, pstk, pbase) :: cs) mm ms),
exists ms',
- exec_function ge f (Vptr pstk pbase) rs ms rs' ms' /\
+ exec_function ge f (Vptr pstk pbase) rs ms t rs' ms' /\
callstack_invariant ((parent, pstk, pbase) :: cs) mm' ms'.
Lemma exec_function_equiv:
- forall f parent rs mm rs' mm',
- Machabstr.exec_function ge f parent rs mm rs' mm' ->
- exec_function_prop f parent rs mm rs' mm'.
+ forall f parent rs mm t rs' mm',
+ Machabstr.exec_function ge f parent rs mm t rs' mm' ->
+ exec_function_prop f parent rs mm t rs' mm'.
Proof.
apply (Machabstr.exec_function_ind4 ge
exec_instr_prop
@@ -1009,16 +1040,22 @@ Proof.
auto.
(* Mcall *)
red in H1.
- assert (WTF': wt_function f').
+ assert (WTF': wt_fundef f').
destruct ros; simpl in H.
- apply (Genv.find_funct_prop wt_function wt_p H).
+ apply (Genv.find_funct_prop wt_fundef wt_p H).
destruct (Genv.find_symbol ge i); try discriminate.
- apply (Genv.find_funct_ptr_prop wt_function wt_p H).
+ apply (Genv.find_funct_ptr_prop wt_fundef wt_p H).
generalize (H1 _ _ _ _ WTF' WTRS WTFR CSI).
intros [ms' [EXF CSI']].
exists ms'. split. apply exec_Mcall with f'; auto.
rewrite SP. auto.
auto.
+ (* Malloc *)
+ generalize (callstack_alloc _ _ _ _ _ _ _ CSI H0).
+ intros [ms' [ALLOC' CSI']].
+ exists ms'; split.
+ eapply exec_Malloc; eauto.
+ auto.
(* Mgoto *)
exists ms. split. constructor; auto. auto.
(* Mcond *)
@@ -1033,7 +1070,7 @@ Proof.
exists ms'. split. apply exec_one; auto. auto.
(* trans *)
generalize (subject_reduction_instrs p wt_p
- _ _ _ _ _ _ _ _ _ _ _ H WTF INCL WTRS WTFR WTPA).
+ _ _ _ _ _ _ _ _ _ _ _ _ H WTF INCL WTRS WTFR WTPA).
intros [INCL2 [WTRS2 WTFR2]].
generalize (H0 _ _ _ _ _ _ WTF INCL WTRS WTFR WTPA CSI SP).
intros [ms1 [EX1 CSI1]].
@@ -1069,7 +1106,7 @@ Proof.
generalize (H3 _ _ _ _ _ _ WTF (incl_refl _) WTRS WTFR2 WTPA
CSI3 (refl_equal _)).
intros [ms4 [EXEC CSI4]].
- generalize (exec_instrs_link_invariant _ _ _ _ _ _ _ _ _ _ _ _
+ generalize (exec_instrs_link_invariant _ _ _ _ _ _ _ _ _ _ _ _ _
H2 WTF (incl_refl _)).
intros [INCL LINKINV].
exists (free ms4 stk). split.
@@ -1082,39 +1119,42 @@ Proof.
apply LINKINV. rewrite FOUR. omega. eapply slot_gss; eauto. auto.
eapply callstack_function_return; eauto.
- (* function *)
+ (* internal function *)
+ inversion WTF. subst f0.
generalize (H0 (Vptr pstk pbase) Vzero I I
- ms pstk pbase cs WTF WTRS WTPA I (refl_equal _) CSI).
+ ms pstk pbase cs H2 WTRS WTPA I (refl_equal _) CSI).
intros [ms' [EXEC CSI']].
exists ms'. split. constructor. intros.
generalize (H0 (Vptr pstk pbase) ra I H1
- ms pstk pbase cs WTF WTRS WTPA H1 (refl_equal _) CSI).
+ ms pstk pbase cs H2 WTRS WTPA H1 (refl_equal _) CSI).
intros [ms1 [EXEC1 CSI1]].
rewrite (callstack_exten _ _ _ _ CSI' CSI1). auto.
auto.
+
+ (* external function *)
+ exists ms; split. econstructor; eauto. auto.
Qed.
End SIMULATION.
Theorem exec_program_equiv:
- forall p r,
+ forall p t r,
wt_program p ->
- Machabstr.exec_program p r ->
- Mach.exec_program p r.
+ Machabstr.exec_program p t r ->
+ Mach.exec_program p t r.
Proof.
- intros p r WTP [fptr [f [rs [mm [FINDPTR [FINDF [SIG [EXEC RES]]]]]]]].
- assert (WTF: wt_function f).
- apply (Genv.find_funct_ptr_prop wt_function WTP FINDF).
+ intros p t r WTP [fptr [f [rs [mm [FINDPTR [FINDF [EXEC RES]]]]]]].
+ assert (WTF: wt_fundef f).
+ apply (Genv.find_funct_ptr_prop wt_fundef WTP FINDF).
assert (WTRS: wt_regset (Regmap.init Vundef)).
red; intros. rewrite Regmap.gi; simpl; auto.
assert (WTFR: wt_frame empty_frame).
red; intros. simpl. auto.
generalize (exec_function_equiv p WTP
f empty_frame
- (Regmap.init Vundef) (Genv.init_mem p) rs mm
+ (Regmap.init Vundef) (Genv.init_mem p) t rs mm
EXEC (Genv.init_mem p) nullptr Int.zero nil
WTF WTRS WTFR (callstack_init p)).
intros [ms' [EXEC' CSI]].
- red. exists fptr; exists f; exists rs; exists ms'.
- intuition. rewrite SIG in RES. exact RES.
+ red. exists fptr; exists f; exists rs; exists ms'. tauto.
Qed.
diff --git a/backend/Machtyping.v b/backend/Machtyping.v
index 987269ba..92f283b6 100644
--- a/backend/Machtyping.v
+++ b/backend/Machtyping.v
@@ -7,6 +7,7 @@ Require Import Mem.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -57,6 +58,8 @@ Inductive wt_instr : instruction -> Prop :=
forall sig ros,
match ros with inl r => mreg_type r = Tint | inr s => True end ->
wt_instr (Mcall sig ros)
+ | wt_Malloc:
+ wt_instr Malloc
| wt_Mgoto:
forall lbl,
wt_instr (Mgoto lbl)
@@ -78,8 +81,16 @@ Record wt_function (f: function) : Prop := mk_wt_function {
f.(fn_framesize) <= -Int.min_signed
}.
+Inductive wt_fundef: fundef -> Prop :=
+ | wt_fundef_external: forall ef,
+ Conventions.sig_external_ok ef.(ef_sig) ->
+ wt_fundef (External ef)
+ | wt_function_internal: forall f,
+ wt_function f ->
+ wt_fundef (Internal f).
+
Definition wt_program (p: program) : Prop :=
- forall i f, In (i, f) (prog_funct p) -> wt_function f.
+ forall i f, In (i, f) (prog_funct p) -> wt_fundef f.
(** * Type soundness *)
@@ -89,8 +100,8 @@ Require Import Machabstr.
of Mach: for a well-typed Mach program, if a transition is taken
from a state where registers hold values of their static types,
registers in the final state hold values of their static types
- as well. This is a progress theorem for our type system.
- It is used in the proof of implcation from the abstract Mach
+ as well. This is a subject reduction theorem for our type system.
+ It is used in the proof of implication from the abstract Mach
semantics to the concrete Mach semantics (file [Machabstr2mach]).
*)
@@ -183,6 +194,14 @@ Proof.
apply incl_tl; auto.
Qed.
+Lemma wt_event_match:
+ forall ef args t res,
+ event_match ef args t res ->
+ Val.has_type res (proj_sig_res ef.(ef_sig)).
+Proof.
+ induction 1. inversion H0; exact I.
+Qed.
+
Section SUBJECT_REDUCTION.
Variable p: program.
@@ -191,7 +210,7 @@ Let ge := Genv.globalenv p.
Definition exec_instr_prop
(f: function) (sp: val) (parent: frame)
- (c1: code) (rs1: regset) (fr1: frame) (m1: mem)
+ (c1: code) (rs1: regset) (fr1: frame) (m1: mem) (t: trace)
(c2: code) (rs2: regset) (fr2: frame) (m2: mem) :=
forall (WTF: wt_function f)
(INCL: incl c1 f.(fn_code))
@@ -201,7 +220,7 @@ Definition exec_instr_prop
incl c2 f.(fn_code) /\ wt_regset rs2 /\ wt_frame fr2.
Definition exec_function_body_prop
(f: function) (parent: frame) (link ra: val)
- (rs1: regset) (m1: mem) (rs2: regset) (m2: mem) :=
+ (rs1: regset) (m1: mem) (t: trace) (rs2: regset) (m2: mem) :=
forall (WTF: wt_function f)
(WTRS: wt_regset rs1)
(WTPA: wt_frame parent)
@@ -209,26 +228,26 @@ Definition exec_function_body_prop
(WTRA: Val.has_type ra Tint),
wt_regset rs2.
Definition exec_function_prop
- (f: function) (parent: frame)
- (rs1: regset) (m1: mem) (rs2: regset) (m2: mem) :=
- forall (WTF: wt_function f)
+ (f: fundef) (parent: frame)
+ (rs1: regset) (m1: mem) (t: trace) (rs2: regset) (m2: mem) :=
+ forall (WTF: wt_fundef f)
(WTRS: wt_regset rs1)
(WTPA: wt_frame parent),
wt_regset rs2.
Lemma subject_reduction:
- (forall f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2,
- exec_instr ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- exec_instr_prop f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2)
-/\ (forall f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2,
- exec_instrs ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- exec_instr_prop f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2)
-/\ (forall f parent link ra rs1 m1 rs2 m2,
- exec_function_body ge f parent link ra rs1 m1 rs2 m2 ->
- exec_function_body_prop f parent link ra rs1 m1 rs2 m2)
-/\ (forall f parent rs1 m1 rs2 m2,
- exec_function ge f parent rs1 m1 rs2 m2 ->
- exec_function_prop f parent rs1 m1 rs2 m2).
+ (forall f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2,
+ exec_instr ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2 ->
+ exec_instr_prop f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2)
+/\ (forall f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2,
+ exec_instrs ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2 ->
+ exec_instr_prop f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2)
+/\ (forall f parent link ra rs1 m1 t rs2 m2,
+ exec_function_body ge f parent link ra rs1 m1 t rs2 m2 ->
+ exec_function_body_prop f parent link ra rs1 m1 t rs2 m2)
+/\ (forall f parent rs1 m1 t rs2 m2,
+ exec_function ge f parent rs1 m1 t rs2 m2 ->
+ exec_function_prop f parent rs1 m1 t rs2 m2).
Proof.
apply exec_mutual_induction; red; intros;
try (generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))); intro;
@@ -249,7 +268,7 @@ Proof.
subst args; subst op. simpl in H.
replace v with Vundef. simpl; auto. congruence.
replace (mreg_type res) with (snd (type_of_operation op)).
- apply type_of_operation_sound with function ge rs##args sp; auto.
+ apply type_of_operation_sound with fundef ge rs##args sp; auto.
rewrite <- H6; reflexivity.
apply wt_setreg; auto. inversion H1. rewrite H7.
@@ -257,11 +276,13 @@ Proof.
apply H1; auto.
destruct ros; simpl in H.
- apply (Genv.find_funct_prop wt_function wt_p H).
+ apply (Genv.find_funct_prop wt_fundef wt_p H).
destruct (Genv.find_symbol ge i).
- apply (Genv.find_funct_ptr_prop wt_function wt_p H).
+ apply (Genv.find_funct_ptr_prop wt_fundef wt_p H).
discriminate.
+ apply wt_setreg; auto. exact I.
+
apply incl_find_label with lbl; auto.
apply incl_find_label with lbl; auto.
@@ -277,26 +298,34 @@ Proof.
generalize (H3 WTF (incl_refl _) WTRS WTFR2 WTPA).
tauto.
- apply (H0 Vzero Vzero). exact I. exact I. auto. auto. auto.
+ apply (H0 Vzero Vzero). exact I. exact I.
+ inversion WTF; auto. auto. auto.
exact I. exact I.
+
+ rewrite H1. apply wt_setreg; auto.
+ generalize (wt_event_match _ _ _ _ H).
+ unfold proj_sig_res, Conventions.loc_result.
+ destruct (sig_res (ef_sig ef)).
+ destruct t; simpl; auto.
+ simpl; auto.
Qed.
Lemma subject_reduction_instr:
- forall f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2,
- exec_instr ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- exec_instr_prop f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2.
+ forall f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2,
+ exec_instr ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2 ->
+ exec_instr_prop f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2.
Proof (proj1 subject_reduction).
Lemma subject_reduction_instrs:
- forall f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2,
- exec_instrs ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
- exec_instr_prop f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2.
+ forall f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2,
+ exec_instrs ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2 ->
+ exec_instr_prop f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2.
Proof (proj1 (proj2 subject_reduction)).
Lemma subject_reduction_function:
- forall f parent rs1 m1 rs2 m2,
- exec_function ge f parent rs1 m1 rs2 m2 ->
- exec_function_prop f parent rs1 m1 rs2 m2.
+ forall f parent rs1 m1 t rs2 m2,
+ exec_function ge f parent rs1 m1 t rs2 m2 ->
+ exec_function_prop f parent rs1 m1 t rs2 m2.
Proof (proj2 (proj2 (proj2 subject_reduction))).
End SUBJECT_REDUCTION.
@@ -335,8 +364,8 @@ Proof.
Qed.
Lemma exec_instr_link_invariant:
- forall ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2,
- exec_instr ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
+ forall ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2,
+ exec_instr ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2 ->
wt_function f ->
incl c1 f.(fn_code) ->
incl c2 f.(fn_code) /\ link_invariant fr1 fr2.
@@ -351,8 +380,8 @@ Proof.
Qed.
Lemma exec_instrs_link_invariant:
- forall ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2,
- exec_instrs ge f sp parent c1 rs1 fr1 m1 c2 rs2 fr2 m2 ->
+ forall ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2,
+ exec_instrs ge f sp parent c1 rs1 fr1 m1 t c2 rs2 fr2 m2 ->
wt_function f ->
incl c1 f.(fn_code) ->
incl c2 f.(fn_code) /\ link_invariant fr1 fr2.
@@ -360,8 +389,8 @@ Proof.
induction 1; intros.
split. auto. apply link_invariant_refl.
eapply exec_instr_link_invariant; eauto.
- generalize (IHexec_instrs1 H1 H2); intros [A B].
- generalize (IHexec_instrs2 H1 A); intros [C D].
+ generalize (IHexec_instrs1 H2 H3); intros [A B].
+ generalize (IHexec_instrs2 H2 A); intros [C D].
split. auto. red; auto.
Qed.
diff --git a/backend/Main.v b/backend/Main.v
index 80a0577f..95dc4e6c 100644
--- a/backend/Main.v
+++ b/backend/Main.v
@@ -78,34 +78,34 @@ Notation "a @@ b" :=
The translation of a Cminor function to a PPC function is as follows. *)
-Definition transf_cminor_function (f: Cminor.function) : option PPC.code :=
+Definition transf_cminor_fundef (f: Cminor.fundef) : option PPC.fundef :=
Some f
- @@@ RTLgen.transl_function
- @@ Constprop.transf_function
- @@ CSE.transf_function
- @@@ Allocation.transf_function
- @@ Tunneling.tunnel_function
- @@ Linearize.transf_function
- @@@ Stacking.transf_function
- @@@ PPCgen.transf_function.
+ @@@ RTLgen.transl_fundef
+ @@ Constprop.transf_fundef
+ @@ CSE.transf_fundef
+ @@@ Allocation.transf_fundef
+ @@ Tunneling.tunnel_fundef
+ @@ Linearize.transf_fundef
+ @@@ Stacking.transf_fundef
+ @@@ PPCgen.transf_fundef.
(** And here is the translation for Csharpminor functions. *)
-Definition transf_csharpminor_function
- (gce: Cminorgen.compilenv) (f: Csharpminor.function) : option PPC.code :=
+Definition transf_csharpminor_fundef
+ (gce: Cminorgen.compilenv) (f: Csharpminor.fundef) : option PPC.fundef :=
Some f
- @@@ Cminorgen.transl_function gce
- @@@ transf_cminor_function.
+ @@@ Cminorgen.transl_fundef gce
+ @@@ transf_cminor_fundef.
(** The corresponding translations for whole program follow. *)
Definition transf_cminor_program (p: Cminor.program) : option PPC.program :=
- transform_partial_program transf_cminor_function p.
+ transform_partial_program transf_cminor_fundef p.
Definition transf_csharpminor_program (p: Csharpminor.program) : option PPC.program :=
let gce := Cminorgen.build_global_compilenv p in
transform_partial_program
- (transf_csharpminor_function gce)
+ (transf_csharpminor_fundef gce)
(Csharpminor.program_of_program p).
(** * Equivalence with whole program transformations *)
@@ -194,7 +194,7 @@ Qed.
Lemma transf_cminor_program_equiv:
forall p, transf_cminor_program2 p = transf_cminor_program p.
Proof.
- intro. unfold transf_cminor_program, transf_cminor_program2, transf_cminor_function.
+ intro. unfold transf_cminor_program, transf_cminor_program2, transf_cminor_fundef.
simpl.
unfold RTLgen.transl_program,
Constprop.transf_program, RTL.program.
@@ -223,7 +223,7 @@ Lemma transf_csharpminor_program_equiv:
forall p, transf_csharpminor_program2 p = transf_csharpminor_program p.
Proof.
intros.
- unfold transf_csharpminor_program2, transf_csharpminor_program, transf_csharpminor_function.
+ unfold transf_csharpminor_program2, transf_csharpminor_program, transf_csharpminor_fundef.
simpl.
replace transf_cminor_program2 with transf_cminor_program.
unfold transf_cminor_program, Cminorgen.transl_program, Cminor.program, PPC.program.
@@ -237,10 +237,10 @@ Qed.
composes the semantic preservation results for each pass. *)
Lemma transf_cminor_program2_correct:
- forall p tp n,
+ forall p tp t n,
transf_cminor_program2 p = Some tp ->
- Cminor.exec_program p (Vint n) ->
- PPC.exec_program tp (Vint n).
+ Cminor.exec_program p t (Vint n) ->
+ PPC.exec_program tp t (Vint n).
Proof.
intros until n.
unfold transf_cminor_program2.
@@ -274,10 +274,10 @@ Proof.
Qed.
Lemma transf_csharpminor_program2_correct:
- forall p tp n,
+ forall p tp t n,
transf_csharpminor_program2 p = Some tp ->
- Csharpminor.exec_program p (Vint n) ->
- PPC.exec_program tp (Vint n).
+ Csharpminor.exec_program p t (Vint n) ->
+ PPC.exec_program tp t (Vint n).
Proof.
intros until n; unfold transf_csharpminor_program2; simpl.
caseEq (Cminorgen.transl_program p).
@@ -291,20 +291,20 @@ Qed.
(** It follows that the whole compiler is semantic-preserving. *)
Theorem transf_cminor_program_correct:
- forall p tp n,
+ forall p tp t n,
transf_cminor_program p = Some tp ->
- Cminor.exec_program p (Vint n) ->
- PPC.exec_program tp (Vint n).
+ Cminor.exec_program p t (Vint n) ->
+ PPC.exec_program tp t (Vint n).
Proof.
intros. apply transf_cminor_program2_correct with p.
rewrite transf_cminor_program_equiv. assumption. assumption.
Qed.
Theorem transf_csharpminor_program_correct:
- forall p tp n,
+ forall p tp t n,
transf_csharpminor_program p = Some tp ->
- Csharpminor.exec_program p (Vint n) ->
- PPC.exec_program tp (Vint n).
+ Csharpminor.exec_program p t (Vint n) ->
+ PPC.exec_program tp t (Vint n).
Proof.
intros. apply transf_csharpminor_program2_correct with p.
rewrite transf_csharpminor_program_equiv. assumption. assumption.
diff --git a/backend/Mem.v b/backend/Mem.v
index 26d4c499..7af696e1 100644
--- a/backend/Mem.v
+++ b/backend/Mem.v
@@ -619,6 +619,88 @@ Qed.
Hint Resolve store_in_bounds store_inv.
+(** Build a block filled with the given initialization data. *)
+
+Fixpoint contents_init_data (pos: Z) (id: list init_data) {struct id}: contentmap :=
+ match id with
+ | nil => (fun y => Undef)
+ | Init_int8 n :: id' =>
+ store_contents Size8 (contents_init_data (pos + 1) id') pos (Vint n)
+ | Init_int16 n :: id' =>
+ store_contents Size16 (contents_init_data (pos + 2) id') pos (Vint n)
+ | Init_int32 n :: id' =>
+ store_contents Size32 (contents_init_data (pos + 4) id') pos (Vint n)
+ | Init_float32 f :: id' =>
+ store_contents Size32 (contents_init_data (pos + 4) id') pos (Vfloat f)
+ | Init_float64 f :: id' =>
+ store_contents Size64 (contents_init_data (pos + 8) id') pos (Vfloat f)
+ | Init_space n :: id' =>
+ contents_init_data (pos + Zmax n 0) id'
+ end.
+
+Definition size_init_data (id: init_data) : Z :=
+ match id with
+ | Init_int8 _ => 1
+ | Init_int16 _ => 2
+ | Init_int32 _ => 4
+ | Init_float32 _ => 4
+ | Init_float64 _ => 8
+ | Init_space n => Zmax n 0
+ end.
+
+Definition size_init_data_list (id: list init_data): Z :=
+ List.fold_right (fun id sz => size_init_data id + sz) 0 id.
+
+Remark size_init_data_list_pos:
+ forall id, size_init_data_list id >= 0.
+Proof.
+ induction id; simpl.
+ omega.
+ assert (size_init_data a >= 0). destruct a; simpl; try omega.
+ generalize (Zmax2 z 0). omega. omega.
+Qed.
+
+Remark contents_init_data_undef_outside:
+ forall id pos x,
+ x < pos \/ x >= pos + size_init_data_list id ->
+ contents_init_data pos id x = Undef.
+Proof.
+ induction id; simpl; intros.
+ auto.
+ generalize (size_init_data_list_pos id); intro.
+ assert (forall n delta dt,
+ delta = 1 + Z_of_nat n ->
+ x < pos \/ x >= pos + (delta + size_init_data_list id) ->
+ setN n pos dt (contents_init_data (pos + delta) id) x = Undef).
+ intros. unfold setN. rewrite update_o.
+ transitivity (contents_init_data (pos + delta) id x).
+ apply set_cont_outside. omega.
+ apply IHid. omega. omega.
+ unfold size_init_data in H; destruct a;
+ try (apply H1; [reflexivity|assumption]).
+ apply IHid. generalize (Zmax2 z 0). omega.
+Qed.
+
+Definition block_init_data (id: list init_data) : block_contents :=
+ mkblock 0 (size_init_data_list id)
+ (contents_init_data 0 id)
+ (contents_init_data_undef_outside id 0).
+
+Definition alloc_init_data (m: mem) (id: list init_data) : mem * block :=
+ (mkmem (update m.(nextblock)
+ (block_init_data id)
+ m.(blocks))
+ (Zsucc m.(nextblock))
+ (succ_nextblock_pos m),
+ m.(nextblock)).
+
+Remark block_init_data_empty:
+ block_init_data nil = empty_block 0 0.
+Proof.
+ unfold block_init_data, empty_block. simpl.
+ decEq. apply proof_irrelevance.
+Qed.
+
(** * Properties of the memory operations *)
(** ** Properties related to block validity *)
@@ -2032,6 +2114,32 @@ Proof.
apply free_first_list_inject. auto.
Qed.
+Lemma contents_init_data_inject:
+ forall id, contentmap_inject (contents_init_data 0 id) (contents_init_data 0 id) 0 (size_init_data_list id) 0.
+Proof.
+ intro id0.
+ set (sz0 := size_init_data_list id0).
+ assert (forall id pos,
+ 0 <= pos -> pos + size_init_data_list id <= sz0 ->
+ contentmap_inject (contents_init_data pos id) (contents_init_data pos id) 0 sz0 0).
+ induction id; simpl; intros.
+ red; intros; constructor.
+ assert (forall n dt,
+ size_init_data a = 1 + Z_of_nat n ->
+ content_inject dt dt ->
+ contentmap_inject (setN n pos dt (contents_init_data (pos + size_init_data a) id))
+ (setN n pos dt (contents_init_data (pos + size_init_data a) id))
+ 0 sz0 0).
+ intros. set (pos' := pos) in |- * at 3. replace pos' with (pos + 0).
+ apply setN_inject. apply IHid. omega. omega. auto. auto.
+ generalize (size_init_data_list_pos id). omega. unfold pos'; omega.
+ destruct a;
+ try (apply H1; [reflexivity|repeat constructor]).
+ apply IHid. generalize (Zmax2 z 0). omega. simpl in H0; omega.
+
+ apply H. omega. unfold sz0. omega.
+Qed.
+
End MEM_INJECT.
Hint Resolve val_inject_int val_inject_float val_inject_ptr val_inject_undef.
@@ -2251,3 +2359,27 @@ Proof.
right. intros. omega. left. auto.
apply mi_no_overlap0; auto.
Qed.
+
+Lemma alloc_parallel_inject:
+ forall f m1 m2 lo hi m1' m2' b1 b2,
+ mem_inject f m1 m2 ->
+ alloc m1 lo hi = (m1', b1) ->
+ alloc m2 lo hi = (m2', b2) ->
+ Int.min_signed <= lo -> hi <= Int.max_signed ->
+ mem_inject (extend_inject b1 (Some(b2,0)) f) m1' m2' /\
+ inject_incr f (extend_inject b1 (Some(b2,0)) f).
+Proof.
+ intros.
+ generalize (low_bound_alloc _ _ b2 _ _ _ H1). rewrite zeq_true; intro LOW.
+ generalize (high_bound_alloc _ _ b2 _ _ _ H1). rewrite zeq_true; intro HIGH.
+ eapply alloc_mapped_inject; eauto.
+ eapply alloc_right_inject; eauto.
+ eapply valid_new_block; eauto.
+ compute. intuition congruence.
+ rewrite LOW; auto.
+ rewrite HIGH; auto.
+ rewrite LOW; omega.
+ rewrite HIGH; omega.
+ intros. elim (mi_mappedblocks _ _ _ H _ _ _ H4); intros.
+ injection H1; intros. rewrite H7 in H5. omegaContradiction.
+Qed.
diff --git a/backend/Op.v b/backend/Op.v
index e0dcfa46..efd0d9ce 100644
--- a/backend/Op.v
+++ b/backend/Op.v
@@ -46,7 +46,9 @@ Inductive operation : Set :=
| Oundef: operation (**r set [rd] to undefined value *)
(*c Integer arithmetic: *)
| Ocast8signed: operation (**r [rd] is 8-bit sign extension of [r1] *)
+ | Ocast8unsigned: operation (**r [rd] is 8-bit zero extension of [r1] *)
| Ocast16signed: operation (**r [rd] is 16-bit sign extension of [r1] *)
+ | Ocast16unsigned: operation (**r [rd] is 16-bit zero extension of [r1] *)
| Oadd: operation (**r [rd = r1 + r2] *)
| Oaddimm: int -> operation (**r [rd = r1 + n] *)
| Osub: operation (**r [rd = r1 - r2] *)
@@ -166,8 +168,10 @@ Definition eval_operation
end
| Oaddrstack ofs, nil => offset_sp sp ofs
| Oundef, nil => Some Vundef
- | Ocast8signed, Vint n1 :: nil => Some (Vint (Int.cast8signed n1))
- | Ocast16signed, Vint n1 :: nil => Some (Vint (Int.cast16signed n1))
+ | Ocast8signed, v1 :: nil => Some (Val.cast8signed v1)
+ | Ocast8unsigned, v1 :: nil => Some (Val.cast8unsigned v1)
+ | Ocast16signed, v1 :: nil => Some (Val.cast16signed v1)
+ | Ocast16unsigned, v1 :: nil => Some (Val.cast16unsigned v1)
| Oadd, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.add n1 n2))
| Oadd, Vint n1 :: Vptr b2 n2 :: nil => Some (Vptr b2 (Int.add n2 n1))
| Oadd, Vptr b1 n1 :: Vint n2 :: nil => Some (Vptr b1 (Int.add n1 n2))
@@ -215,8 +219,8 @@ Definition eval_operation
Some (Vfloat (Float.add (Float.mul f1 f2) f3))
| Omulsubf, Vfloat f1 :: Vfloat f2 :: Vfloat f3 :: nil =>
Some (Vfloat (Float.sub (Float.mul f1 f2) f3))
- | Osingleoffloat, Vfloat f1 :: nil =>
- Some (Vfloat (Float.singleoffloat f1))
+ | Osingleoffloat, v1 :: nil =>
+ Some (Val.singleoffloat v1)
| Ointoffloat, Vfloat f1 :: nil =>
Some (Vint (Float.intoffloat f1))
| Ofloatofint, Vint n1 :: nil =>
@@ -396,7 +400,9 @@ Definition type_of_operation (op: operation) : list typ * typ :=
| Oaddrstack _ => (nil, Tint)
| Oundef => (nil, Tint) (* treated specially *)
| Ocast8signed => (Tint :: nil, Tint)
+ | Ocast8unsigned => (Tint :: nil, Tint)
| Ocast16signed => (Tint :: nil, Tint)
+ | Ocast16unsigned => (Tint :: nil, Tint)
| Oadd => (Tint :: Tint :: nil, Tint)
| Oaddimm _ => (Tint :: nil, Tint)
| Osub => (Tint :: Tint :: nil, Tint)
@@ -476,6 +482,10 @@ Proof.
destruct (Genv.find_symbol genv i); simplify_eq H1; intro; subst v; exact I.
simpl. unfold offset_sp in H1. destruct sp; try discriminate.
inversion H1. exact I.
+ destruct v0; exact I.
+ destruct v0; exact I.
+ destruct v0; exact I.
+ destruct v0; exact I.
destruct (eq_block b b0). injection H1; intro; subst v; exact I.
discriminate.
destruct (Int.eq i0 Int.zero). discriminate.
@@ -492,6 +502,7 @@ Proof.
injection H1; intro; subst v; exact I. discriminate.
destruct (Int.ltu i0 (Int.repr 32)).
injection H1; intro; subst v; exact I. discriminate.
+ destruct v0; exact I.
destruct (eval_condition c vl).
destruct b; injection H1; intro; subst v; exact I.
discriminate.
@@ -550,7 +561,9 @@ Definition eval_operation_total (sp: val) (op: operation) (vl: list val) : val :
| Oaddrstack ofs, nil => Val.add sp (Vint ofs)
| Oundef, nil => Vundef
| Ocast8signed, v1::nil => Val.cast8signed v1
+ | Ocast8unsigned, v1::nil => Val.cast8unsigned v1
| Ocast16signed, v1::nil => Val.cast16signed v1
+ | Ocast16unsigned, v1::nil => Val.cast16unsigned v1
| Oadd, v1::v2::nil => Val.add v1 v2
| Oaddimm n, v1::nil => Val.add v1 (Vint n)
| Osub, v1::v2::nil => Val.sub v1 v2
diff --git a/backend/PPC.v b/backend/PPC.v
index 64bd90a8..37f882b3 100644
--- a/backend/PPC.v
+++ b/backend/PPC.v
@@ -7,6 +7,7 @@ Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
(** * Abstract syntax *)
@@ -85,6 +86,7 @@ Inductive instruction : Set :=
| Paddi: ireg -> ireg -> constant -> instruction (**r add immediate *)
| Paddis: ireg -> ireg -> constant -> instruction (**r add immediate high *)
| Paddze: ireg -> ireg -> instruction (**r add Carry bit *)
+ | Pallocblock: instruction (**r allocate new heap block *)
| Pallocframe: Z -> Z -> instruction (**r allocate new stack frame *)
| Pand_: ireg -> ireg -> ireg -> instruction (**r bitwise and *)
| Pandc: ireg -> ireg -> ireg -> instruction (**r bitwise and-complement *)
@@ -260,6 +262,12 @@ lbl: .long 0x43300000, 0x00000000
Again, our memory model cannot comprehend that this operation
frees (logically) the current stack frame.
+- [Pallocheap]: in the formal semantics, this pseudo-instruction
+ allocates a heap block of size the contents of [GPR3], and leaves
+ a pointer to this block in [GPR3]. In the generated assembly code,
+ it is turned into a call to the allocation function of the run-time
+ system.
+
- [Piundef] and [Pfundef]: set an integer or float register (respectively)
to the [Vundef] value. Expand to nothing, as the PowerPC processor has
no notion of ``undefined value''. These two pseudo-instructions are used
@@ -277,7 +285,8 @@ lbl: .long 0x43300000, 0x00000000
*)
Definition code := list instruction.
-Definition program := AST.program code.
+Definition fundef := AST.fundef code.
+Definition program := AST.program fundef.
(** * Operational semantics *)
@@ -317,7 +326,7 @@ Module Pregmap := EMap(PregEq).
[Vzero] or [Vone]. *)
Definition regset := Pregmap.t val.
-Definition genv := Genv.t code.
+Definition genv := Genv.t fundef.
Notation "a # b" := (a b) (at level 1, only parsing).
Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level).
@@ -540,6 +549,14 @@ Definition exec_instr (c: code) (i: instruction) (rs: regset) (m: mem) : outcome
OK (nextinstr (rs#rd <- (Val.add (gpr_or_zero rs r1) (const_high cst)))) m
| Paddze rd r1 =>
OK (nextinstr (rs#rd <- (Val.add rs#r1 rs#CARRY))) m
+ | Pallocblock =>
+ match rs#GPR3 with
+ | Vint n =>
+ let (m', b) := Mem.alloc m 0 (Int.signed n) in
+ OK (nextinstr (rs#GPR3 <- (Vptr b Int.zero)
+ #LR <- (Val.add rs#PC Vone))) m'
+ | _ => Error
+ end
| Pallocframe lo hi =>
let (m1, stk) := Mem.alloc m lo hi in
let sp := Vptr stk (Int.repr lo) in
@@ -735,36 +752,84 @@ Definition exec_instr (c: code) (i: instruction) (rs: regset) (m: mem) : outcome
OK (nextinstr rs) m
end.
+(** Calling conventions for external functions. These are compatible with
+ the calling conventions in module [Conventions], except that
+ we use PPC registers instead of locations. *)
+
+Fixpoint loc_external_arguments_rec
+ (tyl: list typ) (iregl: list ireg) (fregl: list freg)
+ {struct tyl} : list preg :=
+ match tyl with
+ | nil => nil
+ | Tint :: tys =>
+ match iregl with
+ | nil => IR GPR11 (* should not happen *)
+ | ireg :: _ => IR ireg
+ end ::
+ loc_external_arguments_rec tys (list_drop1 iregl) fregl
+ | Tfloat :: tys =>
+ match fregl with
+ | nil => IR GPR11 (* should not happen *)
+ | freg :: _ => FR freg
+ end ::
+ loc_external_arguments_rec tys (list_drop2 iregl) (list_drop1 fregl)
+ end.
+
+Definition int_param_regs :=
+ GPR3 :: GPR4 :: GPR5 :: GPR6 :: GPR7 :: GPR8 :: GPR9 :: GPR10 :: nil.
+Definition float_param_regs :=
+ FPR1 :: FPR2 :: FPR3 :: FPR4 :: FPR5 :: FPR6 :: FPR7 :: FPR8 :: FPR9 :: FPR10 :: nil.
+
+Definition loc_external_arguments (s: signature) : list preg :=
+ loc_external_arguments_rec s.(sig_args) int_param_regs float_param_regs.
+
+Definition loc_external_result (s: signature) : preg :=
+ match s.(sig_res) with
+ | None => GPR3
+ | Some Tint => GPR3
+ | Some Tfloat => FPR1
+ end.
+
(** Execution of the instruction at [rs#PC]. *)
-Inductive exec_step: regset -> mem -> regset -> mem -> Prop :=
- | exec_step_intro:
+Inductive exec_step: regset -> mem -> trace -> regset -> mem -> Prop :=
+ | exec_step_internal:
forall b ofs c i rs m rs' m',
rs PC = Vptr b ofs ->
- Genv.find_funct_ptr ge b = Some c ->
+ Genv.find_funct_ptr ge b = Some (Internal c) ->
find_instr (Int.unsigned ofs) c = Some i ->
exec_instr c i rs m = OK rs' m' ->
- exec_step rs m rs' m'.
+ exec_step rs m E0 rs' m'
+ | exec_step_external:
+ forall b ef args res rs m t rs',
+ rs PC = Vptr b Int.zero ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ event_match ef args t res ->
+ args = List.map (fun r => rs#r) (loc_external_arguments ef.(ef_sig)) ->
+ rs' = (rs#(loc_external_result ef.(ef_sig)) <- res
+ #PC <- (rs LR)) ->
+ exec_step rs m t rs' m.
(** Execution of zero, one or several instructions starting at [rs#PC]. *)
-Inductive exec_steps: regset -> mem -> regset -> mem -> Prop :=
+Inductive exec_steps: regset -> mem -> trace -> regset -> mem -> Prop :=
| exec_refl:
forall rs m,
- exec_steps rs m rs m
+ exec_steps rs m E0 rs m
| exec_one:
- forall rs m rs' m',
- exec_step rs m rs' m' ->
- exec_steps rs m rs' m'
+ forall rs m t rs' m',
+ exec_step rs m t rs' m' ->
+ exec_steps rs m t rs' m'
| exec_trans:
- forall rs1 m1 rs2 m2 rs3 m3,
- exec_steps rs1 m1 rs2 m2 ->
- exec_steps rs2 m2 rs3 m3 ->
- exec_steps rs1 m1 rs3 m3.
+ forall rs1 m1 t1 rs2 m2 t2 rs3 m3 t3,
+ exec_steps rs1 m1 t1 rs2 m2 ->
+ exec_steps rs2 m2 t2 rs3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_steps rs1 m1 t3 rs3 m3.
End RELSEM.
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
let rs0 :=
@@ -772,4 +837,4 @@ Definition exec_program (p: program) (r: val) : Prop :=
# LR <- Vzero
# GPR1 <- (Vptr Mem.nullptr Int.zero) in
exists rs, exists m,
- exec_steps ge rs0 m0 rs m /\ rs#PC = Vzero /\ rs#GPR3 = r.
+ exec_steps ge rs0 m0 t rs m /\ rs#PC = Vzero /\ rs#GPR3 = r.
diff --git a/backend/PPCgen.v b/backend/PPCgen.v
index dc8ed40f..6cf06991 100644
--- a/backend/PPCgen.v
+++ b/backend/PPCgen.v
@@ -275,8 +275,12 @@ Definition transl_op
end
| Ocast8signed, a1 :: nil =>
Pextsb (ireg_of r) (ireg_of a1) :: k
+ | Ocast8unsigned, a1 :: nil =>
+ Prlwinm (ireg_of r) (ireg_of a1) Int.zero (Int.repr 255) :: k
| Ocast16signed, a1 :: nil =>
Pextsh (ireg_of r) (ireg_of a1) :: k
+ | Ocast16unsigned, a1 :: nil =>
+ Prlwinm (ireg_of r) (ireg_of a1) Int.zero (Int.repr 65535) :: k
| Oadd, a1 :: a2 :: nil =>
Padd (ireg_of r) (ireg_of a1) (ireg_of a2) :: k
| Oaddimm n, a1 :: nil =>
@@ -470,6 +474,8 @@ Definition transl_instr (i: Mach.instruction) (k: code) :=
Pmtctr (ireg_of r) :: Pbctrl :: k
| Mcall sig (inr symb) =>
Pbl symb :: k
+ | Malloc =>
+ Pallocblock :: k
| Mlabel lbl =>
Plabel lbl :: k
| Mgoto lbl =>
@@ -504,11 +510,14 @@ Fixpoint code_size (c: code) : Z :=
| instr :: c' => code_size c' + 1
end.
-Definition transf_function (f: Mach.function) :=
+Definition transf_function (f: Mach.function) : option PPC.code :=
let c := transl_function f in
if zlt Int.max_unsigned (code_size c)
then None
else Some c.
+Definition transf_fundef (f: Mach.fundef) : option PPC.fundef :=
+ transf_partial_fundef transf_function f.
+
Definition transf_program (p: Mach.program) : option PPC.program :=
- transform_partial_program transf_function p.
+ transform_partial_program transf_fundef p.
diff --git a/backend/PPCgenproof.v b/backend/PPCgenproof.v
index 99aa4c83..32649998 100644
--- a/backend/PPCgenproof.v
+++ b/backend/PPCgenproof.v
@@ -7,6 +7,7 @@ Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -29,53 +30,44 @@ Lemma symbols_preserved:
forall id, Genv.find_symbol tge id = Genv.find_symbol ge id.
Proof.
intros. unfold ge, tge.
- apply Genv.find_symbol_transf_partial with transf_function.
+ apply Genv.find_symbol_transf_partial with transf_fundef.
exact TRANSF.
Qed.
Lemma functions_translated:
forall f b,
Genv.find_funct_ptr ge b = Some f ->
- Genv.find_funct_ptr tge b = Some (transl_function f).
+ exists tf, Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = Some tf.
Proof.
intros.
- generalize (Genv.find_funct_ptr_transf_partial
- transf_function TRANSF H).
- intros [A B].
- unfold tge. change code with (list instruction). rewrite A.
- generalize B. unfold transf_function.
- case (zlt Int.max_unsigned (code_size (transl_function f))); intro.
- tauto. auto.
+ generalize (Genv.find_funct_ptr_transf_partial transf_fundef TRANSF H).
+ case (transf_fundef f).
+ intros f' [A B]. exists f'; split. assumption. auto.
+ tauto.
Qed.
-Lemma functions_translated_2:
- forall f v,
- Genv.find_funct ge v = Some f ->
- Genv.find_funct tge v = Some (transl_function f).
+Lemma functions_transl:
+ forall f b,
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ Genv.find_funct_ptr tge b = Some (Internal (transl_function f)).
Proof.
intros.
- generalize (Genv.find_funct_transf_partial
- transf_function TRANSF H).
- intros [A B].
- unfold tge. change code with (list instruction). rewrite A.
- generalize B. unfold transf_function.
+ destruct (functions_translated _ _ H) as [tf [A B]].
+ rewrite A. generalize B. unfold transf_fundef, transf_partial_fundef, transf_function.
case (zlt Int.max_unsigned (code_size (transl_function f))); intro.
- tauto. auto.
+ congruence. auto.
Qed.
-Lemma functions_translated_no_overflow:
+Lemma functions_transl_no_overflow:
forall b f,
- Genv.find_funct_ptr ge b = Some f ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
code_size (transl_function f) <= Int.max_unsigned.
Proof.
- intros.
- generalize (Genv.find_funct_ptr_transf_partial
- transf_function TRANSF H).
- intros [A B].
- generalize B.
- unfold transf_function.
+ intros.
+ destruct (functions_translated _ _ H) as [tf [A B]].
+ generalize B. unfold transf_fundef, transf_partial_fundef, transf_function.
case (zlt Int.max_unsigned (code_size (transl_function f))); intro.
- tauto. intro. omega.
+ congruence. intro; omega.
Qed.
(** * Properties of control flow *)
@@ -180,7 +172,7 @@ Qed.
Inductive transl_code_at_pc: val -> Mach.function -> Mach.code -> Prop :=
transl_code_at_pc_intro:
forall b ofs f c,
- Genv.find_funct_ptr ge b = Some f ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
code_tail (Int.unsigned ofs) (transl_function f) = transl_code c ->
transl_code_at_pc (Vptr b ofs) f c.
@@ -194,19 +186,20 @@ Lemma exec_straight_steps_1:
code_size fn <= Int.max_unsigned ->
forall b ofs,
rs#PC = Vptr b ofs ->
- Genv.find_funct_ptr tge b = Some fn ->
+ Genv.find_funct_ptr tge b = Some (Internal fn) ->
code_tail (Int.unsigned ofs) fn = c ->
- exec_steps tge rs m rs' m'.
+ exec_steps tge rs m E0 rs' m'.
Proof.
induction 1.
intros. apply exec_refl.
- intros. apply exec_trans with rs2 m2.
+ intros. apply exec_trans with E0 rs2 m2 E0.
apply exec_one; econstructor; eauto.
rewrite find_instr_tail. rewrite H5. auto.
apply IHexec_straight with b (Int.add ofs Int.one).
auto. rewrite H0. rewrite H3. reflexivity.
auto.
apply code_tail_next_int with i; auto.
+ traceEq.
Qed.
Lemma exec_straight_steps_2:
@@ -215,7 +208,7 @@ Lemma exec_straight_steps_2:
code_size fn <= Int.max_unsigned ->
forall b ofs,
rs#PC = Vptr b ofs ->
- Genv.find_funct_ptr tge b = Some fn ->
+ Genv.find_funct_ptr tge b = Some (Internal fn) ->
code_tail (Int.unsigned ofs) fn = c ->
exists ofs',
rs'#PC = Vptr b ofs'
@@ -233,11 +226,11 @@ Lemma exec_straight_steps:
transl_code_at_pc (rs PC) f c ->
exec_straight tge (transl_function f)
(transl_code c) rs m (transl_code c') rs' m' ->
- exec_steps tge rs m rs' m' /\ transl_code_at_pc (rs' PC) f c'.
+ exec_steps tge rs m E0 rs' m' /\ transl_code_at_pc (rs' PC) f c'.
Proof.
intros. inversion H.
- generalize (functions_translated_no_overflow _ _ H2). intro.
- generalize (functions_translated _ _ H2). intro.
+ generalize (functions_transl_no_overflow _ _ H2). intro.
+ generalize (functions_transl _ _ H2). intro.
split. eapply exec_straight_steps_1; eauto.
generalize (exec_straight_steps_2 _ _ _ _ _ _ _
H0 H6 _ _ (sym_equal H1) H7 H3).
@@ -490,7 +483,7 @@ End TRANSL_LABEL.
Lemma find_label_goto_label:
forall f lbl rs m c' b ofs,
- Genv.find_funct_ptr ge b = Some f ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
rs PC = Vptr b ofs ->
Mach.find_label lbl f.(fn_code) = Some c' ->
exists rs',
@@ -509,7 +502,7 @@ Proof.
split. rewrite Pregmap.gss. constructor. auto.
rewrite Int.unsigned_repr. replace (pos' - 0) with pos' in B.
auto. omega.
- generalize (functions_translated_no_overflow _ _ H).
+ generalize (functions_transl_no_overflow _ _ H).
omega.
intros. apply Pregmap.gso; auto.
Qed.
@@ -575,7 +568,7 @@ Qed.
Definition exec_instr_prop
(f: Mach.function) (sp: val)
- (c1: Mach.code) (ms1: Mach.regset) (m1: mem)
+ (c1: Mach.code) (ms1: Mach.regset) (m1: mem) (t: trace)
(c2: Mach.code) (ms2: Mach.regset) (m2: mem) :=
forall rs1
(WTF: wt_function f)
@@ -584,36 +577,36 @@ Definition exec_instr_prop
(AG: agree ms1 sp rs1),
exists rs2,
agree ms2 sp rs2
- /\ exec_steps tge rs1 m1 rs2 m2
+ /\ exec_steps tge rs1 m1 t rs2 m2
/\ transl_code_at_pc (rs2 PC) f c2.
Definition exec_function_body_prop
(f: Mach.function) (parent: val) (ra: val)
- (ms1: Mach.regset) (m1: mem)
+ (ms1: Mach.regset) (m1: mem) (t: trace)
(ms2: Mach.regset) (m2: mem) :=
forall rs1
(WTRA: Val.has_type ra Tint)
(RALR: rs1 LR = ra)
(WTF: wt_function f)
- (AT: Genv.find_funct ge (rs1 PC) = Some f)
+ (AT: Genv.find_funct ge (rs1 PC) = Some (Internal f))
(AG: agree ms1 parent rs1),
exists rs2,
agree ms2 parent rs2
- /\ exec_steps tge rs1 m1 rs2 m2
+ /\ exec_steps tge rs1 m1 t rs2 m2
/\ rs2 PC = rs1 LR.
Definition exec_function_prop
- (f: Mach.function) (parent: val)
- (ms1: Mach.regset) (m1: mem)
+ (f: Mach.fundef) (parent: val)
+ (ms1: Mach.regset) (m1: mem) (t: trace)
(ms2: Mach.regset) (m2: mem) :=
forall rs1
- (WTF: wt_function f)
+ (WTF: wt_fundef f)
(AT: Genv.find_funct ge (rs1 PC) = Some f)
(AG: agree ms1 parent rs1)
(WTRA: Val.has_type (rs1 LR) Tint),
exists rs2,
agree ms2 parent rs2
- /\ exec_steps tge rs1 m1 rs2 m2
+ /\ exec_steps tge rs1 m1 t rs2 m2
/\ rs2 PC = rs1 LR.
(** We show each case of the inductive proof of simulation as a separate
@@ -622,7 +615,7 @@ Definition exec_function_prop
Lemma exec_Mlabel_prop:
forall (f : function) (sp : val) (lbl : Mach.label)
(c : list Mach.instruction) (rs : Mach.regset) (m : mem),
- exec_instr_prop f sp (Mlabel lbl :: c) rs m c rs m.
+ exec_instr_prop f sp (Mlabel lbl :: c) rs m E0 c rs m.
Proof.
intros; red; intros.
assert (exec_straight tge (transl_function f)
@@ -637,7 +630,7 @@ Lemma exec_Mgetstack_prop:
forall (f : function) (sp : val) (ofs : int) (ty : typ) (dst : mreg)
(c : list Mach.instruction) (ms : Mach.regset) (m : mem) (v : val),
load_stack m sp ty ofs = Some v ->
- exec_instr_prop f sp (Mgetstack ofs ty dst :: c) ms m c (Regmap.set dst v ms) m.
+ exec_instr_prop f sp (Mgetstack ofs ty dst :: c) ms m E0 c (Regmap.set dst v ms) m.
Proof.
intros; red; intros.
unfold load_stack in H.
@@ -661,7 +654,7 @@ Lemma exec_Msetstack_prop:
forall (f : function) (sp : val) (src : mreg) (ofs : int) (ty : typ)
(c : list Mach.instruction) (ms : mreg -> val) (m m' : mem),
store_stack m sp ty ofs (ms src) = Some m' ->
- exec_instr_prop f sp (Msetstack src ofs ty :: c) ms m c ms m'.
+ exec_instr_prop f sp (Msetstack src ofs ty :: c) ms m E0 c ms m'.
Proof.
intros; red; intros.
unfold store_stack in H.
@@ -684,7 +677,7 @@ Lemma exec_Mgetparam_prop:
(m : mem) (v : val),
load_stack m sp Tint (Int.repr 0) = Some parent ->
load_stack m parent ty ofs = Some v ->
- exec_instr_prop f sp (Mgetparam ofs ty dst :: c) ms m c (Regmap.set dst v ms) m.
+ exec_instr_prop f sp (Mgetparam ofs ty dst :: c) ms m E0 c (Regmap.set dst v ms) m.
Proof.
intros; red; intros.
set (rs2 := nextinstr (rs1#GPR2 <- parent)).
@@ -723,7 +716,7 @@ Lemma exec_straight_exec_prop:
/\ agree ms' sp rs2) ->
(exists rs2,
agree ms' sp rs2
- /\ exec_steps tge rs1 m1 rs2 m2
+ /\ exec_steps tge rs1 m1 E0 rs2 m2
/\ transl_code_at_pc (rs2 PC) f c2).
Proof.
intros until ms'. intros TRANS1 [rs2 [EX AG]].
@@ -736,7 +729,7 @@ Lemma exec_Mop_prop:
(res : mreg) (c : list Mach.instruction) (ms: Mach.regset)
(m : mem) (v: val),
eval_operation ge sp op ms ## args = Some v ->
- exec_instr_prop f sp (Mop op args res :: c) ms m c (Regmap.set res v ms) m.
+ exec_instr_prop f sp (Mop op args res :: c) ms m E0 c (Regmap.set res v ms) m.
Proof.
intros; red; intros.
generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))).
@@ -753,7 +746,7 @@ Lemma exec_Mload_prop:
(a v : val),
eval_addressing ge sp addr ms ## args = Some a ->
loadv chunk m a = Some v ->
- exec_instr_prop f sp (Mload chunk addr args dst :: c) ms m c (Regmap.set dst v ms) m.
+ exec_instr_prop f sp (Mload chunk addr args dst :: c) ms m E0 c (Regmap.set dst v ms) m.
Proof.
intros; red; intros.
generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))).
@@ -801,7 +794,7 @@ Lemma exec_Mstore_prop:
(a : val),
eval_addressing ge sp addr ms ## args = Some a ->
storev chunk m a (ms src) = Some m' ->
- exec_instr_prop f sp (Mstore chunk addr args src :: c) ms m c ms m'.
+ exec_instr_prop f sp (Mstore chunk addr args src :: c) ms m E0 c ms m'.
Proof.
intros; red; intros.
generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))).
@@ -820,23 +813,23 @@ Hypothesis wt_prog: wt_program prog.
Lemma exec_Mcall_prop:
forall (f : function) (sp : val) (sig : signature)
(mos : mreg + ident) (c : list Mach.instruction) (ms : Mach.regset)
- (m : mem) (f' : function) (ms' : Mach.regset) (m' : mem),
+ (m : mem) (f' : Mach.fundef) (t: trace) (ms' : Mach.regset) (m' : mem),
find_function ge mos ms = Some f' ->
- exec_function ge f' sp ms m ms' m' ->
- exec_function_prop f' sp ms m ms' m' ->
- exec_instr_prop f sp (Mcall sig mos :: c) ms m c ms' m'.
+ exec_function ge f' sp ms m t ms' m' ->
+ exec_function_prop f' sp ms m t ms' m' ->
+ exec_instr_prop f sp (Mcall sig mos :: c) ms m t c ms' m'.
Proof.
intros; red; intros.
generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))).
intro WTI. inversion WTI.
inversion AT.
- assert (WTF': wt_function f').
+ assert (WTF': wt_fundef f').
destruct mos; simpl in H.
- apply (Genv.find_funct_prop wt_function wt_prog H).
+ apply (Genv.find_funct_prop wt_fundef wt_prog H).
destruct (Genv.find_symbol ge i); try discriminate.
- apply (Genv.find_funct_ptr_prop wt_function wt_prog H).
+ apply (Genv.find_funct_ptr_prop wt_fundef wt_prog H).
assert (NOOV: code_size (transl_function f) <= Int.max_unsigned).
- eapply functions_translated_no_overflow; eauto.
+ eapply functions_transl_no_overflow; eauto.
destruct mos; simpl in H; simpl transl_code in H7.
(* Indirect call *)
generalize (code_tail_next_int _ _ _ _ NOOV H7). intro CT1.
@@ -853,19 +846,19 @@ Proof.
generalize (H1 rs3 WTF' TFIND AG3 WTRA).
intros [rs4 [AG4 [EXF' PC4]]].
exists rs4. split. auto. split.
- apply exec_trans with rs2 m. apply exec_one. econstructor.
- eauto. apply functions_translated. eexact H6.
+ apply exec_trans with E0 rs2 m t. apply exec_one. econstructor.
+ eauto. apply functions_transl. eexact H6.
rewrite find_instr_tail. rewrite H7. reflexivity.
simpl. rewrite <- (ireg_val ms sp rs1); auto.
- apply exec_trans with rs3 m. apply exec_one. econstructor.
+ apply exec_trans with E0 rs3 m t. apply exec_one. econstructor.
unfold rs2, nextinstr. rewrite Pregmap.gss.
rewrite Pregmap.gso. rewrite <- H5. simpl. reflexivity.
- discriminate. apply functions_translated. eexact H6.
+ discriminate. apply functions_transl. eexact H6.
rewrite find_instr_tail. rewrite CT1. reflexivity.
simpl. replace (rs2 CTR) with (ms m0). reflexivity.
unfold rs2. rewrite nextinstr_inv. rewrite Pregmap.gss.
auto. discriminate.
- exact EXF'.
+ exact EXF'. traceEq. traceEq.
rewrite PC4. unfold rs3. rewrite Pregmap.gso. rewrite Pregmap.gss.
unfold rs2, nextinstr. rewrite Pregmap.gss. rewrite Pregmap.gso.
rewrite <- H5. simpl. constructor. auto. auto.
@@ -888,23 +881,36 @@ Proof.
generalize (H1 rs2 WTF' TFIND AG2 WTRA).
intros [rs3 [AG3 [EXF' PC3]]].
exists rs3. split. auto. split.
- apply exec_trans with rs2 m. apply exec_one. econstructor.
- eauto. apply functions_translated. eexact H6.
+ apply exec_trans with E0 rs2 m t. apply exec_one. econstructor.
+ eauto. apply functions_transl. eexact H6.
rewrite find_instr_tail. rewrite H7. reflexivity.
- simpl. reflexivity.
- exact EXF'.
+ simpl. reflexivity.
+ exact EXF'. traceEq.
rewrite PC3. unfold rs2. rewrite Pregmap.gso. rewrite Pregmap.gss.
rewrite <- H5. simpl. constructor. auto. auto.
discriminate.
intro FINDS. rewrite FINDS in H. discriminate.
Qed.
+Lemma exec_Malloc_prop:
+ forall (f : function) (sp : val) (c : list Mach.instruction)
+ (ms : Mach.regset) (m : mem) (sz : int) (m' : mem) (blk : block),
+ ms Conventions.loc_alloc_argument = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', blk) ->
+ exec_instr_prop f sp (Malloc :: c) ms m E0 c
+ (Regmap.set Conventions.loc_alloc_result (Vptr blk Int.zero) ms) m'.
+Proof.
+ intros; red; intros.
+ eapply exec_straight_exec_prop; eauto.
+ simpl. eapply transl_alloc_correct; eauto.
+Qed.
+
Lemma exec_Mgoto_prop:
forall (f : function) (sp : val) (lbl : Mach.label)
(c : list Mach.instruction) (ms : Mach.regset) (m : mem)
(c' : Mach.code),
Mach.find_label lbl (fn_code f) = Some c' ->
- exec_instr_prop f sp (Mgoto lbl :: c) ms m c' ms m.
+ exec_instr_prop f sp (Mgoto lbl :: c) ms m E0 c' ms m.
Proof.
intros; red; intros.
inversion AT.
@@ -912,7 +918,7 @@ Proof.
intros [rs2 [GOTO [AT2 INV]]].
exists rs2. split. apply agree_exten_2 with rs1; auto.
split. inversion AT. apply exec_one. econstructor; eauto.
- apply functions_translated; eauto.
+ apply functions_transl; eauto.
rewrite find_instr_tail. rewrite H7. simpl. reflexivity.
simpl. rewrite GOTO. auto. auto.
Qed.
@@ -923,7 +929,7 @@ Lemma exec_Mcond_true_prop:
(ms: Mach.regset) (m : mem) (c' : Mach.code),
eval_condition cond ms ## args = Some true ->
Mach.find_label lbl (fn_code f) = Some c' ->
- exec_instr_prop f sp (Mcond cond args lbl :: c) ms m c' ms m.
+ exec_instr_prop f sp (Mcond cond args lbl :: c) ms m E0 c' ms m.
Proof.
intros; red; intros.
generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))).
@@ -936,8 +942,8 @@ Proof.
cond args k1 ms sp rs1 m true H2 AG H).
simpl. intros [rs2 [EX [RES AG2]]].
inversion AT.
- generalize (functions_translated _ _ H6); intro FN.
- generalize (functions_translated_no_overflow _ _ H6); intro NOOV.
+ generalize (functions_transl _ _ H6); intro FN.
+ generalize (functions_transl_no_overflow _ _ H6); intro NOOV.
simpl in H7.
generalize (exec_straight_steps_2 _ _ _ _ _ _ _ EX
NOOV _ _ (sym_equal H5) FN H7).
@@ -955,7 +961,7 @@ Proof.
apply exec_one. econstructor; eauto.
rewrite find_instr_tail. rewrite CT2. unfold k1. rewrite ISSET. reflexivity.
simpl. rewrite RES. simpl. auto.
- auto.
+ traceEq. auto.
Qed.
Lemma exec_Mcond_false_prop:
@@ -963,7 +969,7 @@ Lemma exec_Mcond_false_prop:
(args : list mreg) (lbl : Mach.label) (c : list Mach.instruction)
(ms : Mach.regset) (m : mem),
eval_condition cond ms ## args = Some false ->
- exec_instr_prop f sp (Mcond cond args lbl :: c) ms m c ms m.
+ exec_instr_prop f sp (Mcond cond args lbl :: c) ms m E0 c ms m.
Proof.
intros; red; intros.
generalize (wt_function_instrs _ WTF _ (INCL _ (in_eq _ _))).
@@ -989,8 +995,8 @@ Proof.
Qed.
Lemma exec_instr_incl:
- forall f sp c rs m c' rs' m',
- Mach.exec_instr ge f sp c rs m c' rs' m' ->
+ forall f sp c rs m t c' rs' m',
+ Mach.exec_instr ge f sp c rs m t c' rs' m' ->
incl c f.(fn_code) -> incl c' f.(fn_code).
Proof.
induction 1; intros; eauto with coqlib.
@@ -999,8 +1005,8 @@ Proof.
Qed.
Lemma exec_instrs_incl:
- forall f sp c rs m c' rs' m',
- Mach.exec_instrs ge f sp c rs m c' rs' m' ->
+ forall f sp c rs m t c' rs' m',
+ Mach.exec_instrs ge f sp c rs m t c' rs' m' ->
incl c f.(fn_code) -> incl c' f.(fn_code).
Proof.
induction 1; intros.
@@ -1011,7 +1017,7 @@ Qed.
Lemma exec_refl_prop:
forall (f : function) (sp : val) (c : Mach.code) (ms : Mach.regset)
- (m : mem), exec_instr_prop f sp c ms m c ms m.
+ (m : mem), exec_instr_prop f sp c ms m E0 c ms m.
Proof.
intros; red; intros.
exists rs1. split. auto. split. apply exec_refl. auto.
@@ -1019,28 +1025,29 @@ Qed.
Lemma exec_one_prop:
forall (f : function) (sp : val) (c : Mach.code) (ms : Mach.regset)
- (m : mem) (c' : Mach.code) (ms' : Mach.regset) (m' : mem),
- Mach.exec_instr ge f sp c ms m c' ms' m' ->
- exec_instr_prop f sp c ms m c' ms' m' ->
- exec_instr_prop f sp c ms m c' ms' m'.
+ (m : mem) (t: trace) (c' : Mach.code) (ms' : Mach.regset) (m' : mem),
+ Mach.exec_instr ge f sp c ms m t c' ms' m' ->
+ exec_instr_prop f sp c ms m t c' ms' m' ->
+ exec_instr_prop f sp c ms m t c' ms' m'.
Proof.
auto.
Qed.
Lemma exec_trans_prop:
forall (f : function) (sp : val) (c1 : Mach.code) (ms1 : Mach.regset)
- (m1 : mem) (c2 : Mach.code) (ms2 : Mach.regset) (m2 : mem)
- (c3 : Mach.code) (ms3 : Mach.regset) (m3 : mem),
- exec_instrs ge f sp c1 ms1 m1 c2 ms2 m2 ->
- exec_instr_prop f sp c1 ms1 m1 c2 ms2 m2 ->
- exec_instrs ge f sp c2 ms2 m2 c3 ms3 m3 ->
- exec_instr_prop f sp c2 ms2 m2 c3 ms3 m3 ->
- exec_instr_prop f sp c1 ms1 m1 c3 ms3 m3.
+ (m1 : mem) (t1: trace) (c2 : Mach.code) (ms2 : Mach.regset) (m2 : mem)
+ (t2: trace) (c3 : Mach.code) (ms3 : Mach.regset) (m3 : mem) (t3: trace),
+ exec_instrs ge f sp c1 ms1 m1 t1 c2 ms2 m2 ->
+ exec_instr_prop f sp c1 ms1 m1 t1 c2 ms2 m2 ->
+ exec_instrs ge f sp c2 ms2 m2 t2 c3 ms3 m3 ->
+ exec_instr_prop f sp c2 ms2 m2 t2 c3 ms3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_instr_prop f sp c1 ms1 m1 t3 c3 ms3 m3.
Proof.
intros; red; intros.
generalize (H0 rs1 WTF INCL AT AG).
intros [rs2 [AG2 [EX2 AT2]]].
- generalize (exec_instrs_incl _ _ _ _ _ _ _ _ H INCL). intro INCL2.
+ generalize (exec_instrs_incl _ _ _ _ _ _ _ _ _ H INCL). intro INCL2.
generalize (H2 rs2 WTF INCL2 AT2 AG2).
intros [rs3 [AG3 [EX3 AT3]]].
exists rs3. split. auto. split. eapply exec_trans; eauto. auto.
@@ -1048,23 +1055,23 @@ Qed.
Lemma exec_function_body_prop_:
forall (f : function) (parent ra : val) (ms : Mach.regset) (m : mem)
- (ms' : Mach.regset) (m1 m2 m3 m4 : mem) (stk : block)
+ (t: trace) (ms' : Mach.regset) (m1 m2 m3 m4 : mem) (stk : block)
(c : list Mach.instruction),
alloc m (- fn_framesize f)
(align_16_top (- fn_framesize f) (fn_stacksize f)) = (m1, stk) ->
let sp := Vptr stk (Int.repr (- fn_framesize f)) in
store_stack m1 sp Tint (Int.repr 0) parent = Some m2 ->
store_stack m2 sp Tint (Int.repr 4) ra = Some m3 ->
- exec_instrs ge f sp (fn_code f) ms m3 (Mreturn :: c) ms' m4 ->
- exec_instr_prop f sp (fn_code f) ms m3 (Mreturn :: c) ms' m4 ->
+ exec_instrs ge f sp (fn_code f) ms m3 t (Mreturn :: c) ms' m4 ->
+ exec_instr_prop f sp (fn_code f) ms m3 t (Mreturn :: c) ms' m4 ->
load_stack m4 sp Tint (Int.repr 0) = Some parent ->
load_stack m4 sp Tint (Int.repr 4) = Some ra ->
- exec_function_body_prop f parent ra ms m ms' (free m4 stk).
+ exec_function_body_prop f parent ra ms m t ms' (free m4 stk).
Proof.
intros; red; intros.
generalize (Genv.find_funct_inv AT). intros [b EQPC].
generalize AT. rewrite EQPC. rewrite Genv.find_funct_find_funct_ptr. intro FN.
- generalize (functions_translated_no_overflow _ _ FN); intro NOOV.
+ generalize (functions_transl_no_overflow _ _ FN); intro NOOV.
set (rs2 := nextinstr (rs1#GPR1 <- sp #GPR2 <- Vundef)).
set (rs3 := nextinstr (rs2#GPR2 <- ra)).
set (rs4 := nextinstr rs3).
@@ -1132,18 +1139,18 @@ Proof.
elim AG7; auto. auto with ppcgen. auto with ppcgen.
unfold rs9; auto with ppcgen.
(* execution *)
- split. apply exec_trans with rs4 m3.
+ split. apply exec_trans with E0 rs4 m3 t.
eapply exec_straight_steps_1; eauto.
- apply functions_translated; auto.
- apply exec_trans with rs5 m4. assumption.
+ apply functions_transl; auto.
+ apply exec_trans with t rs5 m4 E0. assumption.
inversion AT5.
- apply exec_trans with rs8 (free m4 stk).
+ apply exec_trans with E0 rs8 (free m4 stk) E0.
eapply exec_straight_steps_1; eauto.
- apply functions_translated; auto.
+ apply functions_transl; auto.
apply exec_one. econstructor.
change rs8#PC with (Val.add (Val.add (Val.add rs5#PC Vone) Vone) Vone).
rewrite <- H8. simpl. reflexivity.
- apply functions_translated; eauto.
+ apply functions_transl; eauto.
assert (code_tail (Int.unsigned (Int.add (Int.add (Int.add ofs Int.one) Int.one) Int.one))
(transl_function f) = Pblr :: transl_code c).
eapply code_tail_next_int; auto.
@@ -1153,35 +1160,62 @@ Proof.
rewrite find_instr_tail. rewrite H13.
reflexivity.
reflexivity.
+ traceEq. traceEq. traceEq.
(* LR preservation *)
change rs9#PC with ra. auto.
Qed.
-Lemma exec_function_prop_:
+Lemma exec_function_internal_prop:
forall (f : function) (parent : val) (ms : Mach.regset) (m : mem)
- (ms' : Mach.regset) (m' : mem),
+ (t: trace) (ms' : Mach.regset) (m' : mem),
(forall ra : val,
Val.has_type ra Tint ->
- exec_function_body ge f parent ra ms m ms' m') ->
+ exec_function_body ge f parent ra ms m t ms' m') ->
(forall ra : val, Val.has_type ra Tint ->
- exec_function_body_prop f parent ra ms m ms' m') ->
- exec_function_prop f parent ms m ms' m'.
+ exec_function_body_prop f parent ra ms m t ms' m') ->
+ exec_function_prop (Internal f) parent ms m t ms' m'.
Proof.
intros; red; intros.
- apply (H0 rs1#LR WTRA rs1 WTRA (refl_equal _) WTF AT AG).
+ inversion WTF. subst f0.
+ apply (H0 rs1#LR WTRA rs1 WTRA (refl_equal _) H2 AT AG).
+Qed.
+
+Lemma exec_function_external_prop:
+ forall (ef : external_function) (parent : val) (args : list val)
+ (res : val) (ms1 ms2: Mach.regset) (m : mem)
+ (t : trace),
+ event_match ef args t res ->
+ args = ms1 ## (Conventions.loc_external_arguments (ef_sig ef)) ->
+ ms2 = Regmap.set (Conventions.loc_result (ef_sig ef)) res ms1 ->
+ exec_function_prop (External ef) parent ms1 m t ms2 m.
+Proof.
+ intros; red; intros.
+ destruct (Genv.find_funct_inv AT) as [b EQ].
+ rewrite EQ in AT. rewrite Genv.find_funct_find_funct_ptr in AT.
+ exists (rs1#(loc_external_result (ef_sig ef)) <- res #PC <- (rs1 LR)).
+ split. rewrite loc_external_result_match. rewrite H1. auto with ppcgen.
+ split. apply exec_one. eapply exec_step_external; eauto.
+ destruct (functions_translated _ _ AT) as [tf [A B]].
+ simpl in B. congruence.
+ rewrite H0. rewrite loc_external_arguments_match.
+ rewrite list_map_compose. apply list_map_exten; intros.
+ symmetry; eapply preg_val; eauto.
+ reflexivity.
Qed.
(** We then conclude by induction on the structure of the Mach
execution derivation. *)
Theorem transf_function_correct:
- forall f parent ms m ms' m',
- Mach.exec_function ge f parent ms m ms' m' ->
- exec_function_prop f parent ms m ms' m'.
+ forall f parent ms m t ms' m',
+ Mach.exec_function ge f parent ms m t ms' m' ->
+ exec_function_prop f parent ms m t ms' m'.
Proof
(Mach.exec_function_ind4 ge
- exec_instr_prop exec_instr_prop
- exec_function_body_prop exec_function_prop
+ exec_instr_prop
+ exec_instr_prop
+ exec_function_body_prop
+ exec_function_prop
exec_Mlabel_prop
exec_Mgetstack_prop
@@ -1191,6 +1225,7 @@ Proof
exec_Mload_prop
exec_Mstore_prop
exec_Mcall_prop
+ exec_Malloc_prop
exec_Mgoto_prop
exec_Mcond_true_prop
exec_Mcond_false_prop
@@ -1198,21 +1233,22 @@ Proof
exec_one_prop
exec_trans_prop
exec_function_body_prop_
- exec_function_prop_).
+ exec_function_internal_prop
+ exec_function_external_prop).
End PRESERVATION.
Theorem transf_program_correct:
- forall (p: Mach.program) (tp: PPC.program) (r: val),
+ forall (p: Mach.program) (tp: PPC.program) (t: trace) (r: val),
wt_program p ->
transf_program p = Some tp ->
- Mach.exec_program p r ->
- PPC.exec_program tp r.
+ Mach.exec_program p t r ->
+ PPC.exec_program tp t r.
Proof.
intros.
destruct H1 as [fptr [f [ms [m [FINDS [FINDF [EX RES]]]]]]].
- assert (WTF: wt_function f).
- apply (Genv.find_funct_ptr_prop wt_function H FINDF).
+ assert (WTF: wt_fundef f).
+ apply (Genv.find_funct_ptr_prop wt_fundef H FINDF).
set (ge := Genv.globalenv p) in *.
set (ms0 := Regmap.init Vundef) in *.
set (tge := Genv.globalenv tp).
@@ -1232,7 +1268,7 @@ Proof.
assert (WTRA: Val.has_type (rs0 LR) Tint).
exact I.
generalize (transf_function_correct p tp H0 H
- _ _ _ _ _ _ EX rs0 WTF AT AG WTRA).
+ _ _ _ _ _ _ _ EX rs0 WTF AT AG WTRA).
intros [rs [AG' [EX' RPC]]].
red. exists rs; exists m.
split. rewrite (Genv.init_mem_transf_partial _ _ H0). exact EX'.
diff --git a/backend/PPCgenproof1.v b/backend/PPCgenproof1.v
index 30eb3368..4a9ac948 100644
--- a/backend/PPCgenproof1.v
+++ b/backend/PPCgenproof1.v
@@ -14,6 +14,7 @@ Require Import Mach.
Require Import Machtyping.
Require Import PPC.
Require Import PPCgen.
+Require Conventions.
(** * Properties of low half/high half decomposition *)
@@ -430,6 +431,75 @@ Proof.
Qed.
Hint Resolve gpr_or_zero_not_zero gpr_or_zero_zero: ppcgen.
+(** Connection between Mach and PPC calling conventions for external
+ functions. *)
+
+Lemma loc_external_result_match:
+ forall sg,
+ PPC.loc_external_result sg = preg_of (Conventions.loc_result sg).
+Proof.
+ intros. destruct sg as [sargs sres].
+ destruct sres. destruct t; reflexivity. reflexivity.
+Qed.
+
+Remark list_map_drop1:
+ forall (A B: Set) (f: A -> B) (l: list A), list_drop1 (map f l) = map f (list_drop1 l).
+Proof.
+ intros; destruct l; reflexivity.
+Qed.
+
+Remark list_map_drop2:
+ forall (A B: Set) (f: A -> B) (l: list A), list_drop2 (map f l) = map f (list_drop2 l).
+Proof.
+ intros; destruct l. reflexivity. destruct l; reflexivity.
+Qed.
+
+Lemma loc_external_arguments_rec_match:
+ forall tyl iregl fregl ofs,
+ (forall r, In r iregl -> mreg_type r = Tint) ->
+ (forall r, In r fregl -> mreg_type r = Tfloat) ->
+ PPC.loc_external_arguments_rec tyl (map ireg_of iregl) (map freg_of fregl) =
+ List.map
+ (fun l => preg_of (match l with R r => r | S _ => IT1 end))
+ (Conventions.loc_arguments_rec tyl iregl fregl ofs).
+Proof.
+ induction tyl; intros; simpl.
+ auto.
+ destruct a; simpl; apply (f_equal2 (@cons preg)).
+ destruct iregl; simpl.
+ reflexivity. unfold preg_of; rewrite (H m); auto with coqlib.
+ rewrite list_map_drop1. apply IHtyl.
+ intros; apply H; apply list_drop1_incl; auto.
+ assumption.
+ destruct fregl; simpl.
+ reflexivity. unfold preg_of; rewrite (H0 m); auto with coqlib.
+ rewrite list_map_drop1. rewrite list_map_drop2. apply IHtyl.
+ intros; apply H; apply list_drop2_incl; auto.
+ intros; apply H0; apply list_drop1_incl; auto.
+Qed.
+
+Ltac ElimOrEq :=
+ match goal with
+ | |- (?x = ?y) \/ _ -> _ =>
+ let H := fresh in
+ (intro H; elim H; clear H;
+ [intro H; rewrite <- H; clear H | ElimOrEq])
+ | |- False -> _ =>
+ let H := fresh in (intro H; contradiction)
+ end.
+
+Lemma loc_external_arguments_match:
+ forall sg,
+ PPC.loc_external_arguments sg = List.map preg_of (Conventions.loc_external_arguments sg).
+Proof.
+ intros. destruct sg as [sgargs sgres]; unfold loc_external_arguments, Conventions.loc_external_arguments.
+ rewrite list_map_compose. unfold Conventions.loc_arguments.
+ rewrite <- loc_external_arguments_rec_match.
+ reflexivity.
+ intro; simpl; ElimOrEq; reflexivity.
+ intro; simpl; ElimOrEq; reflexivity.
+Qed.
+
(** * Execution of straight-line code *)
Section STRAIGHTLINE.
@@ -1198,6 +1268,22 @@ Proof.
rewrite (sp_val ms sp rs). auto. auto.
(* Oundef again *)
congruence.
+ (* Ocast8unsigned *)
+ exists (nextinstr (rs#(ireg_of res) <- (Val.rolm (ms m0) Int.zero (Int.repr 255)))).
+ split. apply exec_straight_one. repeat (rewrite (ireg_val ms sp rs)); auto. reflexivity.
+ replace (Val.cast8unsigned (ms m0))
+ with (Val.rolm (ms m0) Int.zero (Int.repr 255)).
+ auto with ppcgen.
+ unfold Val.rolm, Val.cast8unsigned. destruct (ms m0); auto.
+ rewrite Int.rolm_zero. rewrite Int.cast8unsigned_and. auto.
+ (* Ocast16unsigned *)
+ exists (nextinstr (rs#(ireg_of res) <- (Val.rolm (ms m0) Int.zero (Int.repr 65535)))).
+ split. apply exec_straight_one. repeat (rewrite (ireg_val ms sp rs)); auto. reflexivity.
+ replace (Val.cast16unsigned (ms m0))
+ with (Val.rolm (ms m0) Int.zero (Int.repr 65535)).
+ auto with ppcgen.
+ unfold Val.rolm, Val.cast16unsigned. destruct (ms m0); auto.
+ rewrite Int.rolm_zero. rewrite Int.cast16unsigned_and. auto.
(* Oaddimm *)
generalize (addimm_correct (ireg_of res) (ireg_of m0) i k rs m
(ireg_of_not_GPR2 m0)).
@@ -1562,5 +1648,31 @@ Proof.
auto. auto.
Qed.
+(** Translation of allocations *)
+
+Lemma transl_alloc_correct:
+ forall ms sp rs sz m m' blk k,
+ agree ms sp rs ->
+ ms Conventions.loc_alloc_argument = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', blk) ->
+ let ms' := Regmap.set Conventions.loc_alloc_result (Vptr blk Int.zero) ms in
+ exists rs',
+ exec_straight (Pallocblock :: k) rs m k rs' m'
+ /\ agree ms' sp rs'.
+Proof.
+ intros.
+ pose (rs' := nextinstr (rs#GPR3 <- (Vptr blk Int.zero) #LR <- (Val.add rs#PC Vone))).
+ exists rs'; split.
+ apply exec_straight_one. unfold exec_instr.
+ generalize (preg_val _ _ _ Conventions.loc_alloc_argument H).
+ unfold preg_of; intro. simpl in H2. rewrite <- H2. rewrite H0.
+ rewrite H1. reflexivity.
+ reflexivity.
+ unfold ms', rs'. apply agree_nextinstr. apply agree_set_other.
+ change (IR GPR3) with (preg_of Conventions.loc_alloc_result).
+ apply agree_set_mreg. auto.
+ simpl. tauto.
+Qed.
+
End STRAIGHTLINE.
diff --git a/backend/Parallelmove.v b/backend/Parallelmove.v
index f95416eb..b2ec930b 100644
--- a/backend/Parallelmove.v
+++ b/backend/Parallelmove.v
@@ -1,2529 +1,290 @@
-(** Translation of parallel moves into sequences of individual moves *)
+Require Import Coqlib.
+Require Parmov.
+Require Import Values.
+Require Import Events.
+Require Import AST.
+Require Import Locations.
+Require Import Conventions.
-(** The ``parallel move'' problem, also known as ``parallel assignment'',
- is the following. We are given a list of (source, destination) pairs
- of locations. The goal is to find a sequence of elementary
- moves ([loc <- loc] assignments) such that, at the end of the sequence,
- location [dst] contains the value of location [src] at the beginning of
- the sequence, for each ([src], [dst]) pairs in the given problem.
- Moreover, other locations should keep their values, except one register
- of each type, which can be used as temporaries in the generated sequences.
+Definition temp_for (l: loc) : loc :=
+ match Loc.type l with Tint => R IT2 | Tfloat => R FT2 end.
- The parallel move problem is trivial if the sources and destinations do
- not overlap. For instance,
-<<
- (R1, R2) <- (R3, R4) becomes R1 <- R3; R2 <- R4
->>
- However, arbitrary overlap is allowed between sources and destinations.
- This requires some care in ordering the individual moves, as in
-<<
- (R1, R2) <- (R3, R1) becomes R2 <- R1; R1 <- R3
->>
- Worse, cycles (permutations) can require the use of temporaries, as in
-<<
- (R1, R2, R3) <- (R2, R3, R1) becomes T <- R1; R1 <- R2; R2 <- R3; R3 <- T;
->>
- An amazing fact is that for any parallel move problem, at most one temporary
- (or in our case one integer temporary and one float temporary) is needed.
+Definition parmove (srcs dsts: list loc) :=
+ Parmov.parmove2 loc temp_for Loc.eq srcs dsts.
- The development in this section was contributed by Laurence Rideau and
- Bernard Serpette. It is described in their paper
- ``Coq à la conquête des moulins'', JFLA 2005,
- #<A HREF="http://www-sop.inria.fr/lemme/Laurence.Rideau/RideauSerpetteJFLA05.ps">#
- http://www-sop.inria.fr/lemme/Laurence.Rideau/RideauSerpetteJFLA05.ps
- #</A>#
-*)
+Definition moves := (list (loc * loc))%type.
-Require Omega.
-Require Import Wf_nat.
-Require Import Conventions.
-Require Import Coqlib.
-Require Import Bool_nat.
-Require Import TheoryList.
-Require Import Bool.
-Require Import Arith.
-Require Import Peano_dec.
-Require Import EqNat.
-Require Import Values.
-Require Import LTL.
-Require Import EqNat.
-Require Import Locations.
-Require Import AST.
-
-Section pmov.
-
-Ltac caseEq name := generalize (refl_equal name); pattern name at -1; case name.
-Hint Resolve beq_nat_eq .
-
-Lemma neq_is_neq: forall (x y : nat), x <> y -> beq_nat x y = false.
-Proof.
-unfold not; intros.
-caseEq (beq_nat x y); auto.
-intro.
-elim H; auto.
-Qed.
-Hint Resolve neq_is_neq .
-
-Lemma app_nil: forall (A : Set) (l : list A), l ++ nil = l.
-Proof.
-intros A l; induction l as [|a l Hrecl]; auto; simpl; rewrite Hrecl; auto.
-Qed.
-
-Lemma app_cons:
- forall (A : Set) (l1 l2 : list A) (a : A), (a :: l1) ++ l2 = a :: (l1 ++ l2).
-Proof.
-auto.
-Qed.
-
-Lemma app_app:
- forall (A : Set) (l1 l2 l3 : list A), l1 ++ (l2 ++ l3) = (l1 ++ l2) ++ l3.
-Proof.
-intros A l1; induction l1 as [|a l1 Hrecl1]; simpl; auto; intros;
- rewrite Hrecl1; auto.
-Qed.
-
-Lemma app_rewrite:
- forall (A : Set) (l : list A) (x : A),
- (exists y : A , exists r : list A , l ++ (x :: nil) = y :: r ).
-Proof.
-intros A l x; induction l as [|a l Hrecl].
-exists x; exists (nil (A:=A)); auto.
-inversion Hrecl; inversion H.
-exists a; exists (l ++ (x :: nil)); auto.
-Qed.
-
-Lemma app_rewrite2:
- forall (A : Set) (l l2 : list A) (x : A),
- (exists y : A , exists r : list A , l ++ (x :: l2) = y :: r ).
-Proof.
-intros A l l2 x; induction l as [|a l Hrecl].
-exists x; exists l2; auto.
-inversion Hrecl; inversion H.
-exists a; exists (l ++ (x :: l2)); auto.
-Qed.
-
-Lemma app_rewriter:
- forall (A : Set) (l : list A) (x : A),
- (exists y : A , exists r : list A , x :: l = r ++ (y :: nil) ).
-Proof.
-intros A l x; induction l as [|a l Hrecl].
-exists x; exists (nil (A:=A)); auto.
-inversion Hrecl; inversion H.
-generalize H0; case x1; simpl; intros; inversion H1.
-exists a; exists (x0 :: nil); simpl; auto.
-exists x0; exists (a0 :: (a :: l0)); simpl; auto.
-Qed.
-Hint Resolve app_rewriter .
-
-Definition Reg := loc.
-
-Definition T :=
- fun (r : loc) =>
- match Loc.type r with Tint => R IT2 | Tfloat => R FT2 end.
-
-Definition notemporary := fun (r : loc) => forall x, Loc.diff r (T x).
-
-Definition Move := (Reg * Reg)%type.
-
-Definition Moves := list Move.
-
-Definition State := ((Moves * Moves) * Moves)%type.
-
-Definition StateToMove (r : State) : Moves :=
- match r with ((t, b), l) => t end.
-
-Definition StateBeing (r : State) : Moves :=
- match r with ((t, b), l) => b end.
-
-Definition StateDone (r : State) : Moves :=
- match r with ((t, b), l) => l end.
-
-Fixpoint noRead (p : Moves) (r : Reg) {struct p} : Prop :=
- match p with
- nil => True
- | (s, d) :: l => Loc.diff s r /\ noRead l r
- end.
-
-Lemma noRead_app:
- forall (l1 l2 : Moves) (r : Reg),
- noRead l1 r -> noRead l2 r -> noRead (l1 ++ l2) r.
-Proof.
-intros; induction l1 as [|a l1 Hrecl1]; simpl; auto.
-destruct a.
-elim H; intros; split; auto.
-Qed.
-
-Inductive step : State -> State -> Prop :=
- step_nop:
- forall (r : Reg) (t1 t2 l : Moves),
- step (t1 ++ ((r, r) :: t2), nil, l) (t1 ++ t2, nil, l)
- | step_start:
- forall (t1 t2 l : Moves) (m : Move),
- step (t1 ++ (m :: t2), nil, l) (t1 ++ t2, m :: nil, l)
- | step_push:
- forall (t1 t2 b l : Moves) (s d r : Reg),
- step
- (t1 ++ ((d, r) :: t2), (s, d) :: b, l)
- (t1 ++ t2, (d, r) :: ((s, d) :: b), l)
- | step_loop:
- forall (t b l : Moves) (s d r0 r0ounon : Reg),
- step
- (t, (s, r0ounon) :: (b ++ ((r0, d) :: nil)), l)
- (t, (s, r0ounon) :: (b ++ ((T r0, d) :: nil)), (r0, T r0) :: l)
- | step_pop:
- forall (t b l : Moves) (s0 d0 sn dn : Reg),
- noRead t dn ->
- Loc.diff dn s0 ->
- step
- (t, (sn, dn) :: (b ++ ((s0, d0) :: nil)), l)
- (t, b ++ ((s0, d0) :: nil), (sn, dn) :: l)
- | step_last:
- forall (t l : Moves) (s d : Reg),
- noRead t d -> step (t, (s, d) :: nil, l) (t, nil, (s, d) :: l) .
-Hint Resolve step_nop step_start step_push step_loop step_pop step_last .
-
-Fixpoint path (l : Moves) : Prop :=
- match l with
- nil => True
- | (s, d) :: l =>
- match l with
- nil => True
- | (ss, dd) :: l2 => s = dd /\ path l
- end
- end.
-
-Lemma path_pop: forall (m : Move) (l : Moves), path (m :: l) -> path l.
-Proof.
-simpl; intros m l; destruct m as [ms md]; case l; auto.
-intros m0; destruct m0; intros; inversion H; auto.
-Qed.
-
-Fixpoint noWrite (p : Moves) (r : Reg) {struct p} : Prop :=
- match p with
- nil => True
- | (s, d) :: l => Loc.diff d r /\ noWrite l r
- end.
-
-Lemma noWrite_pop:
- forall (p1 p2 : Moves) (m : Move) (r : Reg),
- noWrite (p1 ++ (m :: p2)) r -> noWrite (p1 ++ p2) r.
-Proof.
-intros; induction p1 as [|a p1 Hrecp1].
-generalize H; simpl; case m; intros; inversion H0; auto.
-generalize H; rewrite app_cons; rewrite app_cons.
-simpl; case a; intros.
-inversion H0; split; auto.
-Qed.
-
-Lemma noWrite_in:
- forall (p1 p2 : Moves) (r0 r1 r2 : Reg),
- noWrite (p1 ++ ((r1, r2) :: p2)) r0 -> Loc.diff r0 r2.
-Proof.
-intros; induction p1 as [|a p1 Hrecp1]; simpl; auto.
-generalize H; simpl; intros; inversion H0; auto.
-apply Loc.diff_sym; auto.
-generalize H; rewrite app_cons; simpl; case a; intros.
-apply Hrecp1; inversion H0; auto.
-Qed.
-
-Lemma noWrite_swap:
- forall (p : Moves) (m1 m2 : Move) (r : Reg),
- noWrite (m1 :: (m2 :: p)) r -> noWrite (m2 :: (m1 :: p)) r.
-Proof.
-intros p m1 m2 r; simpl; case m1; case m2.
-intros; inversion H; inversion H1; split; auto.
-Qed.
-
-Lemma noWrite_movFront:
- forall (p1 p2 : Moves) (m : Move) (r0 : Reg),
- noWrite (p1 ++ (m :: p2)) r0 -> noWrite (m :: (p1 ++ p2)) r0.
-Proof.
-intros p1 p2 m r0; induction p1 as [|a p1 Hrecp1]; auto.
-case a; intros r1 r2; rewrite app_cons; rewrite app_cons.
-intros; apply noWrite_swap; rewrite <- app_cons.
-simpl in H |-; inversion H; unfold noWrite; fold noWrite; auto.
-Qed.
-
-Lemma noWrite_insert:
- forall (p1 p2 : Moves) (m : Move) (r0 : Reg),
- noWrite (m :: (p1 ++ p2)) r0 -> noWrite (p1 ++ (m :: p2)) r0.
-Proof.
-intros p1 p2 m r0; induction p1 as [|a p1 Hrecp1].
-simpl; auto.
-destruct a; simpl.
-destruct m.
-intros [H1 [H2 H3]]; split; auto.
-apply Hrecp1.
-simpl; auto.
-Qed.
-
-Lemma noWrite_tmpLast:
- forall (t : Moves) (r s d : Reg),
- noWrite (t ++ ((s, d) :: nil)) r ->
- forall (x : Reg), noWrite (t ++ ((x, d) :: nil)) r.
-Proof.
-intros; induction t as [|a t Hrect].
-simpl; auto.
-generalize H; simpl; case a; intros; inversion H0; split; auto.
-Qed.
-
-Fixpoint simpleDest (p : Moves) : Prop :=
- match p with
- nil => True
- | (s, d) :: l => noWrite l d /\ simpleDest l
- end.
-
-Lemma simpleDest_Pop:
- forall (m : Move) (l1 l2 : Moves),
- simpleDest (l1 ++ (m :: l2)) -> simpleDest (l1 ++ l2).
-Proof.
-intros; induction l1 as [|a l1 Hrecl1].
-generalize H; simpl; case m; intros; inversion H0; auto.
-generalize H; rewrite app_cons; rewrite app_cons.
-simpl; case a; intros; inversion H0; split; auto.
-apply (noWrite_pop l1 l2 m); auto.
-Qed.
-
-Lemma simpleDest_pop:
- forall (m : Move) (l : Moves), simpleDest (m :: l) -> simpleDest l.
-Proof.
-intros m l; simpl; case m; intros _ r [X Y]; auto.
-Qed.
-
-Lemma simpleDest_right:
- forall (l1 l2 : Moves), simpleDest (l1 ++ l2) -> simpleDest l2.
-Proof.
-intros l1; induction l1 as [|a l1 Hrecl1]; auto.
-intros l2; rewrite app_cons; intros; apply Hrecl1.
-apply (simpleDest_pop a); auto.
-Qed.
-
-Lemma simpleDest_swap:
- forall (m1 m2 : Move) (l : Moves),
- simpleDest (m1 :: (m2 :: l)) -> simpleDest (m2 :: (m1 :: l)).
-Proof.
-intros m1 m2 l; simpl; case m1; case m2.
-intros _ r0 _ r2 [[X Y] [Z U]]; auto.
-(repeat split); auto.
-apply Loc.diff_sym; auto.
-Qed.
-
-Lemma simpleDest_pop2:
- forall (m1 m2 : Move) (l : Moves),
- simpleDest (m1 :: (m2 :: l)) -> simpleDest (m1 :: l).
-Proof.
-intros; apply (simpleDest_pop m2); apply simpleDest_swap; auto.
-Qed.
-
-Lemma simpleDest_movFront:
- forall (p1 p2 : Moves) (m : Move),
- simpleDest (p1 ++ (m :: p2)) -> simpleDest (m :: (p1 ++ p2)).
-Proof.
-intros p1 p2 m; induction p1 as [|a p1 Hrecp1].
-simpl; auto.
-rewrite app_cons; rewrite app_cons.
-case a; intros; simpl in H |-; inversion H.
-apply simpleDest_swap; simpl; auto.
-destruct m.
-cut (noWrite ((r1, r2) :: (p1 ++ p2)) r0).
-cut (simpleDest ((r1, r2) :: (p1 ++ p2))).
-intro; (repeat split); elim H3; elim H2; intros; auto.
-apply Hrecp1; auto.
-apply noWrite_movFront; auto.
-Qed.
-
-Lemma simpleDest_insert:
- forall (p1 p2 : Moves) (m : Move),
- simpleDest (m :: (p1 ++ p2)) -> simpleDest (p1 ++ (m :: p2)).
-Proof.
-intros p1 p2 m; induction p1 as [|a p1 Hrecp1].
-simpl; auto.
-rewrite app_cons; intros.
-simpl.
-destruct a as [a1 a2].
-split.
-destruct m; simpl in H |-.
-apply noWrite_insert.
-simpl; split; elim H; intros [H1 H2] [H3 H4]; auto.
-apply Loc.diff_sym; auto.
-apply Hrecp1.
-apply simpleDest_pop2 with (a1, a2); auto.
-Qed.
-
-Lemma simpleDest_movBack:
- forall (p1 p2 : Moves) (m : Move),
- simpleDest (p1 ++ (m :: p2)) -> simpleDest ((p1 ++ p2) ++ (m :: nil)).
-Proof.
-intros.
-apply (simpleDest_insert (p1 ++ p2) nil m).
-rewrite app_nil; apply simpleDest_movFront; auto.
-Qed.
-
-Lemma simpleDest_swap_app:
- forall (t1 t2 t3 : Moves) (m : Move),
- simpleDest (t1 ++ (m :: (t2 ++ t3))) -> simpleDest ((t1 ++ t2) ++ (m :: t3)).
-Proof.
-intros.
-apply (simpleDest_insert (t1 ++ t2) t3 m).
-rewrite <- app_app.
-apply simpleDest_movFront; auto.
-Qed.
-
-Lemma simpleDest_tmpLast:
- forall (t : Moves) (s d : Reg),
- simpleDest (t ++ ((s, d) :: nil)) ->
- forall (r : Reg), simpleDest (t ++ ((r, d) :: nil)).
-Proof.
-intros t s d; induction t as [|a t Hrect].
-simpl; auto.
-simpl; case a; intros; inversion H; split; auto.
-apply (noWrite_tmpLast t r0 s); auto.
-Qed.
-
-Fixpoint noTmp (b : Moves) : Prop :=
- match b with
- nil => True
- | (s, d) :: l =>
- (forall r, Loc.diff s (T r)) /\
- ((forall r, Loc.diff d (T r)) /\ noTmp l)
- end.
-
-Fixpoint noTmpLast (b : Moves) : Prop :=
- match b with
- nil => True
- | (s, d) :: nil => forall r, Loc.diff d (T r)
- | (s, d) :: l =>
- (forall r, Loc.diff s (T r)) /\
- ((forall r, Loc.diff d (T r)) /\ noTmpLast l)
- end.
-
-Lemma noTmp_app:
- forall (l1 l2 : Moves) (m : Move),
- noTmp l1 -> noTmpLast (m :: l2) -> noTmpLast (l1 ++ (m :: l2)).
-Proof.
-intros.
-induction l1 as [|a l1 Hrecl1].
-simpl; auto.
-simpl.
-caseEq (l1 ++ (m :: l2)); intro.
-destruct a.
-elim H; intros; auto.
-inversion H; auto.
-elim H3; auto.
-intros; destruct a as [a1 a2].
-elim H; intros H2 [H3 H4]; auto.
-(repeat split); auto.
-rewrite H1 in Hrecl1; apply Hrecl1; auto.
-Qed.
-
-Lemma noTmpLast_popBack:
- forall (t : Moves) (m : Move), noTmpLast (t ++ (m :: nil)) -> noTmp t.
-Proof.
-intros.
-induction t as [|a t Hrect].
-simpl; auto.
-destruct a as [a1 a2].
-rewrite app_cons in H.
-simpl.
-simpl in H |-.
-generalize H; caseEq (t ++ (m :: nil)); intros.
-destruct t; inversion H0.
-elim H1.
-intros H2 [H3 H4]; (repeat split); auto.
-rewrite <- H0 in H4.
-apply Hrect; auto.
-Qed.
-
-Fixpoint getsrc (p : Moves) : list Reg :=
- match p with
- nil => nil
- | (s, d) :: l => s :: getsrc l
- end.
-
-Fixpoint getdst (p : Moves) : list Reg :=
- match p with
- nil => nil
- | (s, d) :: l => d :: getdst l
- end.
-
-Fixpoint noOverlap_aux (r : Reg) (l : list Reg) {struct l} : Prop :=
- match l with
- nil => True
- | b :: m => (b = r \/ Loc.diff b r) /\ noOverlap_aux r m
- end.
-
-Definition noOverlap (p : Moves) : Prop :=
- forall l, In l (getsrc p) -> noOverlap_aux l (getdst p).
-
-Definition stepInv (r : State) : Prop :=
- path (StateBeing r) /\
- (simpleDest (StateToMove r ++ StateBeing r) /\
- (noOverlap (StateToMove r ++ StateBeing r) /\
- (noTmp (StateToMove r) /\ noTmpLast (StateBeing r)))).
-
-Definition Value := val.
-
-Definition Env := locset.
-
-Definition get (e : Env) (r : Reg) := Locmap.get r e.
-
-Definition update (e : Env) (r : Reg) (v : Value) : Env := Locmap.set r v e.
-
-Fixpoint sexec (p : Moves) (e : Env) {struct p} : Env :=
- match p with
- nil => e
- | (s, d) :: l => let e' := sexec l e in
- update e' d (get e' s)
- end.
-
-Fixpoint pexec (p : Moves) (e : Env) {struct p} : Env :=
- match p with
- nil => e
- | (s, d) :: l => update (pexec l e) d (get e s)
- end.
-
-Lemma get_update:
- forall (e : Env) (r1 r2 : Reg) (v : Value),
- get (update e r1 v) r2 =
- (if Loc.eq r1 r2 then v else if Loc.overlap r1 r2 then Vundef else get e r2).
-Proof.
-intros.
-unfold update, get, Locmap.get, Locmap.set; trivial.
-Qed.
-
-Lemma get_update_id:
- forall (e : Env) (r1 : Reg) (v : Value), get (update e r1 v) r1 = v.
-Proof.
-intros e r1 v; rewrite (get_update e r1 r1); auto.
-case (Loc.eq r1 r1); auto.
-intros H; elim H; trivial.
-Qed.
-
-Lemma get_update_diff:
- forall (e : Env) (r1 r2 : Reg) (v : Value),
- Loc.diff r1 r2 -> get (update e r1 v) r2 = get e r2.
-Proof.
-intros; unfold update, get, Locmap.get, Locmap.set.
-case (Loc.eq r1 r2); intro.
-absurd (r1 = r2); [apply Loc.diff_not_eq; trivial | trivial].
-caseEq (Loc.overlap r1 r2); intro; trivial.
-absurd (Loc.diff r1 r2); [apply Loc.overlap_not_diff; assumption | assumption].
-Qed.
-
-Lemma get_update_ndiff:
- forall (e : Env) (r1 r2 : Reg) (v : Value),
- r1 <> r2 -> not (Loc.diff r1 r2) -> get (update e r1 v) r2 = Vundef.
-Proof.
-intros; unfold update, get, Locmap.get, Locmap.set.
-case (Loc.eq r1 r2); intro.
-absurd (r1 = r2); assumption.
-caseEq (Loc.overlap r1 r2); intro; trivial.
-absurd (Loc.diff r1 r2); (try assumption).
-apply Loc.non_overlap_diff; assumption.
-Qed.
-
-Lemma pexec_swap:
- forall (m1 m2 : Move) (t : Moves),
- simpleDest (m1 :: (m2 :: t)) ->
- forall (e : Env) (r : Reg),
- get (pexec (m1 :: (m2 :: t)) e) r = get (pexec (m2 :: (m1 :: t)) e) r.
-Proof.
-intros; destruct m1 as [m1s m1d]; destruct m2 as [m2s m2d].
-generalize H; simpl; intros [[NEQ NW] [NW2 HSD]]; clear H.
-case (Loc.eq m1d r); case (Loc.eq m2d r); intros.
-absurd (m1d = m2d);
- [apply Loc.diff_not_eq; apply Loc.diff_sym; assumption |
- rewrite e0; rewrite e1; trivial].
-caseEq (Loc.overlap m2d r); intro.
-absurd (Loc.diff m2d m1d); [apply Loc.overlap_not_diff; rewrite e0 | idtac];
- (try assumption).
-subst m1d; rewrite get_update_id; rewrite get_update_diff;
- (try rewrite get_update_id); auto.
-caseEq (Loc.overlap m1d r); intro.
-absurd (Loc.diff m1d m2d);
- [apply Loc.overlap_not_diff; rewrite e0 | apply Loc.diff_sym]; assumption.
-subst m2d; (repeat rewrite get_update_id); rewrite get_update_diff;
- [rewrite get_update_id; trivial | apply Loc.diff_sym; trivial].
-caseEq (Loc.overlap m1d r); caseEq (Loc.overlap m2d r); intros.
-(repeat rewrite get_update_ndiff);
- (try (apply Loc.overlap_not_diff; assumption)); trivial.
-assert (~ Loc.diff m1d r);
- [apply Loc.overlap_not_diff; assumption |
- intros; rewrite get_update_ndiff; auto].
-rewrite get_update_diff;
- [rewrite get_update_ndiff; auto | apply Loc.non_overlap_diff; auto].
-cut (~ Loc.diff m2d r); [idtac | apply Loc.overlap_not_diff; auto].
-cut (Loc.diff m1d r); [idtac | apply Loc.non_overlap_diff; auto].
-intros; rewrite get_update_diff; auto.
-(repeat rewrite get_update_ndiff); auto.
-cut (Loc.diff m1d r); [idtac | apply Loc.non_overlap_diff; auto].
-cut (Loc.diff m2d r); [idtac | apply Loc.non_overlap_diff; auto].
-intros; (repeat rewrite get_update_diff); auto.
-Qed.
-
-Lemma pexec_add:
- forall (t1 t2 : Moves) (r : Reg) (e : Env),
- get (pexec t1 e) r = get (pexec t2 e) r ->
- forall (a : Move), get (pexec (a :: t1) e) r = get (pexec (a :: t2) e) r.
-Proof.
-intros.
-case a.
-simpl.
-intros a1 a2.
-unfold get, update, Locmap.set, Locmap.get.
-case (Loc.eq a2 r); case (Loc.overlap a2 r); auto.
-Qed.
-
-Lemma pexec_movBack:
- forall (t1 t2 : Moves) (m : Move),
- simpleDest (m :: (t1 ++ t2)) ->
- forall (e : Env) (r : Reg),
- get (pexec (m :: (t1 ++ t2)) e) r = get (pexec (t1 ++ (m :: t2)) e) r.
-Proof.
-intros t1 t2 m; induction t1 as [|a t1 Hrect1].
-simpl; auto.
-rewrite app_cons.
-intros; rewrite pexec_swap; auto; rewrite app_cons; auto.
-apply pexec_add.
-apply Hrect1.
-apply (simpleDest_pop2 m a); auto.
-Qed.
-
-Lemma pexec_movFront:
- forall (t1 t2 : Moves) (m : Move),
- simpleDest (t1 ++ (m :: t2)) ->
- forall (e : Env) (r : Reg),
- get (pexec (t1 ++ (m :: t2)) e) r = get (pexec (m :: (t1 ++ t2)) e) r.
-Proof.
-intros; rewrite <- pexec_movBack; eauto.
-apply simpleDest_movFront; auto.
-Qed.
-
-Lemma pexec_mov:
- forall (t1 t2 t3 : Moves) (m : Move),
- simpleDest ((t1 ++ (m :: t2)) ++ t3) ->
- forall (e : Env) (r : Reg),
- get (pexec ((t1 ++ (m :: t2)) ++ t3) e) r =
- get (pexec ((t1 ++ t2) ++ (m :: t3)) e) r.
-Proof.
-intros t1 t2 t3 m.
-rewrite <- app_app.
-rewrite app_cons.
-intros.
-rewrite pexec_movFront; auto.
-cut (simpleDest (m :: (t1 ++ (t2 ++ t3)))).
-rewrite app_app.
-rewrite <- pexec_movFront; auto.
-apply simpleDest_swap_app; auto.
-apply simpleDest_movFront; auto.
-Qed.
-
-Definition diff_dec:
- forall (x y : Reg), ({ Loc.diff x y }) + ({ not (Loc.diff x y) }).
-intros.
-case (Loc.eq x y).
-intros heq; right.
-red; intro; absurd (x = y); auto.
-apply Loc.diff_not_eq; auto.
-intro; caseEq (Loc.overlap x y).
-intro; right.
-apply Loc.overlap_not_diff; auto.
-intro; left; apply Loc.non_overlap_diff; auto.
-Defined.
-
-Lemma get_pexec_id_noWrite:
- forall (t : Moves) (d : Reg),
- noWrite t d ->
- forall (e : Env) (v : Value), v = get (pexec t (update e d v)) d.
-Proof.
-intros.
-induction t as [|a t Hrect].
-simpl.
-rewrite get_update_id; auto.
-generalize H; destruct a as [a1 a2]; simpl; intros [NEQ R].
-rewrite get_update_diff; auto.
-Qed.
-
-Lemma pexec_nop:
- forall (t : Moves) (r : Reg) (e : Env) (x : Reg),
- Loc.diff r x -> get (pexec ((r, r) :: t) e) x = get (pexec t e) x.
-Proof.
-intros.
-simpl.
-rewrite get_update_diff; auto.
-Qed.
-
-Lemma sD_nW: forall t r s, simpleDest ((s, r) :: t) -> noWrite t r.
-Proof.
-induction t.
-simpl; auto.
-simpl.
-destruct a.
-intros r1 r2 H; split; [try assumption | idtac].
-elim H;
- [intros H0 H1; elim H0; [intros H2 H3; (try clear H0 H); (try exact H2)]].
-elim H;
- [intros H0 H1; elim H0; [intros H2 H3; (try clear H0 H); (try exact H3)]].
-Qed.
-
-Lemma sD_pexec:
- forall (t : Moves) (s d : Reg),
- simpleDest ((s, d) :: t) -> forall (e : Env), get (pexec t e) d = get e d.
-Proof.
-intros.
-induction t as [|a t Hrect]; simpl; auto.
-destruct a as [a1 a2].
-simpl in H |-; elim H; intros [H0 H1] [H2 H3]; clear H.
-case (Loc.eq a2 d); intro.
-absurd (a2 = d); [apply Loc.diff_not_eq | trivial]; assumption.
-rewrite get_update_diff; (try assumption).
-apply Hrect.
-simpl; (split; assumption).
-Qed.
-
-Lemma noOverlap_nil: noOverlap nil.
-Proof.
-unfold noOverlap, noOverlap_aux, getsrc, getdst; trivial.
-Qed.
-
-Lemma getsrc_add:
- forall (m : Move) (l1 l2 : Moves) (l : Reg),
- In l (getsrc (l1 ++ l2)) -> In l (getsrc (l1 ++ (m :: l2))).
-Proof.
-intros m l1 l2 l; destruct m; induction l1; simpl; auto.
-destruct a; simpl; intros.
-elim H; intros H0; [left | right]; auto.
-Qed.
-
-Lemma getdst_add:
- forall (r1 r2 : Reg) (l1 l2 : Moves),
- getdst (l1 ++ ((r1, r2) :: l2)) = getdst l1 ++ (r2 :: getdst l2).
-Proof.
-intros r1 r2 l1 l2; induction l1; simpl; auto.
-destruct a; simpl; rewrite IHl1; auto.
-Qed.
-
-Lemma getdst_app:
- forall (l1 l2 : Moves), getdst (l1 ++ l2) = getdst l1 ++ getdst l2.
-Proof.
-intros; induction l1; simpl; auto.
-destruct a; simpl; rewrite IHl1; auto.
-Qed.
-
-Lemma noOverlap_auxpop:
- forall (x r : Reg) (l : list Reg),
- noOverlap_aux x (r :: l) -> noOverlap_aux x l.
-Proof.
-induction l; simpl; auto.
-intros [H1 [H2 H3]]; split; auto.
-Qed.
-
-Lemma noOverlap_auxPop:
- forall (x r : Reg) (l1 l2 : list Reg),
- noOverlap_aux x (l1 ++ (r :: l2)) -> noOverlap_aux x (l1 ++ l2).
-Proof.
-intros x r l1 l2; (try assumption).
-induction l1 as [|a l1 Hrecl1]; simpl app.
-intro; apply (noOverlap_auxpop x r); auto.
-(repeat rewrite app_cons); simpl.
-intros [H1 H2]; split; auto.
-Qed.
-
-Lemma noOverlap_pop:
- forall (m : Move) (l : Moves), noOverlap (m :: l) -> noOverlap l.
-Proof.
-induction l.
-intro; apply noOverlap_nil.
-unfold noOverlap; simpl; destruct m; destruct a; simpl; intros.
-elim (H l0); intros; (try assumption).
-elim H0; intros H1; right; [left | right]; assumption.
-Qed.
-
-Lemma noOverlap_Pop:
- forall (m : Move) (l1 l2 : Moves),
- noOverlap (l1 ++ (m :: l2)) -> noOverlap (l1 ++ l2).
-Proof.
-intros m l1 l2; induction l1 as [|a l1 Hrecl1]; simpl.
-simpl; apply noOverlap_pop.
-(repeat rewrite app_cons); unfold noOverlap; destruct a; simpl.
-intros H l H0; split.
-elim (H l); [intros H1 H2 | idtac]; auto.
-elim H0; [intros H3; left | intros H3; right; apply getsrc_add]; auto.
-unfold noOverlap in Hrecl1 |-.
-elim H0; intros H1; clear H0.
-destruct m; rewrite getdst_app; apply noOverlap_auxPop with ( r := r2 ).
-rewrite getdst_add in H.
-elim H with ( l := l ); [intros H0 H2; (try clear H); (try exact H2) | idtac].
-left; (try assumption).
-apply Hrecl1 with ( l := l ); auto.
-intros l0 H0; (try assumption).
-elim H with ( l := l0 ); [intros H2 H3; (try clear H); (try exact H3) | idtac];
- auto.
-Qed.
-
-Lemma noOverlap_right:
- forall (l1 l2 : Moves), noOverlap (l1 ++ l2) -> noOverlap l2.
-Proof.
-intros l1; induction l1 as [|a l1 Hrecl1]; auto.
-intros l2; rewrite app_cons; intros; apply Hrecl1.
-apply (noOverlap_pop a); auto.
-Qed.
-
-Lemma pexec_update:
- forall t e d r v,
- Loc.diff d r ->
- noRead t d -> get (pexec t (update e d v)) r = get (pexec t e) r.
-Proof.
-induction t; simpl.
-intros; rewrite get_update_diff; auto.
-destruct a as [a1 a2]; intros; case (Loc.eq a2 r); intro.
-subst a2; (repeat rewrite get_update_id).
-rewrite get_update_diff; auto; apply Loc.diff_sym; elim H0; auto.
-case (diff_dec a2 r); intro.
-(repeat rewrite get_update_diff); auto.
-apply IHt; auto.
-elim H0; auto.
-(repeat rewrite get_update_ndiff); auto.
-Qed.
-
-Lemma pexec_push:
- forall (t l : Moves) (s d : Reg),
- noRead t d ->
- simpleDest ((s, d) :: t) ->
- forall (e : Env) (r : Reg),
- r = d \/ Loc.diff d r ->
- get (pexec ((s, d) :: t) (sexec l e)) r =
- get (pexec t (sexec ((s, d) :: l) e)) r.
-Proof.
-intros; simpl.
-elim H1; intros e1.
-rewrite e1; rewrite get_update_id; auto.
-rewrite (sD_pexec t s d); auto; rewrite get_update_id; auto.
-rewrite pexec_update; auto.
-rewrite get_update_diff; auto.
-Qed.
-
-Definition exec (s : State) (e : Env) :=
- pexec (StateToMove s ++ StateBeing s) (sexec (StateDone s) e).
-
-Definition sameEnv (e1 e2 : Env) :=
- forall (r : Reg), notemporary r -> get e1 r = get e2 r.
-
-Definition NoOverlap (r : Reg) (s : State) :=
- noOverlap ((r, r) :: (StateToMove s ++ StateBeing s)).
-
-Lemma noOverlapaux_swap2:
- forall (l1 l2 : list Reg) (m l : Reg),
- noOverlap_aux l (l1 ++ (m :: l2)) -> noOverlap_aux l (m :: (l1 ++ l2)).
-Proof.
-intros l1 l2 m l; induction l1; simpl noOverlap_aux; auto.
-intros; elim H; intros H0 H1; (repeat split); auto.
-simpl in IHl1 |-.
-elim IHl1; [intros H2 H3; (try exact H2) | idtac]; auto.
-apply (noOverlap_auxpop l m).
-apply IHl1; auto.
-Qed.
-
-Lemma noTmp_noReadTmp: forall t, noTmp t -> forall s, noRead t (T s).
-Proof.
-induction t; simpl; auto.
-destruct a as [a1 a2]; intros.
-split; [idtac | apply IHt]; elim H; intros H1 [H2 H3]; auto.
-Qed.
-
-Lemma noRead_by_path:
- forall (b t : Moves) (r0 r1 r7 r8 : Reg),
- simpleDest ((r7, r8) :: (b ++ ((r0, r1) :: nil))) ->
- path (b ++ ((r0, r1) :: nil)) -> Loc.diff r8 r0 -> noRead b r8.
-Proof.
-intros; induction b as [|a b Hrecb]; simpl; auto.
-destruct a as [a1 a2]; generalize H H0; rewrite app_cons; intros; split.
-simpl in H3 |-; caseEq (b ++ ((r0, r1) :: nil)); intro.
-destruct b; inversion H4.
-intros l H4.
-rewrite H4 in H3.
-destruct m.
-rewrite H4 in H2; simpl in H2 |-.
-elim H3; [intros H5 H6; (try clear H3); (try exact H5)].
-rewrite H5.
-elim H2; intros [H3 [H7 H8]] [H9 [H10 H11]]; (try assumption).
-apply Hrecb.
-apply (simpleDest_pop (a1, a2)); apply simpleDest_swap; auto.
-apply (path_pop (a1, a2)); auto.
-Qed.
-
-Lemma noOverlap_swap:
- forall (m1 m2 : Move) (l : Moves),
- noOverlap (m1 :: (m2 :: l)) -> noOverlap (m2 :: (m1 :: l)).
-Proof.
-intros m1 m2 l; simpl; destruct m1 as [m1s m1d]; destruct m2 as [m2s m2d].
-unfold noOverlap; simpl; intros.
-assert (m1s = l0 \/ (m2s = l0 \/ In l0 (getsrc l))).
-elim H0; [intros H1 | intros [H1|H2]].
-right; left; (try assumption).
-left; (try assumption).
-right; right; (try assumption).
-(repeat split);
- (elim (H l0); [intros H2 H3; elim H3; [intros H4 H5] | idtac]; auto).
-Qed.
-
-Lemma getsrc_add1:
- forall (r1 r2 : Reg) (l1 l2 : Moves),
- getsrc (l1 ++ ((r1, r2) :: l2)) = getsrc l1 ++ (r1 :: getsrc l2).
-Proof.
-intros r1 r2 l1 l2; induction l1; simpl; auto.
-destruct a; simpl; rewrite IHl1; auto.
-Qed.
-
-Lemma getsrc_app:
- forall (l1 l2 : Moves), getsrc (l1 ++ l2) = getsrc l1 ++ getsrc l2.
-Proof.
-intros; induction l1; simpl; auto.
-destruct a; simpl; rewrite IHl1; auto.
-Qed.
-
-Lemma Ingetsrc_swap:
- forall (m : Move) (l1 l2 : Moves) (l : Reg),
- In l (getsrc (m :: (l1 ++ l2))) -> In l (getsrc (l1 ++ (m :: l2))).
-Proof.
-intros; destruct m as [m1 m2]; simpl; auto.
-simpl in H |-.
-elim H; intros H0; auto.
-rewrite H0; rewrite getsrc_add1; auto.
-apply (in_or_app (getsrc l1) (l :: getsrc l2)); auto.
-right; apply in_eq; auto.
-apply getsrc_add; auto.
-Qed.
-
-Lemma noOverlap_movFront:
- forall (p1 p2 : Moves) (m : Move),
- noOverlap (p1 ++ (m :: p2)) -> noOverlap (m :: (p1 ++ p2)).
-Proof.
-intros p1 p2 m; unfold noOverlap.
-destruct m; rewrite getdst_add; simpl getdst; rewrite getdst_app; intros.
-apply noOverlapaux_swap2.
-apply (H l); apply Ingetsrc_swap; auto.
-Qed.
-
-Lemma step_inv_loop_aux:
- forall (t l : Moves) (s d : Reg),
- simpleDest (t ++ ((s, d) :: nil)) ->
- noTmp t ->
- forall (e : Env) (r : Reg),
- notemporary r ->
- d = r \/ Loc.diff d r ->
- get (pexec (t ++ ((s, d) :: nil)) (sexec l e)) r =
- get (pexec (t ++ ((T s, d) :: nil)) (sexec ((s, T s) :: l) e)) r.
-Proof.
-intros; (repeat rewrite pexec_movFront); auto.
-(repeat rewrite app_nil); simpl; elim H2; intros e1.
-subst d; (repeat rewrite get_update_id); auto.
-(repeat rewrite get_update_diff); auto.
-rewrite pexec_update; auto.
-apply Loc.diff_sym; unfold notemporary in H1 |-; auto.
-apply noTmp_noReadTmp; auto.
-apply (simpleDest_tmpLast t s); auto.
-Qed.
-
-Lemma step_inv_loop:
- forall (t l : Moves) (s d : Reg),
- simpleDest (t ++ ((s, d) :: nil)) ->
- noTmpLast (t ++ ((s, d) :: nil)) ->
- forall (e : Env) (r : Reg),
- notemporary r ->
- d = r \/ Loc.diff d r ->
- get (pexec (t ++ ((s, d) :: nil)) (sexec l e)) r =
- get (pexec (t ++ ((T s, d) :: nil)) (sexec ((s, T s) :: l) e)) r.
-Proof.
-intros; apply step_inv_loop_aux; auto.
-apply (noTmpLast_popBack t (s, d)); auto.
-Qed.
-
-Definition sameExec (s1 s2 : State) :=
- forall (e : Env) (r : Reg),
- (let A :=
- getdst
- ((StateToMove s1 ++ StateBeing s1) ++ (StateToMove s2 ++ StateBeing s2))
- in
- notemporary r ->
- (forall x, In x A -> r = x \/ Loc.diff r x) ->
- get (exec s1 e) r = get (exec s2 e) r).
-
-Lemma get_noWrite:
- forall (t : Moves) (d : Reg),
- noWrite t d -> forall (e : Env), get e d = get (pexec t e) d.
-Proof.
-intros; induction t as [|a t Hrect]; simpl; auto.
-generalize H; destruct a as [a1 a2]; simpl; intros [NEQ R].
-unfold get, Locmap.get, update, Locmap.set.
-case (Loc.eq a2 d); intro; auto.
-absurd (a2 = d); auto; apply Loc.diff_not_eq; (try assumption).
-caseEq (Loc.overlap a2 d); intro.
-absurd (Loc.diff a2 d); auto; apply Loc.overlap_not_diff; auto.
-unfold get, Locmap.get in Hrect |-; apply Hrect; auto.
-Qed.
-
-Lemma step_sameExec:
- forall (r1 r2 : State), step r1 r2 -> stepInv r1 -> sameExec r1 r2.
-Proof.
-intros r1 r2 STEP; inversion STEP;
- unfold stepInv, sameExec, NoOverlap, exec, StateToMove, StateBeing, StateDone;
- (repeat rewrite app_nil); intros [P [SD [NO [TT TB]]]]; intros.
-rewrite pexec_movFront; simpl; auto.
-case (Loc.eq r r0); intros e0.
-subst r0; rewrite get_update_id; apply get_noWrite; apply sD_nW with r;
- apply simpleDest_movFront; auto.
-elim H2 with ( x := r );
- [intros H3; absurd (r = r0); auto |
- intros H3; rewrite get_update_diff; auto; apply Loc.diff_sym; auto | idtac].
-(repeat (rewrite getdst_app; simpl)); apply in_or_app; left; apply in_or_app;
- right; simpl; auto.
-(repeat rewrite pexec_movFront); auto.
-rewrite app_nil; auto.
-apply simpleDest_movBack; auto.
-apply pexec_mov; auto.
-repeat (rewrite <- app_cons; rewrite app_app).
-apply step_inv_loop; auto.
-repeat (rewrite <- app_app; rewrite app_cons; auto).
-repeat (rewrite <- app_app; rewrite app_cons; auto).
-apply noTmp_app; auto.
-elim H2 with ( x := d );
- [intros H3; left; auto | intros H3; right; apply Loc.diff_sym; auto
- | try clear H2].
-repeat (rewrite getdst_app; simpl).
-apply in_or_app; left; apply in_or_app; right; simpl; right; apply in_or_app;
- right; simpl; left; trivial.
-rewrite pexec_movFront; auto; apply pexec_push; auto.
-apply noRead_app; auto.
-apply noRead_app.
-apply (noRead_by_path b b s0 d0 sn dn); auto.
-apply (simpleDest_right t); auto.
-apply (path_pop (sn, dn)); auto.
-simpl; split; [apply Loc.diff_sym | idtac]; auto.
-apply simpleDest_movFront; auto.
-elim H4 with ( x := dn ); [intros H5 | intros H5 | try clear H4].
-left; (try assumption).
-right; apply Loc.diff_sym; (try assumption).
-repeat (rewrite getdst_app; simpl).
-apply in_or_app; left; apply in_or_app; right; simpl; left; trivial.
-rewrite pexec_movFront; auto.
-rewrite app_nil; auto.
-apply pexec_push; auto.
-rewrite <- (app_nil _ t).
-apply simpleDest_movFront; auto.
-elim (H3 d); (try intros H4).
-left; (try assumption).
-right; apply Loc.diff_sym; (try assumption).
-(repeat rewrite getdst_app); simpl; apply in_or_app; left; apply in_or_app;
- right; simpl; left; trivial.
-Qed.
-
-Lemma path_tmpLast:
- forall (s d : Reg) (l : Moves),
- path (l ++ ((s, d) :: nil)) -> path (l ++ ((T s, d) :: nil)).
-Proof.
-intros; induction l as [|a l Hrecl].
-simpl; auto.
-generalize H; (repeat rewrite app_cons).
-case a; generalize Hrecl; case l; intros; auto.
-destruct m; intros.
-inversion H0; split; auto.
-Qed.
-
-Lemma step_inv_path:
- forall (r1 r2 : State), step r1 r2 -> stepInv r1 -> path (StateBeing r2).
-Proof.
-intros r1 r2 STEP; inversion_clear STEP; unfold stepInv;
- unfold stepInv, sameExec, sameEnv, exec, StateToMove, StateBeing, StateDone;
- intros [P [SD [TT TB]]]; (try (simpl; auto; fail)).
-simpl; case m; auto.
-generalize P; rewrite <- app_cons; rewrite <- app_cons.
-apply (path_tmpLast r0).
-generalize P; apply path_pop.
-Qed.
-
-Lemma step_inv_simpleDest:
- forall (r1 r2 : State),
- step r1 r2 -> stepInv r1 -> simpleDest (StateToMove r2 ++ StateBeing r2).
-Proof.
-intros r1 r2 STEP; inversion_clear STEP; unfold stepInv;
- unfold stepInv, sameExec, sameEnv, exec, StateToMove, StateBeing, StateDone;
- (repeat rewrite app_nil); intros [P [SD [TT TB]]].
-apply (simpleDest_Pop (r, r)); assumption.
-apply simpleDest_movBack; assumption.
-apply simpleDest_insert; rewrite <- app_app; apply simpleDest_movFront.
-rewrite <- app_cons; rewrite app_app; auto.
-generalize SD; (repeat rewrite <- app_cons); (repeat rewrite app_app).
-generalize (simpleDest_tmpLast (t ++ ((s, r0ounon) :: b)) r0 d); auto.
-generalize SD; apply simpleDest_Pop.
-rewrite <- (app_nil _ t); generalize SD; apply simpleDest_Pop.
-Qed.
-
-Lemma noTmp_pop:
- forall (m : Move) (l1 l2 : Moves), noTmp (l1 ++ (m :: l2)) -> noTmp (l1 ++ l2).
-Proof.
-intros; induction l1 as [|a l1 Hrecl1]; generalize H.
-simpl; case m; intros; inversion H0; inversion H2; auto.
-rewrite app_cons; rewrite app_cons; simpl; case a.
-intros; inversion H0; inversion H2; auto.
-Qed.
-
-Lemma step_inv_noTmp:
- forall (r1 r2 : State), step r1 r2 -> stepInv r1 -> noTmp (StateToMove r2).
-Proof.
-intros r1 r2 STEP; inversion_clear STEP; unfold stepInv;
- unfold stepInv, sameExec, sameEnv, exec, StateToMove, StateBeing, StateDone;
- intros [P [SD [NO [TT TB]]]]; generalize TT; (try apply noTmp_pop); auto.
-Qed.
-
-Lemma noTmp_noTmpLast: forall (l : Moves), noTmp l -> noTmpLast l.
-Proof.
-intros; induction l as [|a l Hrecl]; (try (simpl; auto; fail)).
-generalize H; simpl; case a; generalize Hrecl; case l;
- (intros; inversion H0; inversion H2; auto).
-Qed.
-
-Lemma noTmpLast_pop:
- forall (m : Move) (l : Moves), noTmpLast (m :: l) -> noTmpLast l.
-Proof.
-intros m l; simpl; case m; case l.
-simpl; auto.
-intros; inversion H; inversion H1; auto.
-Qed.
-
-Lemma noTmpLast_Pop:
- forall (m : Move) (l1 l2 : Moves),
- noTmpLast (l1 ++ (m :: l2)) -> noTmpLast (l1 ++ l2).
-Proof.
-intros; induction l1 as [|a l1 Hrecl1]; generalize H.
-simpl; case m; case l2.
-simpl; auto.
-intros.
-elim H0; [intros H1 H2; elim H2; [intros H3 H4; (try exact H4)]].
-(repeat rewrite app_cons); simpl; case a.
-generalize Hrecl1; case l1.
-simpl; case m; case l2; intros; inversion H0; inversion H2; auto.
-intros m0 l R r r0; rewrite app_cons; rewrite app_cons.
-intros; inversion H0; inversion H2; auto.
-Qed.
-
-Lemma noTmpLast_push:
- forall (m : Move) (t1 t2 t3 : Moves),
- noTmp (t1 ++ (m :: t2)) -> noTmpLast t3 -> noTmpLast (m :: t3).
-Proof.
-intros; induction t1 as [|a t1 Hrect1]; generalize H.
-simpl; case m; intros r r0 [N1 [N2 NT]]; generalize H0; case t3; auto.
-rewrite app_cons; intros; apply Hrect1.
-generalize H1.
-simpl; case m; case a; intros; inversion H2; inversion H4; auto.
-Qed.
-
-Lemma noTmpLast_tmpLast:
- forall (s d : Reg) (l : Moves),
- noTmpLast (l ++ ((s, d) :: nil)) -> noTmpLast (l ++ ((T s, d) :: nil)).
+Definition exec_seq (m: moves) (e: Locmap.t) : Locmap.t :=
+ Parmov.exec_seq loc val Loc.eq m e.
+
+Lemma temp_for_charact:
+ forall l, temp_for l = R IT2 \/ temp_for l = R FT2.
Proof.
-intros; induction l as [|a l Hrecl].
-simpl; auto.
-generalize H; rewrite app_cons; rewrite app_cons; simpl.
-case a; generalize Hrecl; case l.
-simpl; auto.
-intros m l0 REC r r0; generalize REC; rewrite app_cons; rewrite app_cons.
-case m; intros; inversion H0; inversion H2; split; auto.
+ intro; unfold temp_for. destruct (Loc.type l); tauto.
Qed.
-
-Lemma step_inv_noTmpLast:
- forall (r1 r2 : State), step r1 r2 -> stepInv r1 -> noTmpLast (StateBeing r2).
-Proof.
-intros r1 r2 STEP; inversion_clear STEP; unfold stepInv;
- unfold stepInv, sameExec, sameEnv, exec, StateToMove, StateBeing, StateDone;
- intros [P [SD [NO [TT TB]]]]; auto.
-apply (noTmpLast_push m t1 t2); auto.
-apply (noTmpLast_push (d, r) t1 t2); auto.
-generalize TB; rewrite <- app_cons; rewrite <- app_cons; apply noTmpLast_tmpLast.
-apply (noTmpLast_pop (sn, dn)); auto.
-Qed.
-
-Lemma noOverlapaux_insert:
- forall (l1 l2 : list Reg) (r x : Reg),
- noOverlap_aux x (r :: (l1 ++ l2)) -> noOverlap_aux x (l1 ++ (r :: l2)).
-Proof.
-simpl; intros; induction l1; simpl; split.
-elim H; [intros H0 H1; (try exact H0)].
-elim H; [intros H0 H1; (try exact H1)].
-simpl in H |-.
-elim H;
- [intros H0 H1; elim H1; [intros H2 H3; (try clear H1 H); (try exact H2)]].
-apply IHl1.
-split.
-elim H; [intros H0 H1; (try exact H0)].
-rewrite app_cons in H.
-apply noOverlap_auxpop with ( r := a ).
-elim H; [intros H0 H1; (try exact H1)].
-Qed.
-
-Lemma Ingetsrc_swap2:
- forall (m : Move) (l1 l2 : Moves) (l : Reg),
- In l (getsrc (l1 ++ (m :: l2))) -> In l (getsrc (m :: (l1 ++ l2))).
-Proof.
-intros; destruct m as [m1 m2]; simpl; auto.
-induction l1; simpl.
-simpl in H |-; auto.
-destruct a; simpl.
-simpl in H |-.
-elim H; [intros H0 | intros H0; (try exact H0)].
-right; left; (try assumption).
-elim IHl1; intros; auto.
-Qed.
-
-Lemma noOverlap_insert:
- forall (p1 p2 : Moves) (m : Move),
- noOverlap (m :: (p1 ++ p2)) -> noOverlap (p1 ++ (m :: p2)).
-Proof.
-unfold noOverlap; destruct m; rewrite getdst_add; simpl getdst;
- rewrite getdst_app.
-intros.
-apply noOverlapaux_insert.
-generalize (H l); intros H1; lapply H1;
- [intros H2; (try clear H1); (try exact H2) | idtac].
-simpl getsrc.
-generalize (Ingetsrc_swap2 (r, r0)); simpl; (intros; auto).
-Qed.
-
-Lemma noOverlap_movBack:
- forall (p1 p2 : Moves) (m : Move),
- noOverlap (p1 ++ (m :: p2)) -> noOverlap ((p1 ++ p2) ++ (m :: nil)).
-Proof.
-intros.
-apply (noOverlap_insert (p1 ++ p2) nil m).
-rewrite app_nil; apply noOverlap_movFront; auto.
-Qed.
-
-Lemma noOverlap_movBack0:
- forall (t : Moves) (s d : Reg),
- noOverlap ((s, d) :: t) -> noOverlap (t ++ ((s, d) :: nil)).
-Proof.
-intros t s d H; (try assumption).
-apply noOverlap_insert.
-rewrite app_nil; auto.
-Qed.
-
-Lemma noOverlap_Front0:
- forall (t : Moves) (s d : Reg),
- noOverlap (t ++ ((s, d) :: nil)) -> noOverlap ((s, d) :: t).
-Proof.
-intros t s d H; (try assumption).
-cut ((s, d) :: t = (s, d) :: (t ++ nil)).
-intros e; rewrite e.
-apply noOverlap_movFront; auto.
-rewrite app_nil; auto.
-Qed.
-
-Lemma noTmpL_diff:
- forall (t : Moves) (s d : Reg),
- noTmpLast (t ++ ((s, d) :: nil)) -> notemporary d.
-Proof.
-intros t s d; unfold notemporary; induction t; (try (simpl; intros; auto; fail)).
-rewrite app_cons.
-intros; apply IHt.
-apply (noTmpLast_pop a); auto.
-Qed.
-
-Lemma noOverlap_aux_app:
- forall l1 l2 (r : Reg),
- noOverlap_aux r l1 -> noOverlap_aux r l2 -> noOverlap_aux r (l1 ++ l2).
-Proof.
-induction l1; simpl; auto.
-intros; split.
-elim H; [intros H1 H2; (try clear H); (try exact H1)].
-apply IHl1; auto.
-elim H; [intros H1 H2; (try clear H); (try exact H2)].
-Qed.
-
-Lemma noTmP_noOverlap_aux:
- forall t (r : Reg), noTmp t -> noOverlap_aux (T r) (getdst t).
-Proof.
-induction t; simpl; auto.
-destruct a; simpl; (intros; split).
-elim H; intros; elim H1; intros.
-right; apply H2.
-apply IHt; auto.
-elim H;
- [intros H0 H1; elim H1; [intros H2 H3; (try clear H1 H); (try exact H3)]].
-Qed.
-
-Lemma noTmp_append: forall l1 l2, noTmp l1 -> noTmp l2 -> noTmp (l1 ++ l2).
-Proof.
-induction l1; simpl; auto.
-destruct a.
-intros l2 [H1 [H2 H3]] H4.
-(repeat split); auto.
-Qed.
-
-Lemma step_inv_noOverlap:
- forall (r1 r2 : State),
- step r1 r2 -> stepInv r1 -> noOverlap (StateToMove r2 ++ StateBeing r2).
-Proof.
-intros r1 r2 STEP; inversion_clear STEP; unfold stepInv;
- unfold stepInv, sameExec, sameEnv, exec, StateToMove, StateBeing, StateDone;
- (repeat rewrite app_nil); intros [P [SD [NO [TT TB]]]];
- (try (generalize NO; apply noOverlap_Pop; auto; fail)).
-apply noOverlap_movBack; auto.
-apply noOverlap_insert; rewrite <- app_app; apply noOverlap_movFront;
- rewrite <- app_cons; rewrite app_app; auto.
-generalize NO; (repeat rewrite <- app_cons); (repeat rewrite app_app);
- (clear NO; intros NO); apply noOverlap_movBack0.
-assert (noOverlap ((r0, d) :: (t ++ ((s, r0ounon) :: b))));
- [apply noOverlap_Front0; auto | idtac].
-generalize H; unfold noOverlap; simpl; clear H; intros.
-elim H0; intros; [idtac | apply (H l0); (right; (try assumption))].
-split; [right; (try assumption) | idtac].
-generalize TB; simpl; caseEq (b ++ ((r0, d) :: nil)); intro.
-elim (app_eq_nil b ((r0, d) :: nil)); intros; auto; inversion H4.
-subst l0; intros; rewrite <- H1 in TB0.
-elim TB0; [intros H2 H3; elim H3; [intros H4 H5; (try clear H3 TB0)]].
-generalize (noTmpL_diff b r0 d); unfold notemporary; intro; apply H3; auto.
-rewrite <- H1; apply noTmP_noOverlap_aux; apply noTmp_append; auto;
- rewrite <- app_cons in TB; apply noTmpLast_popBack with (r0, d); auto.
-rewrite <- (app_nil _ t); apply (noOverlap_Pop (s, d)); assumption.
-Qed.
-
-Lemma step_inv: forall (r1 r2 : State), step r1 r2 -> stepInv r1 -> stepInv r2.
-Proof.
-intros; unfold stepInv; (repeat split).
-apply (step_inv_path r1 r2); auto.
-apply (step_inv_simpleDest r1 r2); auto.
-apply (step_inv_noOverlap r1 r2); auto.
-apply (step_inv_noTmp r1 r2); auto.
-apply (step_inv_noTmpLast r1 r2); auto.
-Qed.
-
-Definition step_NF (r : State) : Prop := ~ (exists s : State , step r s ).
-
-Inductive stepp : State -> State -> Prop :=
- stepp_refl: forall (r : State), stepp r r
- | stepp_trans:
- forall (r1 r2 r3 : State), step r1 r2 -> stepp r2 r3 -> stepp r1 r3 .
-Hint Resolve stepp_refl stepp_trans .
-
-Lemma stepp_transitive:
- forall (r1 r2 r3 : State), stepp r1 r2 -> stepp r2 r3 -> stepp r1 r3.
-Proof.
-intros; induction H as [r|r1 r2 r0 H H1 HrecH]; eauto.
-Qed.
-
-Lemma step_stepp: forall (s1 s2 : State), step s1 s2 -> stepp s1 s2.
-Proof.
-eauto.
-Qed.
-
-Lemma stepp_inv:
- forall (r1 r2 : State), stepp r1 r2 -> stepInv r1 -> stepInv r2.
-Proof.
-intros; induction H as [r|r1 r2 r3 H H1 HrecH]; auto.
-apply HrecH; auto.
-apply (step_inv r1 r2); auto.
-Qed.
-
-Lemma noTmpLast_lastnoTmp:
- forall l s d, noTmpLast (l ++ ((s, d) :: nil)) -> notemporary d.
-Proof.
-induction l.
-simpl.
-intros; unfold notemporary; auto.
-destruct a as [a1 a2]; intros.
-change (noTmpLast ((a1, a2) :: (l ++ ((s, d) :: nil)))) in H |-.
-apply IHl with s.
-apply noTmpLast_pop with (a1, a2); auto.
+
+Lemma is_not_temp_charact:
+ forall l,
+ Parmov.is_not_temp loc temp_for l <-> l <> R IT2 /\ l <> R FT2.
+Proof.
+ intros. unfold Parmov.is_not_temp.
+ destruct (Loc.eq l (R IT2)).
+ subst l. intuition. apply (H (R IT2)). reflexivity. discriminate.
+ destruct (Loc.eq l (R FT2)).
+ subst l. intuition. apply (H (R FT2)). reflexivity.
+ assert (forall d, l <> temp_for d).
+ intro. elim (temp_for_charact d); congruence.
+ intuition.
Qed.
-
-Lemma step_inv_NoOverlap:
- forall (s1 s2 : State) r,
- step s1 s2 -> notemporary r -> stepInv s1 -> NoOverlap r s1 -> NoOverlap r s2.
+
+Lemma disjoint_temp_not_temp:
+ forall l, Loc.notin l temporaries -> Parmov.is_not_temp loc temp_for l.
Proof.
-intros s1 s2 r STEP notempr; inversion_clear STEP; unfold stepInv;
- unfold stepInv, sameExec, sameEnv, exec, StateToMove, StateBeing, StateDone;
- intros [P [SD [NO [TT TB]]]]; unfold NoOverlap; simpl.
-simpl; (repeat rewrite app_nil); simpl; (repeat rewrite <- app_cons); intro;
- apply noOverlap_Pop with ( m := (r0, r0) ); auto.
-(repeat rewrite app_nil); simpl; rewrite app_ass; (repeat rewrite <- app_cons);
- intro; rewrite ass_app; apply noOverlap_movBack; auto.
-simpl; (repeat (rewrite app_ass; simpl)); (repeat rewrite <- app_cons); intro.
-rewrite ass_app; apply noOverlap_insert; rewrite app_ass;
- apply noOverlap_movFront; auto.
-simpl; (repeat rewrite <- app_cons); intro; rewrite ass_app;
- apply noOverlap_movBack0; auto.
-generalize H; (repeat (rewrite app_ass; simpl)); intro.
-assert (noOverlap ((r0, d) :: (((r, r) :: t) ++ ((s, r0ounon) :: b))));
- [apply noOverlap_Front0 | idtac]; auto.
-generalize H0; (repeat (rewrite app_ass; simpl)); auto.
-generalize H1; unfold noOverlap; simpl; intros.
-elim H3; intros H4; clear H3.
-split.
-right; assert (notemporary d).
-change (noTmpLast (((s, r0ounon) :: b) ++ ((r0, d) :: nil))) in TB |-;
- apply (noTmpLast_lastnoTmp ((s, r0ounon) :: b) r0); auto.
-rewrite <- H4; unfold notemporary in H3 |-; apply H3.
-split.
-right; rewrite <- H4; unfold notemporary in notempr |-; apply notempr.
-rewrite <- H4; apply noTmP_noOverlap_aux; auto.
-apply noTmp_append; auto.
-change (noTmpLast (((s, r0ounon) :: b) ++ ((r0, d) :: nil))) in TB |-;
- apply noTmpLast_popBack with ( m := (r0, d) ); auto.
-apply (H2 l0).
-elim H4; intros H3; right; [left | right]; assumption.
-intro;
- change (noOverlap (((r, r) :: t) ++ ((sn, dn) :: (b ++ ((s0, d0) :: nil))))) in
- H1 |-.
-change (noOverlap (((r, r) :: t) ++ (b ++ ((s0, d0) :: nil))));
- apply (noOverlap_Pop (sn, dn)); auto.
-(repeat rewrite <- app_cons); apply noOverlap_Pop.
+ intros. rewrite is_not_temp_charact.
+ unfold temporaries in H; simpl in H.
+ split; apply Loc.diff_not_eq; tauto.
Qed.
-
-Lemma step_inv_getdst:
- forall (s1 s2 : State) r,
- step s1 s2 ->
- In r (getdst (StateToMove s2 ++ StateBeing s2)) ->
- In r (getdst (StateToMove s1 ++ StateBeing s1)).
+
+Lemma loc_norepet_norepet:
+ forall l, Loc.norepet l -> list_norepet l.
Proof.
-intros s1 s2 r STEP; inversion_clear STEP;
- unfold StateToMove, StateBeing, StateDone.
-(repeat rewrite getdst_app); simpl; (repeat rewrite app_nil); intro;
- apply in_or_app.
-elim (in_app_or (getdst t1) (getdst t2) r); auto.
-intro; right; simpl; right; assumption.
-(repeat rewrite getdst_app); destruct m as [m1 m2]; simpl;
- (repeat rewrite app_nil); intro; apply in_or_app.
-elim (in_app_or (getdst t1 ++ getdst t2) (m2 :: nil) r); auto; intro.
-elim (in_app_or (getdst t1) (getdst t2) r); auto; intro.
-right; simpl; right; assumption.
-elim H0; intros H1; [right; simpl; left; (try assumption) | inversion H1].
-(repeat rewrite getdst_app); simpl; (repeat rewrite app_nil); intro;
- apply in_or_app.
-elim (in_app_or (getdst t1 ++ getdst t2) (r0 :: (d :: getdst b)) r); auto;
- intro.
-elim (in_app_or (getdst t1) (getdst t2) r); auto; intro.
-left; apply in_or_app; left; assumption.
-left; apply in_or_app; right; simpl; right; assumption.
-elim H0; intro.
-left; apply in_or_app; right; simpl; left; trivial.
-elim H1; intro.
-right; (simpl; left; trivial).
-right; simpl; right; assumption.
-(repeat (rewrite getdst_app; simpl)); trivial.
-(repeat (rewrite getdst_app; simpl)); intro.
-elim (in_app_or (getdst t) (getdst b ++ (d0 :: nil)) r); auto; intro;
- apply in_or_app; auto.
-elim (in_app_or (getdst b) (d0 :: nil) r); auto; intro.
-right; simpl; right; apply in_or_app; auto.
-elim H3; intro.
-right; simpl; right; apply in_or_app; right; simpl; auto.
-inversion H4.
-rewrite app_nil; (repeat (rewrite getdst_app; simpl)); intro.
-apply in_or_app; left; assumption.
+ induction 1; constructor.
+ apply Loc.notin_not_in; auto. auto.
Qed.
-
-Lemma stepp_sameExec:
- forall (r1 r2 : State), stepp r1 r2 -> stepInv r1 -> sameExec r1 r2.
-Proof.
-intros; induction H as [r|r1 r2 r3 H H1 HrecH].
-unfold sameExec; intros; auto.
-cut (sameExec r1 r2); [idtac | apply (step_sameExec r1); auto].
-unfold sameExec; unfold sameExec in HrecH |-; intros.
-rewrite H2; auto.
-rewrite HrecH; auto.
-apply (step_inv r1); auto.
-intros x H5; apply H4.
-generalize H5; (repeat rewrite getdst_app); intros; apply in_or_app.
-elim
- (in_app_or
- (getdst (StateToMove r2) ++ getdst (StateBeing r2))
- (getdst (StateToMove r3) ++ getdst (StateBeing r3)) x); auto; intro.
-generalize (step_inv_getdst r1 r2 x); (repeat rewrite getdst_app); intro.
-left; apply H8; auto.
-intros x H5; apply H4.
-generalize H5; (repeat rewrite getdst_app); intros; apply in_or_app.
-elim
- (in_app_or
- (getdst (StateToMove r1) ++ getdst (StateBeing r1))
- (getdst (StateToMove r2) ++ getdst (StateBeing r2)) x); auto; intro.
-generalize (step_inv_getdst r1 r2 x); (repeat rewrite getdst_app); intro.
-left; apply H8; auto.
+
+Lemma parmove_prop_1:
+ forall srcs dsts,
+ List.length srcs = List.length dsts ->
+ Loc.norepet dsts ->
+ Loc.disjoint srcs temporaries ->
+ Loc.disjoint dsts temporaries ->
+ forall e,
+ let e' := exec_seq (parmove srcs dsts) e in
+ List.map e' dsts = List.map e srcs /\
+ forall l, ~In l dsts -> l <> R IT2 -> l <> R FT2 -> e' l = e l.
+Proof.
+ intros.
+ assert (NR: list_norepet dsts) by (apply loc_norepet_norepet; auto).
+ assert (NTS: forall r, In r srcs -> Parmov.is_not_temp loc temp_for r).
+ intros. apply disjoint_temp_not_temp. apply Loc.disjoint_notin with srcs; auto.
+ assert (NTD: forall r, In r dsts -> Parmov.is_not_temp loc temp_for r).
+ intros. apply disjoint_temp_not_temp. apply Loc.disjoint_notin with dsts; auto.
+ generalize (Parmov.parmove2_correctness loc temp_for val Loc.eq srcs dsts H NR NTS NTD e).
+ change (Parmov.exec_seq loc val Loc.eq (Parmov.parmove2 loc temp_for Loc.eq srcs dsts) e) with e'.
+ intros [A B].
+ split. auto. intros. apply B. auto. rewrite is_not_temp_charact; auto.
Qed.
-
-Inductive dstep : State -> State -> Prop :=
- dstep_nop:
- forall (r : Reg) (t l : Moves), dstep ((r, r) :: t, nil, l) (t, nil, l)
- | dstep_start:
- forall (t l : Moves) (s d : Reg),
- s <> d -> dstep ((s, d) :: t, nil, l) (t, (s, d) :: nil, l)
- | dstep_push:
- forall (t1 t2 b l : Moves) (s d r : Reg),
- noRead t1 d ->
- dstep
- (t1 ++ ((d, r) :: t2), (s, d) :: b, l)
- (t1 ++ t2, (d, r) :: ((s, d) :: b), l)
- | dstep_pop_loop:
- forall (t b l : Moves) (s d r0 : Reg),
- noRead t r0 ->
- dstep
- (t, (s, r0) :: (b ++ ((r0, d) :: nil)), l)
- (t, b ++ ((T r0, d) :: nil), (s, r0) :: ((r0, T r0) :: l))
- | dstep_pop:
- forall (t b l : Moves) (s0 d0 sn dn : Reg),
- noRead t dn ->
- Loc.diff dn s0 ->
- dstep
- (t, (sn, dn) :: (b ++ ((s0, d0) :: nil)), l)
- (t, b ++ ((s0, d0) :: nil), (sn, dn) :: l)
- | dstep_last:
- forall (t l : Moves) (s d : Reg),
- noRead t d -> dstep (t, (s, d) :: nil, l) (t, nil, (s, d) :: l) .
-Hint Resolve dstep_nop dstep_start dstep_push .
-Hint Resolve dstep_pop_loop dstep_pop dstep_last .
-
-Lemma dstep_step:
- forall (r1 r2 : State), dstep r1 r2 -> stepInv r1 -> stepp r1 r2.
-Proof.
-intros r1 r2 DS; inversion_clear DS; intros SI; eauto.
-change (stepp (nil ++ ((r, r) :: t), nil, l) (t, nil, l)); apply step_stepp;
- apply (step_nop r nil t).
-change (stepp (nil ++ ((s, d) :: t), nil, l) (t, (s, d) :: nil, l));
- apply step_stepp; apply (step_start nil t l).
-apply
- (stepp_trans
- (t, (s, r0) :: (b ++ ((r0, d) :: nil)), l)
- (t, (s, r0) :: (b ++ ((T r0, d) :: nil)), (r0, T r0) :: l)
- (t, b ++ ((T r0, d) :: nil), (s, r0) :: ((r0, T r0) :: l))); auto.
-apply step_stepp; apply step_pop; auto.
-unfold stepInv in SI |-; generalize SI; intros [X [Y [Z [U V]]]].
-generalize V; unfold StateBeing, noTmpLast.
-case (b ++ ((r0, d) :: nil)); auto.
-intros m l0 [R1 [OK PP]]; auto.
+
+Lemma parmove_prop_2:
+ forall srcs dsts s d,
+ In (s, d) (parmove srcs dsts) ->
+ (In s srcs \/ s = R IT2 \/ s = R FT2)
+ /\ (In d dsts \/ d = R IT2 \/ d = R FT2).
+Proof.
+ intros srcs dsts.
+ set (mu := List.combine srcs dsts).
+ assert (forall s d, Parmov.wf_move loc temp_for mu s d ->
+ (In s srcs \/ s = R IT2 \/ s = R FT2)
+ /\ (In d dsts \/ d = R IT2 \/ d = R FT2)).
+ unfold mu; induction 1.
+ split.
+ left. eapply List.in_combine_l; eauto.
+ left. eapply List.in_combine_r; eauto.
+ split.
+ right. apply temp_for_charact.
+ tauto.
+ split.
+ tauto.
+ right. apply temp_for_charact.
+ intros. apply H.
+ apply (Parmov.parmove2_wf_moves loc temp_for Loc.eq srcs dsts s d H0).
Qed.
-
-Lemma dstep_inv:
- forall (r1 r2 : State), dstep r1 r2 -> stepInv r1 -> stepInv r2.
+
+Lemma loc_type_temp_for:
+ forall l, Loc.type (temp_for l) = Loc.type l.
Proof.
-intros.
-apply (stepp_inv r1 r2); auto.
-apply dstep_step; auto.
+ intros; unfold temp_for. destruct (Loc.type l); reflexivity.
Qed.
-
-Inductive dstepp : State -> State -> Prop :=
- dstepp_refl: forall (r : State), dstepp r r
- | dstepp_trans:
- forall (r1 r2 r3 : State), dstep r1 r2 -> dstepp r2 r3 -> dstepp r1 r3 .
-Hint Resolve dstepp_refl dstepp_trans .
-
-Lemma dstepp_stepp:
- forall (s1 s2 : State), stepInv s1 -> dstepp s1 s2 -> stepp s1 s2.
-Proof.
-intros; induction H0 as [r|r1 r2 r3 H0 H1 HrecH0]; auto.
-apply (stepp_transitive r1 r2 r3); auto.
-apply dstep_step; auto.
-apply HrecH0; auto.
-apply (dstep_inv r1 r2); auto.
+
+Lemma loc_type_combine:
+ forall srcs dsts,
+ List.map Loc.type srcs = List.map Loc.type dsts ->
+ forall s d,
+ In (s, d) (List.combine srcs dsts) ->
+ Loc.type s = Loc.type d.
+Proof.
+ induction srcs; destruct dsts; simpl; intros; try discriminate.
+ elim H0.
+ elim H0; intros. inversion H1; subst. congruence.
+ apply IHsrcs with dsts. congruence. auto.
Qed.
-
-Lemma dstepp_sameExec:
- forall (r1 r2 : State), dstepp r1 r2 -> stepInv r1 -> sameExec r1 r2.
-Proof.
-intros; apply stepp_sameExec; auto.
-apply dstepp_stepp; auto.
+
+Lemma parmove_prop_3:
+ forall srcs dsts,
+ List.map Loc.type srcs = List.map Loc.type dsts ->
+ forall s d,
+ In (s, d) (parmove srcs dsts) -> Loc.type s = Loc.type d.
+Proof.
+ intros srcs dsts TYP.
+ set (mu := List.combine srcs dsts).
+ assert (forall s d, Parmov.wf_move loc temp_for mu s d ->
+ Loc.type s = Loc.type d).
+ unfold mu; induction 1.
+ eapply loc_type_combine; eauto.
+ rewrite loc_type_temp_for; auto.
+ rewrite loc_type_temp_for; auto.
+ intros. apply H.
+ apply (Parmov.parmove2_wf_moves loc temp_for Loc.eq srcs dsts s d H0).
Qed.
-
-End pmov.
-Fixpoint split_move' (m : Moves) (r : Reg) {struct m} :
- option ((Moves * Reg) * Moves) :=
- match m with
- (s, d) :: tail =>
- match diff_dec s r with
- right _ => Some (nil, d, tail)
- | left _ =>
- match split_move' tail r with
- Some ((t1, r2, t2)) => Some ((s, d) :: t1, r2, t2)
- | None => None
- end
- end
- | nil => None
- end.
-
-Fixpoint split_move (m : Moves) (r : Reg) {struct m} :
- option ((Moves * Reg) * Moves) :=
- match m with
- (s, d) :: tail =>
- match Loc.eq s r with
- left _ => Some (nil, d, tail)
- | right _ =>
- match split_move tail r with
- Some ((t1, r2, t2)) => Some ((s, d) :: t1, r2, t2)
- | None => None
- end
- end
- | nil => None
- end.
+Section EQUIVALENCE.
-Definition def : Move := (R IT1, R IT1).
-
-Fixpoint last (M : Moves) : Move :=
- match M with nil => def
- | m :: nil => m
- | m :: tail => last tail end.
-
-Fixpoint head_but_last (M : Moves) : Moves :=
- match M with
- nil => nil
- | m' :: nil => nil
- | m' :: tail => m' :: head_but_last tail
- end.
-
-Fixpoint replace_last_s (M : Moves) : Moves :=
- match M with
- nil => nil
- | m :: nil =>
- match m with (s, d) => (T s, d) :: nil end
- | m :: tail => m :: replace_last_s tail
- end.
-
-Ltac CaseEq name := generalize (refl_equal name); pattern name at -1; case name.
-
-Definition stepf' (S1 : State) : State :=
- match S1 with
- (nil, nil, _) => S1
- | ((s, d) :: tl, nil, l) =>
- match diff_dec s d with
- right _ => (tl, nil, l)
- | left _ => (tl, (s, d) :: nil, l)
- end
- | (t, (s, d) :: b, l) =>
- match split_move t d with
- Some ((t1, r, t2)) =>
- (t1 ++ t2, (d, r) :: ((s, d) :: b), l)
- | None =>
- match b with
- nil => (t, nil, (s, d) :: l)
- | _ =>
- match diff_dec d (fst (last b)) with
- right _ =>
- (t, replace_last_s b, (s, d) :: ((d, T d) :: l))
- | left _ => (t, b, (s, d) :: l)
- end
- end
- end
- end.
-
-Definition stepf (S1 : State) : State :=
- match S1 with
- (nil, nil, _) => S1
- | ((s, d) :: tl, nil, l) =>
- match Loc.eq s d with
- left _ => (tl, nil, l)
- | right _ => (tl, (s, d) :: nil, l)
- end
- | (t, (s, d) :: b, l) =>
- match split_move t d with
- Some ((t1, r, t2)) =>
- (t1 ++ t2, (d, r) :: ((s, d) :: b), l)
- | None =>
- match b with
- nil => (t, nil, (s, d) :: l)
- | _ =>
- match Loc.eq d (fst (last b)) with
- left _ =>
- (t, replace_last_s b, (s, d) :: ((d, T d) :: l))
- | right _ => (t, b, (s, d) :: l)
- end
- end
- end
- end.
-
-Lemma rebuild_l:
- forall (l : Moves) (m : Move),
- m :: l = head_but_last (m :: l) ++ (last (m :: l) :: nil).
-Proof.
-induction l; simpl; auto.
-intros m; rewrite (IHl a); auto.
-Qed.
-
-Lemma splitSome:
- forall (l t1 t2 : Moves) (s d r : Reg),
- noOverlap (l ++ ((r, s) :: nil)) ->
- split_move l s = Some (t1, d, t2) -> noRead t1 s.
-Proof.
-induction l; simpl.
-intros; discriminate.
-destruct a as [a1 a2].
-intros t1 t2 s d r Hno; case (Loc.eq a1 s).
-intros e H1; inversion H1.
-simpl; auto.
-CaseEq (split_move l s).
-intros; (repeat destruct p).
-inversion H0; auto.
-simpl; split; auto.
-change (noOverlap (((a1, a2) :: l) ++ ((r, s) :: nil))) in Hno |-.
-assert (noOverlap ((r, s) :: ((a1, a2) :: l))).
-apply noOverlap_Front0; auto.
-unfold noOverlap in H1 |-; simpl in H1 |-.
-elim H1 with ( l0 := a1 );
- [intros H5 H6; (try clear H1); (try exact H5) | idtac].
-elim H5; [intros H1; (try clear H5); (try exact H1) | intros H1; (try clear H5)].
-absurd (a1 = s); auto.
-apply Loc.diff_sym; auto.
-right; left; trivial.
-apply (IHl m0 m s r0 r); auto.
-apply (noOverlap_pop (a1, a2)); auto.
-intros; discriminate.
-Qed.
-
-Lemma unsplit_move:
- forall (l t1 t2 : Moves) (s d r : Reg),
- noOverlap (l ++ ((r, s) :: nil)) ->
- split_move l s = Some (t1, d, t2) -> l = t1 ++ ((s, d) :: t2).
-Proof.
-induction l.
-simpl; intros; discriminate.
-intros t1 t2 s d r HnoO; destruct a as [a1 a2]; simpl; case (diff_dec a1 s);
- intro.
-case (Loc.eq a1 s); intro.
-absurd (Loc.diff a1 s); auto.
-rewrite e; apply Loc.same_not_diff.
-CaseEq (split_move l s); intros; (try discriminate).
-(repeat destruct p); inversion H0.
-rewrite app_cons; subst t2; subst d; rewrite (IHl m0 m s r0 r); auto.
-apply (noOverlap_pop (a1, a2)); auto.
-case (Loc.eq a1 s); intros e H; inversion H; simpl.
-rewrite e; auto.
-cut (noOverlap_aux a1 (getdst ((r, s) :: nil))).
-intros [[H5|H4] H0]; [try exact H5 | idtac].
-absurd (s = a1); auto.
-absurd (Loc.diff a1 s); auto; apply Loc.diff_sym; auto.
-generalize HnoO; rewrite app_cons; intro.
-assert (noOverlap (l ++ ((a1, a2) :: ((r, s) :: nil))));
- (try (apply noOverlap_insert; assumption)).
-assert (noOverlap ((a1, a2) :: ((r, s) :: nil))).
-apply (noOverlap_right l); auto.
-generalize H2; unfold noOverlap; simpl.
-intros H5; elim (H5 a1); [idtac | left; trivial].
-intros H6 [[H7|H8] H9].
-absurd (s = a1); auto.
-split; [right; (try assumption) | auto].
-Qed.
-
-Lemma cons_replace:
- forall (a : Move) (l : Moves),
- l <> nil -> replace_last_s (a :: l) = a :: replace_last_s l.
-Proof.
-intros; simpl.
-CaseEq l.
-intro; contradiction.
-intros m l0 H0; auto.
-Qed.
-
-Lemma last_replace:
- forall (l : Moves) (s d : Reg),
- replace_last_s (l ++ ((s, d) :: nil)) = l ++ ((T s, d) :: nil).
-Proof.
-induction l; (try (simpl; auto; fail)).
-intros; (repeat rewrite <- app_comm_cons).
-rewrite cons_replace.
-rewrite IHl; auto.
-red; intro.
-elim (app_eq_nil l ((s, d) :: nil)); auto; intros; discriminate.
-Qed.
-
-Lemma last_app: forall (l : Moves) (m : Move), last (l ++ (m :: nil)) = m.
-Proof.
-induction l; simpl; auto.
-intros m; CaseEq (l ++ (m :: nil)).
-intro; elim (app_eq_nil l (m :: nil)); auto; intros; discriminate.
-intros m0 l0 H; (rewrite <- H; apply IHl).
-Qed.
-
-Lemma last_cons:
- forall (l : Moves) (m m0 : Move), last (m0 :: (m :: l)) = last (m :: l).
-Proof.
-intros; simpl; auto.
-Qed.
-
-Lemma stepf_popLoop:
- forall (t b l : Moves) (s d r0 : Reg),
- split_move t d = None ->
- stepf (t, (s, d) :: (b ++ ((d, r0) :: nil)), l) =
- (t, b ++ ((T d, r0) :: nil), (s, d) :: ((d, T d) :: l)).
-Proof.
-intros; simpl; rewrite H; CaseEq (b ++ ((d, r0) :: nil)); intros.
-destruct b; discriminate.
-rewrite <- H0; rewrite last_app; simpl; rewrite last_replace.
-case (Loc.eq d d); intro; intuition.
-destruct t; (try destruct m0); simpl; auto.
-Qed.
-
-Lemma stepf_pop:
- forall (t b l : Moves) (s d r r0 : Reg),
- split_move t d = None ->
- d <> r ->
- stepf (t, (s, d) :: (b ++ ((r, r0) :: nil)), l) =
- (t, b ++ ((r, r0) :: nil), (s, d) :: l).
-Proof.
-intros; simpl; rewrite H; CaseEq (b ++ ((r, r0) :: nil)); intros.
-destruct b; discriminate.
-rewrite <- H1; rewrite last_app; simpl.
-case (Loc.eq d r); intro.
-absurd (d = r); auto.
-destruct t; (try destruct m0); simpl; auto.
-Qed.
-
-Lemma noOverlap_head:
- forall l1 l2 m, noOverlap (l1 ++ (m :: l2)) -> noOverlap (l1 ++ (m :: nil)).
-Proof.
-induction l2; simpl; auto.
-intros; apply IHl2.
-cut (l1 ++ (m :: (a :: l2)) = (l1 ++ (m :: nil)) ++ (a :: l2));
- [idtac | rewrite app_ass; auto].
-intros e; rewrite e in H.
-cut (l1 ++ (m :: l2) = (l1 ++ (m :: nil)) ++ l2);
- [idtac | rewrite app_ass; auto].
-intros e'; rewrite e'; auto.
-apply noOverlap_Pop with a; auto.
-Qed.
-
-Lemma splitNone:
- forall (l : Moves) (s d : Reg),
- split_move l d = None -> noOverlap (l ++ ((s, d) :: nil)) -> noRead l d.
-Proof.
-induction l; intros s d; simpl; auto.
-destruct a as [a1 a2]; case (Loc.eq a1 d); intro; (try (intro; discriminate)).
-CaseEq (split_move l d); intros.
-(repeat destruct p); discriminate.
-split; (try assumption).
-change (noOverlap (((a1, a2) :: l) ++ ((s, d) :: nil))) in H1 |-.
-assert (noOverlap ((s, d) :: ((a1, a2) :: l))).
-apply noOverlap_Front0; auto.
-assert (noOverlap ((a1, a2) :: ((s, d) :: l))).
-apply noOverlap_swap; auto.
-unfold noOverlap in H3 |-; simpl in H3 |-.
-elim H3 with ( l0 := a1 );
- [intros H5 H6; (try clear H1); (try exact H5) | idtac].
-elim H6;
- [intros H1 H4; elim H1;
- [intros H7; (try clear H1 H6); (try exact H7) | intros H7; (try clear H1 H6)]].
-absurd (a1 = d); auto.
-apply Loc.diff_sym; auto.
-left; trivial.
-apply IHl with s; auto.
-apply noOverlap_pop with (a1, a2); auto.
-Qed.
-
-Lemma noO_diff:
- forall l1 l2 s d r r0,
- noOverlap (l1 ++ ((s, d) :: (l2 ++ ((r, r0) :: nil)))) ->
- r = d \/ Loc.diff d r.
-Proof.
-intros.
-assert (noOverlap ((s, d) :: (l2 ++ ((r, r0) :: nil)))); auto.
-apply (noOverlap_right l1); auto.
-assert (noOverlap ((l2 ++ ((r, r0) :: nil)) ++ ((s, d) :: nil))); auto.
-apply (noOverlap_movBack0 (l2 ++ ((r, r0) :: nil))); auto.
-assert
- ((l2 ++ ((r, r0) :: nil)) ++ ((s, d) :: nil) =
- l2 ++ (((r, r0) :: nil) ++ ((s, d) :: nil))); auto.
-rewrite app_ass; auto.
-rewrite H2 in H1.
-simpl in H1 |-.
-assert (noOverlap ((r, r0) :: ((s, d) :: nil))); auto.
-apply (noOverlap_right l2); auto.
-unfold noOverlap in H3 |-.
-generalize (H3 r); simpl.
-intros H4; elim H4; intros; [idtac | left; trivial].
-elim H6; intros [H9|H9] H10; [left | right]; auto.
-Qed.
-
-Lemma f2ind:
- forall (S1 S2 : State),
- (forall (l : Moves), (S1 <> (nil, nil, l))) ->
- noOverlap (StateToMove S1 ++ StateBeing S1) -> stepf S1 = S2 -> dstep S1 S2.
-Proof.
-intros S1 S2 Hneq HnoO; destruct S1 as [[t b] l]; destruct b.
-destruct t.
-elim (Hneq l); auto.
-destruct m; simpl; case (Loc.eq r r0).
-intros.
-rewrite e; rewrite <- H; apply dstep_nop.
-intros n H; rewrite <- H; generalize (dstep_start t l r r0); auto.
-intros H; rewrite <- H; destruct m as [s d].
-CaseEq (split_move t d).
-intros p H0; destruct p as [[t1 s0] t2]; simpl; rewrite H0; destruct t; simpl.
-simpl in H0 |-; discriminate.
-rewrite (unsplit_move (m :: t) t1 t2 d s0 s); auto.
-destruct m; generalize dstep_push; intros H1; apply H1.
-unfold StateToMove, StateBeing in HnoO |-.
-apply (splitSome ((r, r0) :: t) t1 t2 d s0 s); auto.
-apply noOverlap_head with b; auto.
-unfold StateToMove, StateBeing in HnoO |-.
-apply noOverlap_head with b; auto.
-intros H0; destruct b.
-simpl.
-rewrite H0.
-destruct t; (try destruct m); generalize dstep_last; intros H1; apply H1.
-simpl; auto.
-unfold StateToMove, StateBeing in HnoO |-.
-apply splitNone with s; auto.
-unfold StateToMove, StateBeing in HnoO |-.
-generalize HnoO; clear HnoO; rewrite (rebuild_l b m); intros HnoO.
-destruct (last (m :: b)).
-case (Loc.eq d r).
-intros e; rewrite <- e.
-CaseEq (head_but_last (m :: b)); intros; [simpl | idtac];
- (try
- (destruct t; (try destruct m0); rewrite H0;
- (case (Loc.eq d d); intros h; (try (elim h; auto))))).
-generalize (dstep_pop_loop nil nil); simpl; intros H3; apply H3; auto.
-generalize (dstep_pop_loop ((r1, r2) :: t) nil); unfold T; simpl app;
- intros H3; apply H3; clear H3; apply splitNone with s; (try assumption).
-apply noOverlap_head with (head_but_last (m :: b) ++ ((r, r0) :: nil)); auto.
-rewrite stepf_popLoop; auto.
-generalize (dstep_pop_loop t (m0 :: l0)); simpl; intros H3; apply H3; clear H3;
- apply splitNone with s; (try assumption).
-apply noOverlap_head with (head_but_last (m :: b) ++ ((r, r0) :: nil)); auto.
-intro; assert (Loc.diff d r).
-assert (r = d \/ Loc.diff d r).
-apply (noO_diff t (head_but_last (m :: b)) s d r r0); auto.
-elim H1; [intros H2; absurd (d = r); auto | intros H2; auto].
-rewrite stepf_pop; auto.
-generalize (dstep_pop t (head_but_last (m :: b))); intros H3; apply H3; auto.
-clear H3; apply splitNone with s; (try assumption).
-apply noOverlap_head with (head_but_last (m :: b) ++ ((r, r0) :: nil)); auto.
-Qed.
-
-Lemma f2ind':
- forall (S1 : State),
- (forall (l : Moves), (S1 <> (nil, nil, l))) ->
- noOverlap (StateToMove S1 ++ StateBeing S1) -> dstep S1 (stepf S1).
-Proof.
-intros S1 H noO; apply f2ind; auto.
-Qed.
-
-Lemma appcons_length:
- forall (l1 l2 : Moves) (m : Move),
- length (l1 ++ (m :: l2)) = (length (l1 ++ l2) + 1%nat)%nat.
-Proof.
-induction l1; simpl; intros; [omega | idtac].
-rewrite IHl1; omega.
-Qed.
-
-Definition mesure (S0 : State) : nat :=
- let (p, _) := S0 in let (t, b) := p in (2 * length t + length b)%nat.
-
-Lemma step_dec0:
- forall (t1 t2 b1 b2 : Moves) (l1 l2 : Moves),
- dstep (t1, b1, l1) (t2, b2, l2) ->
- (2 * length t2 + length b2 < 2 * length t1 + length b1)%nat.
-Proof.
-intros t1 t2 b1 b2 l1 l2 H; inversion H; simpl; (try omega).
-rewrite appcons_length; omega.
-cut (length (b ++ ((T r0, d) :: nil)) = length (b ++ ((r0, d) :: nil)));
- (try omega).
-induction b; simpl; auto.
-(repeat rewrite appcons_length); auto.
-Qed.
-
-Lemma step_dec:
- forall (S1 S2 : State), dstep S1 S2 -> (mesure S2 < mesure S1)%nat.
-Proof.
-unfold mesure; destruct S1 as [[t1 b1] l1]; destruct S2 as [[t2 b2] l2].
-intro; apply (step_dec0 t1 t2 b1 b2 l1 l2); trivial.
-Qed.
-
-Lemma stepf_dec0:
- forall (S1 S2 : State),
- (forall (l : Moves), (S1 <> (nil, nil, l))) /\
- (S2 = stepf S1 /\ noOverlap (StateToMove S1 ++ StateBeing S1)) ->
- (mesure S2 < mesure S1)%nat.
-Proof.
-intros S1 S2 [H1 [H2 H3]]; apply step_dec.
-apply f2ind; trivial.
-rewrite H2; reflexivity.
-Qed.
-
-Lemma stepf_dec:
- forall (S1 S2 : State),
- S2 = stepf S1 /\
- ((forall (l : Moves), (S1 <> (nil, nil, l))) /\
- noOverlap (StateToMove S1 ++ StateBeing S1)) -> ltof _ mesure S2 S1.
-Proof.
-unfold ltof.
-intros S1 S2 [H1 [H2 H3]]; apply step_dec.
-apply f2ind; trivial.
-rewrite H1; reflexivity.
-Qed.
-
-Lemma replace_last_id:
- forall l m m0, replace_last_s (m :: (m0 :: l)) = m :: replace_last_s (m0 :: l).
-Proof.
-intros; case l; simpl.
-destruct m0; simpl; auto.
-intros; case l0; auto.
-Qed.
-
-Lemma length_replace: forall l, length (replace_last_s l) = length l.
-Proof.
-induction l; simpl; auto.
-destruct l; destruct a; simpl; auto.
-Qed.
-
-Lemma length_app:
- forall (A : Set) (l1 l2 : list A),
- (length (l1 ++ l2) = length l1 + length l2)%nat.
-Proof.
-intros A l1 l2; (try assumption).
-induction l1; simpl; auto.
-Qed.
-
-Lemma split_length:
- forall (l t1 t2 : Moves) (s d : Reg),
- split_move l s = Some (t1, d, t2) ->
- (length l = (length t1 + length t2) + 1)%nat.
-Proof.
-induction l.
-intros; discriminate.
-intros t1 t2 s d; destruct a as [r r0]; simpl; case (Loc.eq r s); intro.
-intros H; inversion H.
-simpl; omega.
-CaseEq (split_move l s); (try (intros; discriminate)).
-(repeat destruct p); intros H H0; inversion H0.
-rewrite H2; rewrite (IHl m0 m s r1); auto.
-rewrite H4; rewrite <- H2; simpl; omega.
-Qed.
-
-Lemma stepf_dec0':
- forall (S1 : State),
- (forall (l : Moves), (S1 <> (nil, nil, l))) ->
- (mesure (stepf S1) < mesure S1)%nat.
-Proof.
-intros S1 H.
-unfold mesure; destruct S1 as [[t1 b1] l1].
-destruct t1.
-destruct b1.
-generalize (H l1); intros H1; elim H1; auto.
-destruct m; simpl.
-destruct b1.
-simpl; auto.
-case (Loc.eq r0 (fst (last (m :: b1)))).
-intros; rewrite length_replace; simpl; omega.
-simpl; case b1; intros; simpl; omega.
-destruct m.
-destruct b1.
-simpl.
-case (Loc.eq r r0); intros; simpl; omega.
-destruct m; simpl; case (Loc.eq r r2).
-intros; simpl; omega.
-CaseEq (split_move t1 r2); intros.
-destruct p; destruct p; simpl.
-rewrite (split_length t1 m0 m r2 r3); auto.
-rewrite length_app; auto.
-omega.
-destruct b1.
-simpl; omega.
-case (Loc.eq r2 (fst (last (m :: b1)))); intros.
-rewrite length_replace; simpl; omega.
-simpl; omega.
-Qed.
-
-Lemma stepf1_dec:
- forall (S1 S2 : State),
- (forall (l : Moves), (S1 <> (nil, nil, l))) ->
- S2 = stepf S1 -> ltof _ mesure S2 S1.
-Proof.
-unfold ltof; intros S1 S2 H H0; rewrite H0.
-apply stepf_dec0'; (try assumption).
-Qed.
-
-Lemma disc1:
- forall (a : Move) (l1 l2 l3 l4 : list Move),
- ((a :: l1, l2, l3) <> (nil, nil, l4)).
-Proof.
-intros; discriminate.
-Qed.
-
-Lemma disc2:
- forall (a : Move) (l1 l2 l3 l4 : list Move),
- ((l1, a :: l2, l3) <> (nil, nil, l4)).
-Proof.
-intros; discriminate.
-Qed.
-Hint Resolve disc1 disc2 .
-
-Lemma sameExec_reflexive: forall (r : State), sameExec r r.
-Proof.
-intros r; unfold sameExec, sameEnv, exec.
-destruct r as [[t b] d]; trivial.
-Qed.
-
-Definition base_case_Pmov_dec:
- forall (s : State),
- ({ exists l : list Move , s = (nil, nil, l) }) +
- ({ forall l, (s <> (nil, nil, l)) }).
-Proof.
-destruct s as [[[|x tl] [|y tl']] l]; (try (right; intro; discriminate)).
-left; exists l; auto.
-Defined.
-
-Definition Pmov :=
- Fix
- (well_founded_ltof _ mesure) (fun _ => State)
- (fun (S1 : State) =>
- fun (Pmov : forall x, ltof _ mesure x S1 -> State) =>
- match base_case_Pmov_dec S1 with
- left h => S1
- | right h => Pmov (stepf S1) (stepf_dec0' S1 h) end).
-
-Theorem Pmov_equation: forall S1, Pmov S1 = match S1 with
- ((nil, nil), _) => S1
- | _ => Pmov (stepf S1)
- end.
-Proof.
-intros S1; unfold Pmov at 1;
- rewrite (Fix_eq
- (well_founded_ltof _ mesure) (fun _ => State)
- (fun (S1 : State) =>
- fun (Pmov : forall x, ltof _ mesure x S1 -> State) =>
- match base_case_Pmov_dec S1 with
- left h => S1
- | right h => Pmov (stepf S1) (stepf_dec0' S1 h) end)).
-fold Pmov.
-destruct S1 as [[[|x tl] [|y tl']] l];
- match goal with
- | |- match ?a with left _ => _ | right _ => _ end = _ => case a end;
- (try (intros [l0 Heq]; discriminate Heq)); auto.
-intros H; elim (H l); auto.
-intros x f g Hfg_ext.
-match goal with
-| |- match ?a with left _ => _ | right _ => _ end = _ => case a end; auto.
-Qed.
-
-Lemma sameExec_transitive:
- forall (r1 r2 r3 : State),
- (forall r,
- In r (getdst (StateToMove r2 ++ StateBeing r2)) ->
- In r (getdst (StateToMove r1 ++ StateBeing r1))) ->
- (forall r,
- In r (getdst (StateToMove r3 ++ StateBeing r3)) ->
- In r (getdst (StateToMove r2 ++ StateBeing r2))) ->
- sameExec r1 r2 -> sameExec r2 r3 -> sameExec r1 r3.
-Proof.
-intros r1 r2 r3; unfold sameExec, exec; (repeat rewrite getdst_app).
-destruct r1 as [[t1 b1] d1]; destruct r2 as [[t2 b2] d2];
- destruct r3 as [[t3 b3] d3]; simpl.
-intros Hin; intros.
-rewrite H0; auto.
-rewrite H1; auto.
-intros.
-apply (H3 x).
-apply in_or_app; auto.
-elim (in_app_or (getdst t2 ++ getdst b2) (getdst t3 ++ getdst b3) x); auto.
-intros.
-apply (H3 x).
-apply in_or_app; auto.
-elim (in_app_or (getdst t1 ++ getdst b1) (getdst t2 ++ getdst b2) x); auto.
-Qed.
-
-Lemma dstep_inv_getdst:
- forall (s1 s2 : State) r,
- dstep s1 s2 ->
- In r (getdst (StateToMove s2 ++ StateBeing s2)) ->
- In r (getdst (StateToMove s1 ++ StateBeing s1)).
-intros s1 s2 r STEP; inversion_clear STEP;
- unfold StateToMove, StateBeing, StateDone; (repeat rewrite app_nil);
- (repeat (rewrite getdst_app; simpl)); intro; auto.
-Proof.
-right; (try assumption).
-elim (in_app_or (getdst t) (d :: nil) r); auto; (simpl; intros [H1|H1]);
- [left; assumption | inversion H1].
-elim (in_app_or (getdst t1 ++ getdst t2) (r0 :: (d :: getdst b)) r); auto;
- (simpl; intros).
-elim (in_app_or (getdst t1) (getdst t2) r); auto; (simpl; intros).
-apply in_or_app; left; apply in_or_app; left; assumption.
-apply in_or_app; left; apply in_or_app; right; simpl; right; assumption.
-elim H1; [intros H2 | intros [H2|H2]].
-apply in_or_app; left; apply in_or_app; right; simpl; left; auto.
-apply in_or_app; right; simpl; left; auto.
-apply in_or_app; right; simpl; right; assumption.
-elim (in_app_or (getdst t) (getdst b ++ (d :: nil)) r); auto; (simpl; intros).
-apply in_or_app; left; assumption.
-elim (in_app_or (getdst b) (d :: nil) r); auto; (simpl; intros).
-apply in_or_app; right; simpl; right; apply in_or_app; left; assumption.
-elim H2; [intros H3 | intros H3; inversion H3].
-apply in_or_app; right; simpl; right; apply in_or_app; right; simpl; auto.
-elim (in_app_or (getdst t) (getdst b ++ (d0 :: nil)) r); auto; (simpl; intros).
-apply in_or_app; left; assumption.
-elim (in_app_or (getdst b) (d0 :: nil) r); auto; simpl;
- [intros H3 | intros [H3|H3]; [idtac | inversion H3]].
-apply in_or_app; right; simpl; right; apply in_or_app; left; assumption.
-apply in_or_app; right; simpl; right; apply in_or_app; right; simpl; auto.
-apply in_or_app; left; assumption.
-Qed.
-
-Theorem STM_Pmov: forall (S1 : State), StateToMove (Pmov S1) = nil.
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1; intros Hrec; destruct S1 as [[t b] d];
- rewrite Pmov_equation; destruct t.
-destruct b; auto.
-apply Hrec; apply stepf1_dec; auto.
-apply Hrec; apply stepf1_dec; auto.
-Qed.
-
-Theorem SB_Pmov: forall (S1 : State), StateBeing (Pmov S1) = nil.
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1; intros Hrec; destruct S1 as [[t b] d];
- rewrite Pmov_equation; destruct t.
-destruct b; auto.
-apply Hrec; apply stepf1_dec; auto.
-apply Hrec; apply stepf1_dec; auto.
-Qed.
-
-Theorem Fpmov_correct:
- forall (S1 : State), stepInv S1 -> sameExec S1 (Pmov S1).
-Proof.
-intros S1; elim S1 using (well_founded_ind (Wf_nat.well_founded_ltof _ mesure)).
-clear S1; intros S1; intros Hrec Hinv; rewrite Pmov_equation;
- destruct S1 as [[t b] d].
-assert
- (forall (r : Reg) S1,
- In r (getdst (StateToMove (Pmov (stepf S1)) ++ StateBeing (Pmov (stepf S1)))) ->
- In r (getdst (StateToMove (stepf S1) ++ StateBeing (stepf S1)))).
-intros r S1; rewrite (STM_Pmov (stepf S1)); rewrite SB_Pmov; simpl; intros.
-inversion H.
-destruct t.
-destruct b.
-apply sameExec_reflexive.
-set (S1:=(nil (A:=Move), m :: b, d)).
-assert (dstep S1 (stepf S1)); (try apply f2ind); unfold S1; auto.
-elim Hinv; intros Hpath [SD [NO NT]]; assumption.
-apply sameExec_transitive with (stepf S1); auto.
-intros r; apply dstep_inv_getdst; auto.
-apply dstepp_sameExec; auto; apply dstepp_trans with (stepf S1); auto.
-apply dstepp_refl; auto.
-apply Hrec; auto.
-unfold ltof; apply step_dec; assumption.
-apply (dstep_inv S1); assumption.
-set (S1:=(m :: t, b, d)).
-assert (dstep S1 (stepf S1)); (try apply f2ind); unfold S1; auto.
-elim Hinv; intros Hpath [SD [NO NT]]; assumption.
-apply sameExec_transitive with (stepf S1); auto.
-intros r; apply dstep_inv_getdst; auto.
-apply dstepp_sameExec; auto; apply dstepp_trans with (stepf S1); auto.
-apply dstepp_refl; auto.
-apply Hrec; auto.
-unfold ltof; apply step_dec; assumption.
-apply (dstep_inv S1); assumption.
-Qed.
-
-Definition P_move := fun (p : Moves) => StateDone (Pmov (p, nil, nil)).
-
-Definition Sexec := sexec.
-
-Definition Get := get.
-
-Fixpoint listsLoc2Moves (src dst : list loc) {struct src} : Moves :=
- match src with
- nil => nil
- | s :: srcs =>
- match dst with
- nil => nil
- | d :: dsts => (s, d) :: listsLoc2Moves srcs dsts
- end
- end.
-
-Definition no_overlap (l1 l2 : list loc) :=
- forall r, In r l1 -> forall s, In s l2 -> r = s \/ Loc.diff r s.
-
-Definition no_overlap_state (S : State) :=
- no_overlap
- (getsrc (StateToMove S ++ StateBeing S))
- (getdst (StateToMove S ++ StateBeing S)).
-
-Definition no_overlap_list := fun l => no_overlap (getsrc l) (getdst l).
-
-Lemma Indst_noOverlap_aux:
- forall l1 l,
- (forall (s : Reg), In s (getdst l1) -> l = s \/ Loc.diff l s) ->
- noOverlap_aux l (getdst l1).
-Proof.
-intros; induction l1; simpl; auto.
-destruct a as [a1 a2]; simpl; split.
-elim (H a2); (try intros H0).
-left; auto.
-right; apply Loc.diff_sym; auto.
-simpl; left; trivial.
-apply IHl1; intros.
-apply H; simpl; right; (try assumption).
-Qed.
-
-Lemma no_overlap_noOverlap:
- forall r, no_overlap_state r -> noOverlap (StateToMove r ++ StateBeing r).
-Proof.
-intros r; unfold noOverlap, no_overlap_state.
-set (l1:=StateToMove r ++ StateBeing r).
-unfold no_overlap; intros H l H0.
-apply Indst_noOverlap_aux; intros; apply H; auto.
-Qed.
-
-Theorem Fpmov_correctMoves:
- forall p e r,
- simpleDest p ->
- no_overlap_list p ->
- noTmp p ->
- notemporary r ->
- (forall (x : Reg), In x (getdst p) -> r = x \/ Loc.diff r x) ->
- get (pexec p e) r = get (sexec (StateDone (Pmov (p, nil, nil))) e) r.
-Proof.
-intros p e r SD no_O notmp notempo.
-generalize (Fpmov_correct (p, nil, nil)); unfold sameExec, exec; simpl;
- rewrite SB_Pmov; rewrite STM_Pmov; simpl.
-(repeat rewrite app_nil); intro.
-apply H; auto.
-unfold stepInv; simpl; (repeat split); (try (rewrite app_nil; assumption)); auto.
-generalize (no_overlap_noOverlap (p, nil, nil)); simpl; intros; auto.
-apply H0; auto; unfold no_overlap_list in H0 |-.
-unfold no_overlap_state; simpl; (repeat rewrite app_nil); auto.
-Qed.
-
-Theorem Fpmov_correct1:
- forall (p : Moves) (e : Env) (r : Reg),
- simpleDest p ->
- no_overlap_list p ->
- noTmp p ->
- notemporary r ->
- (forall (x : Reg), In x (getdst p) -> r = x \/ Loc.diff r x) ->
- noWrite p r -> get e r = get (sexec (StateDone (Pmov (p, nil, nil))) e) r.
-Proof.
-intros p e r Hsd Hno_O HnoTmp Hrnotempo Hrno_Overlap Hnw.
-rewrite <- (Fpmov_correctMoves p e); (try assumption).
-destruct p; auto.
-destruct m as [m1 m2]; simpl; case (Loc.eq m2 r); intros.
-elim Hnw; intros; absurd (Loc.diff m2 r); auto.
-rewrite e0; apply Loc.same_not_diff.
-elim Hnw; intros H1 H2.
-rewrite get_update_diff; (try assumption).
-apply get_noWrite; (try assumption).
-Qed.
-
-Lemma In_SD_diff:
- forall (s d a1 a2 : Reg) (p : Moves),
- In (s, d) p -> simpleDest ((a1, a2) :: p) -> Loc.diff a2 d.
-Proof.
-intros; induction p.
-inversion H.
-elim H; auto.
-intro; subst a; elim H0; intros H1 H2; elim H1; intros; apply Loc.diff_sym;
- assumption.
-intro; apply IHp; auto.
-apply simpleDest_pop2 with a; (try assumption).
-Qed.
-
-Theorem pexec_correct:
- forall (e : Env) (m : Move) (p : Moves),
- In m p -> simpleDest p -> (let (s, d) := m in get (pexec p e) d = get e s).
-Proof.
-induction p; intros.
-elim H.
-destruct m.
-elim (in_inv H); intro.
-rewrite H1; simpl; rewrite get_update_id; auto.
-destruct a as [a1 a2]; simpl.
-rewrite get_update_diff.
-apply IHp; auto.
-apply (simpleDest_pop (a1, a2)); (try assumption).
-apply (In_SD_diff r) with ( p := p ) ( a1 := a1 ); auto.
-Qed.
-
-Lemma In_noTmp_notempo:
- forall (s d : Reg) (p : Moves), In (s, d) p -> noTmp p -> notemporary d.
-Proof.
-intros; unfold notemporary; induction p.
-inversion H.
-elim H; intro.
-subst a; elim H0; intros H1 [H3 H2]; (try assumption).
-intro; apply IHp; auto.
-destruct a; elim H0; intros _ [H2 H3]; (try assumption).
-Qed.
-
-Lemma In_Indst: forall s d p, In (s, d) p -> In d (getdst p).
-Proof.
-intros; induction p; auto.
-destruct a; simpl.
-elim H; intro.
-left; inversion H0; trivial.
-right; apply IHp; auto.
-Qed.
-
-Lemma In_SD_diff':
- forall (d a1 a2 : Reg) (p : Moves),
- In d (getdst p) -> simpleDest ((a1, a2) :: p) -> Loc.diff a2 d.
-Proof.
-intros d a1 a2 p H H0; induction p.
-inversion H.
-destruct a; elim H.
-elim H0; simpl; intros.
-subst r0.
-elim H1; intros H3 H4; apply Loc.diff_sym; assumption.
-intro; apply IHp; (try assumption).
-apply simpleDest_pop2 with (r, r0); (try assumption).
-Qed.
-
-Lemma In_SD_no_o:
- forall (s d : Reg) (p : Moves),
- In (s, d) p ->
- simpleDest p -> forall (x : Reg), In x (getdst p) -> d = x \/ Loc.diff d x.
-Proof.
-intros s d p Hin Hsd; induction p.
-inversion Hin.
-destruct a as [a1 a2]; elim Hin; intros.
-inversion H; subst d; subst s.
-elim H0; intros H1; [left | right]; (try assumption).
-apply (In_SD_diff' x a1 a2 p); auto.
-elim H0.
-intro; subst x.
-right; apply Loc.diff_sym; apply (In_SD_diff s d a1 a2 p); auto.
-intro; apply IHp; auto.
-apply (simpleDest_pop (a1, a2)); assumption.
-Qed.
-
-Lemma getdst_map: forall p, getdst p = map (fun x => snd x) p.
-Proof.
-induction p.
-simpl; auto.
-destruct a; simpl.
-rewrite IHp; auto.
+Variables srcs dsts: list loc.
+Hypothesis LENGTH: List.length srcs = List.length dsts.
+Hypothesis NOREPET: Loc.norepet dsts.
+Hypothesis NO_OVERLAP: Loc.no_overlap srcs dsts.
+Hypothesis NO_SRCS_TEMP: Loc.disjoint srcs temporaries.
+Hypothesis NO_DSTS_TEMP: Loc.disjoint dsts temporaries.
+
+Definition no_overlap_dests (l: loc) : Prop :=
+ forall d, In d dsts -> l = d \/ Loc.diff l d.
+
+Lemma dests_no_overlap_dests:
+ forall l, In l dsts -> no_overlap_dests l.
+Proof.
+ assert (forall d, Loc.norepet d ->
+ forall l1 l2, In l1 d -> In l2 d -> l1 = l2 \/ Loc.diff l1 l2).
+ induction 1; simpl; intros.
+ contradiction.
+ elim H1; intro; elim H2; intro.
+ left; congruence.
+ right. subst l1. eapply Loc.in_notin_diff; eauto.
+ right. subst l2. apply Loc.diff_sym. eapply Loc.in_notin_diff; eauto.
+ eauto.
+ intros; red; intros. eauto.
Qed.
-
-Lemma getsrc_map: forall p, getsrc p = map (fun x => fst x) p.
+
+Lemma notin_dests_no_overlap_dests:
+ forall l, Loc.notin l dsts -> no_overlap_dests l.
Proof.
-induction p.
-simpl; auto.
-destruct a; simpl.
-rewrite IHp; auto.
+ intros; red; intros.
+ right. eapply Loc.in_notin_diff; eauto.
Qed.
-
-Theorem Fpmov_correct2:
- forall (p : Moves) (e : Env) (m : Move),
- In m p ->
- simpleDest p ->
- no_overlap_list p ->
- noTmp p ->
- (let (s, d) := m in get (sexec (StateDone (Pmov (p, nil, nil))) e) d = get e s).
+
+Lemma source_no_overlap_dests:
+ forall s, In s srcs \/ s = R IT2 \/ s = R FT2 -> no_overlap_dests s.
Proof.
-intros p e m Hin Hsd Hno_O HnoTmp; destruct m as [s d];
- generalize (Fpmov_correctMoves p e); intros.
-rewrite <- H; auto.
-apply pexec_correct with ( m := (s, d) ); auto.
-apply (In_noTmp_notempo s d p); auto.
-apply (In_SD_no_o s d p Hin Hsd).
+ intros. elim H; intro. exact (NO_OVERLAP s H0).
+ elim H0; intro; subst s; red; intros;
+ right; apply Loc.diff_sym; apply NO_DSTS_TEMP; auto; simpl; tauto.
Qed.
-
-Lemma notindst_nW: forall a p, Loc.notin a (getdst p) -> noWrite p a.
+
+Lemma source_not_temp1:
+ forall s, In s srcs \/ s = R IT2 \/ s = R FT2 -> Loc.diff s (R IT1) /\ Loc.diff s (R FT1).
Proof.
-induction p; simpl; auto.
-destruct a0 as [a1 a2].
-simpl.
-intros H; elim H; intro; split.
-apply Loc.diff_sym; (try assumption).
-apply IHp; auto.
+ intros. elim H; intro.
+ split; apply NO_SRCS_TEMP; auto; simpl; tauto.
+ elim H0; intro; subst s; simpl; split; congruence.
Qed.
-
-Lemma disjoint_tmp__noTmp:
- forall p,
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries -> noTmp p.
-Proof.
-induction p; simpl; auto.
-destruct a as [a1 a2]; simpl getsrc; simpl getdst; unfold Loc.disjoint; intros;
- (repeat split).
-intro; unfold T; case (Loc.type r); apply H; (try (left; trivial; fail)).
-right; left; trivial.
-right; right; right; right; left; trivial.
-intro; unfold T; case (Loc.type r); apply H0; (try (left; trivial; fail)).
-right; left; trivial.
-right; right; right; right; left; trivial.
-apply IHp.
-apply Loc.disjoint_cons_left with a1; auto.
-apply Loc.disjoint_cons_left with a2; auto.
+
+Lemma dest_noteq_diff:
+ forall d l,
+ In d dsts \/ d = R IT2 \/ d = R FT2 ->
+ l <> d ->
+ no_overlap_dests l ->
+ Loc.diff l d.
+Proof.
+ intros. elim H; intro.
+ elim (H1 d H2); intro. congruence. auto.
+ assert (forall r, l <> R r -> Loc.diff l (R r)).
+ intros. destruct l; simpl. congruence. destruct s; auto.
+ elim H2; intro; subst d; auto.
Qed.
-
-Theorem Fpmov_correct_IT3:
- forall p rs,
- simpleDest p ->
- no_overlap_list p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- (sexec (StateDone (Pmov (p, nil, nil))) rs) (R IT3) = rs (R IT3).
-Proof.
-intros p rs Hsd Hno_O Hdistmpsrc Hdistmpdst.
-generalize (Fpmov_correctMoves p rs); unfold get, Locmap.get; intros H2.
-rewrite <- H2; auto.
-generalize (get_noWrite p (R IT3)); unfold get, Locmap.get; intros.
-rewrite <- H; auto.
-apply notindst_nW.
-apply (Loc.disjoint_notin temporaries).
-apply Loc.disjoint_sym; auto.
-right; right; left; trivial.
-apply disjoint_tmp__noTmp; auto.
-unfold notemporary, T.
-intros x; case (Loc.type x); simpl; intro; discriminate.
-intros x H; right; apply Loc.in_notin_diff with (getdst p); auto.
-apply Loc.disjoint_notin with temporaries; auto.
-apply Loc.disjoint_sym; auto.
-right; right; left; trivial.
+
+Definition locmap_equiv (e1 e2: Locmap.t): Prop :=
+ forall l,
+ no_overlap_dests l -> Loc.diff l (R IT1) -> Loc.diff l (R FT1) -> e2 l = e1 l.
+
+Definition effect_move (src dst: loc) (e e': Locmap.t): Prop :=
+ e' dst = e src /\
+ forall l, Loc.diff l dst -> Loc.diff l (R IT1) -> Loc.diff l (R FT1) -> e' l = e l.
+
+Inductive effect_seqmove: list (loc * loc) -> Locmap.t -> Locmap.t -> Prop :=
+ | effect_seqmove_nil: forall e,
+ effect_seqmove nil e e
+ | effect_seqmove_cons: forall s d m e1 e2 e3,
+ effect_move s d e1 e2 ->
+ effect_seqmove m e2 e3 ->
+ effect_seqmove ((s, d) :: m) e1 e3.
+
+Lemma effect_move_equiv:
+ forall s d e1 e2 e1',
+ (In s srcs \/ s = R IT2 \/ s = R FT2) ->
+ (In d dsts \/ d = R IT2 \/ d = R FT2) ->
+ locmap_equiv e1 e2 -> effect_move s d e1 e1' ->
+ locmap_equiv e1' (Parmov.update loc val Loc.eq d (e2 s) e2).
+Proof.
+ intros. destruct H2. red; intros.
+ unfold Parmov.update. destruct (Loc.eq l d).
+ subst l. elim (source_not_temp1 _ H); intros.
+ rewrite H2. apply H1; auto. apply source_no_overlap_dests; auto.
+ rewrite H3; auto. apply dest_noteq_diff; auto.
Qed.
-
-Theorem Fpmov_correct_map:
- forall p rs,
- simpleDest p ->
- no_overlap_list p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- List.map (sexec (StateDone (Pmov (p, nil, nil))) rs) (getdst p) =
- List.map rs (getsrc p).
-Proof.
-intros; rewrite getsrc_map; rewrite getdst_map; rewrite list_map_compose;
- rewrite list_map_compose; apply list_map_exten; intros.
-generalize (Fpmov_correct2 p rs x).
-destruct x; simpl.
-unfold get, Locmap.get; intros; auto.
-rewrite H4; auto.
-apply disjoint_tmp__noTmp; auto.
+
+Lemma effect_seqmove_equiv:
+ forall mu e1 e1',
+ effect_seqmove mu e1 e1' ->
+ forall e2,
+ (forall s d, In (s, d) mu ->
+ (In s srcs \/ s = R IT2 \/ s = R FT2) /\
+ (In d dsts \/ d = R IT2 \/ d = R FT2)) ->
+ locmap_equiv e1 e2 ->
+ locmap_equiv e1' (exec_seq mu e2).
+Proof.
+ induction 1; intros.
+ simpl. auto.
+ simpl. apply IHeffect_seqmove.
+ intros. apply H1. apply in_cons; auto.
+ destruct (H1 s d (in_eq _ _)).
+ eapply effect_move_equiv; eauto.
Qed.
-
-Theorem Fpmov_correct_ext:
- forall p rs,
- simpleDest p ->
- no_overlap_list p ->
- Loc.disjoint (getsrc p) temporaries ->
- Loc.disjoint (getdst p) temporaries ->
- forall l,
- Loc.notin l (getdst p) ->
- Loc.notin l temporaries ->
- (sexec (StateDone (Pmov (p, nil, nil))) rs) l = rs l.
-Proof.
-intros; generalize (Fpmov_correct1 p rs l); unfold get, Locmap.get; intros.
-rewrite <- H5; auto.
-apply disjoint_tmp__noTmp; auto.
-unfold notemporary; simpl in H4 |-; unfold T; intros x; case (Loc.type x).
-elim H4;
- [intros H6 H7; elim H7; [intros H8 H9; (try clear H7 H4); (try exact H8)]].
-elim H4;
- [intros H6 H7; elim H7;
- [intros H8 H9; elim H9;
- [intros H10 H11; elim H11;
- [intros H12 H13; elim H13;
- [intros H14 H15; (try clear H13 H11 H9 H7 H4); (try exact H14)]]]]].
-unfold no_overlap_list, no_overlap in H0 |-; intros.
-case (Loc.eq l x).
-intros e; left; (try assumption).
-intros n; right; (try assumption).
-apply Loc.in_notin_diff with (getdst p); auto.
-apply notindst_nW; auto.
+
+Lemma effect_parmove:
+ forall e e',
+ effect_seqmove (parmove srcs dsts) e e' ->
+ List.map e' dsts = List.map e srcs /\
+ e' (R IT3) = e (R IT3) /\
+ forall l, Loc.notin l dsts -> Loc.notin l temporaries -> e' l = e l.
+Proof.
+ set (mu := parmove srcs dsts). intros.
+ assert (locmap_equiv e e) by (red; auto).
+ generalize (effect_seqmove_equiv mu e e' H e (parmove_prop_2 srcs dsts) H0).
+ intro.
+ generalize (parmove_prop_1 srcs dsts LENGTH NOREPET NO_SRCS_TEMP NO_DSTS_TEMP e).
+ fold mu. intros [A B].
+ (* e' dsts = e srcs *)
+ split. rewrite <- A. apply list_map_exten; intros.
+ apply H1. apply dests_no_overlap_dests; auto.
+ apply NO_DSTS_TEMP; auto; simpl; tauto.
+ apply NO_DSTS_TEMP; auto; simpl; tauto.
+ (* e' IT3 = e IT3 *)
+ split.
+ assert (Loc.notin (R IT3) dsts).
+ apply Loc.disjoint_notin with temporaries.
+ apply Loc.disjoint_sym; auto. simpl; tauto.
+ transitivity (exec_seq mu e (R IT3)).
+ symmetry. apply H1. apply notin_dests_no_overlap_dests. auto.
+ simpl; congruence. simpl; congruence.
+ apply B. apply Loc.notin_not_in; auto. congruence. congruence.
+ (* other locations *)
+ intros. transitivity (exec_seq mu e l).
+ symmetry. apply H1. apply notin_dests_no_overlap_dests; auto.
+ eapply Loc.in_notin_diff; eauto. simpl; tauto.
+ eapply Loc.in_notin_diff; eauto. simpl; tauto.
+ apply B. apply Loc.notin_not_in; auto.
+ apply Loc.diff_not_eq. eapply Loc.in_notin_diff; eauto. simpl; tauto.
+ apply Loc.diff_not_eq. eapply Loc.in_notin_diff; eauto. simpl; tauto.
Qed.
+
+End EQUIVALENCE.
+
diff --git a/backend/RTL.v b/backend/RTL.v
index ac9a4159..4a3f8e8c 100644
--- a/backend/RTL.v
+++ b/backend/RTL.v
@@ -4,12 +4,13 @@
after Cminor.
*)
-Require Import Relations.
+(*Require Import Relations.*)
Require Import Coqlib.
Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Values.
+Require Import Events.
Require Import Mem.
Require Import Globalenvs.
Require Import Op.
@@ -54,6 +55,10 @@ Inductive instruction: Set :=
function name), giving it the values of registers [args]
as arguments. It stores the return value in [dest] and branches
to [succ]. *)
+ | Ialloc: reg -> reg -> node -> instruction
+ (** [Ialloc arg dest succ] allocates a fresh block of size
+ the contents of register [arg], stores a pointer to this
+ block in [dest], and branches to [succ]. *)
| Icond: condition -> list reg -> node -> node -> instruction
(** [Icond cond args ifso ifnot] evaluates the boolean condition
[cond] over the values of registers [args]. If the condition
@@ -85,11 +90,19 @@ Record function: Set := mkfunction {
in the CFG. [fn_code_wf] asserts that all instructions of the function
have nodes no greater than [fn_nextpc]. *)
-Definition program := AST.program function.
+Definition fundef := AST.fundef function.
+
+Definition program := AST.program fundef.
+
+Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => f.(fn_sig)
+ | External ef => ef.(ef_sig)
+ end.
(** * Operational semantics *)
-Definition genv := Genv.t function.
+Definition genv := Genv.t fundef.
Definition regset := Regmap.t val.
Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
@@ -102,7 +115,8 @@ Section RELSEM.
Variable ge: genv.
-Definition find_function (ros: reg + ident) (rs: regset) : option function :=
+Definition find_function
+ (ros: reg + ident) (rs: regset) : option fundef :=
match ros with
| inl r => Genv.find_funct ge rs#r
| inr symb =>
@@ -132,66 +146,73 @@ Definition find_function (ros: reg + ident) (rs: regset) : option function :=
and memory state [m]. The final state is [pc'], [rs'] and [m']. *)
Inductive exec_instr: code -> val ->
- node -> regset -> mem ->
+ node -> regset -> mem -> trace ->
node -> regset -> mem -> Prop :=
| exec_Inop:
forall c sp pc rs m pc',
c!pc = Some(Inop pc') ->
- exec_instr c sp pc rs m pc' rs m
+ exec_instr c sp pc rs m E0 pc' rs m
| exec_Iop:
forall c sp pc rs m op args res pc' v,
c!pc = Some(Iop op args res pc') ->
eval_operation ge sp op rs##args = Some v ->
- exec_instr c sp pc rs m pc' (rs#res <- v) m
+ exec_instr c sp pc rs m E0 pc' (rs#res <- v) m
| exec_Iload:
forall c sp pc rs m chunk addr args dst pc' a v,
c!pc = Some(Iload chunk addr args dst pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
- exec_instr c sp pc rs m pc' (rs#dst <- v) m
+ exec_instr c sp pc rs m E0 pc' (rs#dst <- v) m
| exec_Istore:
forall c sp pc rs m chunk addr args src pc' a m',
c!pc = Some(Istore chunk addr args src pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.storev chunk m a rs#src = Some m' ->
- exec_instr c sp pc rs m pc' rs m'
+ exec_instr c sp pc rs m E0 pc' rs m'
| exec_Icall:
- forall c sp pc rs m sig ros args res pc' f vres m',
+ forall c sp pc rs m sig ros args res pc' f vres m' t,
c!pc = Some(Icall sig ros args res pc') ->
find_function ros rs = Some f ->
- sig = f.(fn_sig) ->
- exec_function f rs##args m vres m' ->
- exec_instr c sp pc rs m pc' (rs#res <- vres) m'
+ funsig f = sig ->
+ exec_function f rs##args m t vres m' ->
+ exec_instr c sp pc rs m t pc' (rs#res <- vres) m'
+ | exec_Ialloc:
+ forall c sp pc rs m pc' arg res sz m' b,
+ c!pc = Some(Ialloc arg res pc') ->
+ rs#arg = Vint sz ->
+ Mem.alloc m 0 (Int.signed sz) = (m', b) ->
+ exec_instr c sp pc rs m E0 pc' (rs#res <- (Vptr b Int.zero)) m'
| exec_Icond_true:
forall c sp pc rs m cond args ifso ifnot,
c!pc = Some(Icond cond args ifso ifnot) ->
eval_condition cond rs##args = Some true ->
- exec_instr c sp pc rs m ifso rs m
+ exec_instr c sp pc rs m E0 ifso rs m
| exec_Icond_false:
forall c sp pc rs m cond args ifso ifnot,
c!pc = Some(Icond cond args ifso ifnot) ->
eval_condition cond rs##args = Some false ->
- exec_instr c sp pc rs m ifnot rs m
+ exec_instr c sp pc rs m E0 ifnot rs m
(** [exec_instrs ge c sp pc rs m pc' rs' m'] is the reflexive
transitive closure of [exec_instr]. It corresponds to the execution
of zero, one or finitely many transitions. *)
with exec_instrs: code -> val ->
- node -> regset -> mem ->
+ node -> regset -> mem -> trace ->
node -> regset -> mem -> Prop :=
| exec_refl:
forall c sp pc rs m,
- exec_instrs c sp pc rs m pc rs m
+ exec_instrs c sp pc rs m E0 pc rs m
| exec_one:
- forall c sp pc rs m pc' rs' m',
- exec_instr c sp pc rs m pc' rs' m' ->
- exec_instrs c sp pc rs m pc' rs' m'
+ forall c sp pc rs m t pc' rs' m',
+ exec_instr c sp pc rs m t pc' rs' m' ->
+ exec_instrs c sp pc rs m t pc' rs' m'
| exec_trans:
- forall c sp pc1 rs1 m1 pc2 rs2 m2 pc3 rs3 m3,
- exec_instrs c sp pc1 rs1 m1 pc2 rs2 m2 ->
- exec_instrs c sp pc2 rs2 m2 pc3 rs3 m3 ->
- exec_instrs c sp pc1 rs1 m1 pc3 rs3 m3
+ forall c sp pc1 rs1 m1 t1 pc2 rs2 m2 t2 pc3 rs3 m3 t3,
+ exec_instrs c sp pc1 rs1 m1 t1 pc2 rs2 m2 ->
+ exec_instrs c sp pc2 rs2 m2 t2 pc3 rs3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_instrs c sp pc1 rs1 m1 t3 pc3 rs3 m3
(** [exec_function ge f args m res m'] executes a function application.
[f] is the called function, [args] the values of its arguments,
@@ -205,17 +226,21 @@ with exec_instrs: code -> val ->
(Non-parameter registers are initialized to [Vundef].) Before returning,
the stack activation block is freed. *)
-with exec_function: function -> list val -> mem ->
- val -> mem -> Prop :=
- | exec_funct:
- forall f m m1 stk args pc rs m2 or vres,
+with exec_function: fundef -> list val -> mem -> trace ->
+ val -> mem -> Prop :=
+ | exec_funct_internal:
+ forall f m m1 stk args t pc rs m2 or vres,
Mem.alloc m 0 f.(fn_stacksize) = (m1, stk) ->
- exec_instrs f.(fn_code) (Vptr stk Int.zero)
+ exec_instrs f.(fn_code) (Vptr stk Int.zero)
f.(fn_entrypoint) (init_regs args f.(fn_params)) m1
- pc rs m2 ->
+ t pc rs m2 ->
f.(fn_code)!pc = Some(Ireturn or) ->
vres = regmap_optget or Vundef rs ->
- exec_function f args m vres (Mem.free m2 stk).
+ exec_function (Internal f) args m t vres (Mem.free m2 stk)
+ | exec_funct_external:
+ forall ef args m t res,
+ event_match ef args t res ->
+ exec_function (External ef) args m t res m.
Scheme exec_instr_ind_3 := Minimality for exec_instr Sort Prop
with exec_instrs_ind_3 := Minimality for exec_instrs Sort Prop
@@ -224,12 +249,13 @@ Scheme exec_instr_ind_3 := Minimality for exec_instr Sort Prop
(** Some derived execution rules. *)
Lemma exec_step:
- forall c sp pc1 rs1 m1 pc2 rs2 m2 pc3 rs3 m3,
- exec_instr c sp pc1 rs1 m1 pc2 rs2 m2 ->
- exec_instrs c sp pc2 rs2 m2 pc3 rs3 m3 ->
- exec_instrs c sp pc1 rs1 m1 pc3 rs3 m3.
+ forall c sp pc1 rs1 m1 t1 pc2 rs2 m2 t2 pc3 rs3 m3 t3,
+ exec_instr c sp pc1 rs1 m1 t1 pc2 rs2 m2 ->
+ exec_instrs c sp pc2 rs2 m2 t2 pc3 rs3 m3 ->
+ t3 = t1 ** t2 ->
+ exec_instrs c sp pc1 rs1 m1 t3 pc3 rs3 m3.
Proof.
- intros. eapply exec_trans. apply exec_one. eauto. eauto.
+ intros. eapply exec_trans. apply exec_one. eauto. eauto. auto.
Qed.
Lemma exec_Iop':
@@ -237,7 +263,7 @@ Lemma exec_Iop':
c!pc = Some(Iop op args res pc') ->
eval_operation ge sp op rs##args = Some v ->
rs' = (rs#res <- v) ->
- exec_instr c sp pc rs m pc' rs' m.
+ exec_instr c sp pc rs m E0 pc' rs' m.
Proof.
intros. subst rs'. eapply exec_Iop; eauto.
Qed.
@@ -248,7 +274,7 @@ Lemma exec_Iload':
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = (rs#dst <- v) ->
- exec_instr c sp pc rs m pc' rs' m.
+ exec_instr c sp pc rs m E0 pc' rs' m.
Proof.
intros. subst rs'. eapply exec_Iload; eauto.
Qed.
@@ -257,16 +283,16 @@ Qed.
is defined in the CFG. *)
Lemma exec_instr_present:
- forall c sp pc rs m pc' rs' m',
- exec_instr c sp pc rs m pc' rs' m' ->
+ forall c sp pc rs m t pc' rs' m',
+ exec_instr c sp pc rs m t pc' rs' m' ->
c!pc <> None.
Proof.
induction 1; congruence.
Qed.
Lemma exec_instrs_present:
- forall c sp pc rs m pc' rs' m',
- exec_instrs c sp pc rs m pc' rs' m' ->
+ forall c sp pc rs m t pc' rs' m',
+ exec_instrs c sp pc rs m t pc' rs' m' ->
c!pc' <> None -> c!pc <> None.
Proof.
induction 1; intros.
@@ -280,14 +306,14 @@ End RELSEM.
(** Execution of whole programs. As in Cminor, we call the ``main'' function
with no arguments and observe its return value. *)
-Definition exec_program (p: program) (r: val) : Prop :=
+Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
let ge := Genv.globalenv p in
let m0 := Genv.init_mem p in
exists b, exists f, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\
- f.(fn_sig) = mksignature nil (Some Tint) /\
- exec_function ge f nil m0 r m.
+ funsig f = mksignature nil (Some Tint) /\
+ exec_function ge f nil m0 t r m.
(** * Operations on RTL abstract syntax *)
@@ -304,14 +330,15 @@ Definition successors (f: function) (pc: node) : list node :=
| Iload chunk addr args dst s => s :: nil
| Istore chunk addr args src s => s :: nil
| Icall sig ros args res s => s :: nil
+ | Ialloc arg res s => s :: nil
| Icond cond args ifso ifnot => ifso :: ifnot :: nil
| Ireturn optarg => nil
end
end.
Lemma successors_correct:
- forall ge f sp pc rs m pc' rs' m',
- exec_instr ge f.(fn_code) sp pc rs m pc' rs' m' ->
+ forall ge f sp pc rs m t pc' rs' m',
+ exec_instr ge f.(fn_code) sp pc rs m t pc' rs' m' ->
In pc' (successors f pc).
Proof.
intros ge f. unfold successors. generalize (fn_code f).
@@ -345,5 +372,5 @@ Definition transf_function (f: function) : function :=
f.(fn_entrypoint)
f.(fn_nextpc)
(transf_code_wf f.(fn_code) f.(fn_nextpc) f.(fn_code_wf)).
-
+
End TRANSF.
diff --git a/backend/RTLgen.v b/backend/RTLgen.v
index 38b19a01..a5c3ae7a 100644
--- a/backend/RTLgen.v
+++ b/backend/RTLgen.v
@@ -28,6 +28,7 @@ Fixpoint mutated_expr (a: expr) : list ident :=
| Econdition b c d => mutated_condexpr b ++ mutated_expr c ++ mutated_expr d
| Elet b c => mutated_expr b ++ mutated_expr c
| Eletvar n => nil
+ | Ealloc b => mutated_expr b
end
with mutated_condexpr (a: condexpr) : list ident :=
@@ -328,6 +329,10 @@ Fixpoint transl_expr (map: mapping) (mut: list ident)
transl_expr map mut b r nc
| Eletvar n =>
do r <- find_letvar map n; add_move r rd nd
+ | Ealloc a =>
+ do r <- alloc_reg map mut a;
+ do no <- add_instr (Ialloc r rd nd);
+ transl_expr map mut a r no
end
(** Translation of a conditional expression. Similar to [transl_expr],
@@ -473,16 +478,18 @@ Definition transl_function (f: Cminor.function) : option RTL.function :=
| Error => None
| OK (nentry, rparams) s =>
Some (RTL.mkfunction
- f.(Cminor.fn_sig)
- rparams
- f.(Cminor.fn_stackspace)
- s.(st_code)
- nentry
- s.(st_nextnode)
- s.(st_wf))
+ f.(Cminor.fn_sig)
+ rparams
+ f.(Cminor.fn_stackspace)
+ s.(st_code)
+ nentry
+ s.(st_nextnode)
+ s.(st_wf))
end.
+Definition transl_fundef := transf_partial_fundef transl_function.
+
(** Translation of a whole program. *)
Definition transl_program (p: Cminor.program) : option RTL.program :=
- transform_partial_program transl_function p.
+ transform_partial_program transl_fundef p.
diff --git a/backend/RTLgenproof.v b/backend/RTLgenproof.v
index d34bae96..24cc41b4 100644
--- a/backend/RTLgenproof.v
+++ b/backend/RTLgenproof.v
@@ -6,6 +6,7 @@ Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Registers.
@@ -30,45 +31,59 @@ Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
intro. unfold ge, tge.
- apply Genv.find_symbol_transf_partial with transl_function.
+ apply Genv.find_symbol_transf_partial with transl_fundef.
exact TRANSL.
Qed.
Lemma function_ptr_translated:
- forall (b: block) (f: Cminor.function),
+ forall (b: block) (f: Cminor.fundef),
Genv.find_funct_ptr ge b = Some f ->
exists tf,
- Genv.find_funct_ptr tge b = Some tf /\ transl_function f = Some tf.
+ Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Some tf.
Proof.
intros.
generalize
- (Genv.find_funct_ptr_transf_partial transl_function TRANSL H).
- case (transl_function f).
+ (Genv.find_funct_ptr_transf_partial transl_fundef TRANSL H).
+ case (transl_fundef f).
intros tf [A B]. exists tf. tauto.
intros [A B]. elim B. reflexivity.
Qed.
Lemma functions_translated:
- forall (v: val) (f: Cminor.function),
+ forall (v: val) (f: Cminor.fundef),
Genv.find_funct ge v = Some f ->
exists tf,
- Genv.find_funct tge v = Some tf /\ transl_function f = Some tf.
+ Genv.find_funct tge v = Some tf /\ transl_fundef f = Some tf.
Proof.
intros.
generalize
- (Genv.find_funct_transf_partial transl_function TRANSL H).
- case (transl_function f).
+ (Genv.find_funct_transf_partial transl_fundef TRANSL H).
+ case (transl_fundef f).
intros tf [A B]. exists tf. tauto.
intros [A B]. elim B. reflexivity.
Qed.
+Lemma sig_transl_function:
+ forall (f: Cminor.fundef) (tf: RTL.fundef),
+ transl_fundef f = Some tf ->
+ RTL.funsig tf = Cminor.funsig f.
+Proof.
+ intros until tf. unfold transl_fundef, transf_partial_fundef.
+ case f; intro.
+ unfold transl_function.
+ case (transl_fun f0 init_state); intros.
+ discriminate.
+ destruct p. inversion H. reflexivity.
+ intro. inversion H. reflexivity.
+Qed.
+
(** Correctness of the code generated by [add_move]. *)
Lemma add_move_correct:
forall r1 r2 sp nd s ns s' rs m,
add_move r1 r2 nd s = OK ns s' ->
exists rs',
- exec_instrs tge s'.(st_code) sp ns rs m nd rs' m /\
+ exec_instrs tge s'.(st_code) sp ns rs m E0 nd rs' m /\
rs'#r2 = rs#r1 /\
(forall r, r <> r2 -> rs'#r = rs#r).
Proof.
@@ -118,7 +133,7 @@ Qed.
Definition transl_expr_correct
(sp: val) (le: letenv) (e: env) (m: mem) (a: expr)
- (e': env) (m': mem) (v: val) : Prop :=
+ (t: trace) (e': env) (m': mem) (v: val) : Prop :=
forall map mut rd nd s ns s' rs
(MWF: map_wf map s)
(TE: transl_expr map mut a rd nd s = OK ns s')
@@ -126,7 +141,7 @@ Definition transl_expr_correct
(MUT: incl (mutated_expr a) mut)
(TRG: target_reg_ok s map mut a rd),
exists rs',
- exec_instrs tge s'.(st_code) sp ns rs m nd rs' m'
+ exec_instrs tge s'.(st_code) sp ns rs m t nd rs' m'
/\ match_env map e' le rs'
/\ rs'#rd = v
/\ (forall r,
@@ -139,7 +154,7 @@ Definition transl_expr_correct
Definition transl_exprlist_correct
(sp: val) (le: letenv) (e: env) (m: mem) (al: exprlist)
- (e': env) (m': mem) (vl: list val) : Prop :=
+ (t: trace) (e': env) (m': mem) (vl: list val) : Prop :=
forall map mut rl nd s ns s' rs
(MWF: map_wf map s)
(TE: transl_exprlist map mut al rl nd s = OK ns s')
@@ -147,7 +162,7 @@ Definition transl_exprlist_correct
(MUT: incl (mutated_exprlist al) mut)
(TRG: target_regs_ok s map mut al rl),
exists rs',
- exec_instrs tge s'.(st_code) sp ns rs m nd rs' m'
+ exec_instrs tge s'.(st_code) sp ns rs m t nd rs' m'
/\ match_env map e' le rs'
/\ rs'##rl = vl
/\ (forall r,
@@ -157,14 +172,14 @@ Definition transl_exprlist_correct
Definition transl_condition_correct
(sp: val) (le: letenv) (e: env) (m: mem) (a: condexpr)
- (e': env) (m': mem) (vb: bool) : Prop :=
+ (t: trace) (e': env) (m': mem) (vb: bool) : Prop :=
forall map mut ntrue nfalse s ns s' rs
(MWF: map_wf map s)
(TE: transl_condition map mut a ntrue nfalse s = OK ns s')
(ME: match_env map e le rs)
(MUT: incl (mutated_condexpr a) mut),
exists rs',
- exec_instrs tge s'.(st_code) sp ns rs m (if vb then ntrue else nfalse) rs' m'
+ exec_instrs tge s'.(st_code) sp ns rs m t (if vb then ntrue else nfalse) rs' m'
/\ match_env map e' le rs'
/\ (forall r,
reg_valid r s -> ~(mutated_reg map mut r) ->
@@ -233,7 +248,7 @@ Definition match_return_outcome
Definition transl_stmt_correct
(sp: val) (e: env) (m: mem) (a: stmt)
- (e': env) (m': mem) (out: outcome) : Prop :=
+ (t: trace) (e': env) (m': mem) (out: outcome) : Prop :=
forall map ncont nexits nret rret s ns s' nd rs
(MWF: map_wf map s)
(TE: transl_stmt map a ncont nexits nret rret s = OK ns s')
@@ -241,7 +256,7 @@ Definition transl_stmt_correct
(OUT: outcome_node out ncont nexits nret nd)
(RRG: return_reg_ok s map rret),
exists rs',
- exec_instrs tge s'.(st_code) sp ns rs m nd rs' m'
+ exec_instrs tge s'.(st_code) sp ns rs m t nd rs' m'
/\ match_env map e' nil rs'
/\ match_return_outcome out rret rs'.
@@ -251,11 +266,11 @@ Definition transl_stmt_correct
the same modifications on the memory state. *)
Definition transl_function_correct
- (m: mem) (f: Cminor.function) (vargs: list val)
- (m':mem) (vres: val) : Prop :=
+ (m: mem) (f: Cminor.fundef) (vargs: list val)
+ (t: trace) (m':mem) (vres: val) : Prop :=
forall tf
- (TE: transl_function f = Some tf),
- exec_function tge tf vargs m vres m'.
+ (TE: transl_fundef f = Some tf),
+ exec_function tge tf vargs m t vres m'.
(** The correctness of the translation is a huge induction over
the Cminor evaluation derivation for the source program. To keep
@@ -268,7 +283,7 @@ Definition transl_function_correct
Lemma transl_expr_Evar_correct:
forall (sp: val) (le: letenv) (e: env) (m: mem) (id: ident) (v: val),
e!id = Some v ->
- transl_expr_correct sp le e m (Evar id) e m v.
+ transl_expr_correct sp le e m (Evar id) E0 e m v.
Proof.
intros; red; intros. monadInv TE. intro.
generalize EQ; unfold find_var. caseEq (map_vars map)!id.
@@ -303,12 +318,12 @@ Qed.
Lemma transl_expr_Eop_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem) (op : operation)
- (al : exprlist) (e1 : env) (m1 : mem) (vl : list val)
+ (al : exprlist) (t: trace) (e1 : env) (m1 : mem) (vl : list val)
(v: val),
- eval_exprlist ge sp le e m al e1 m1 vl ->
- transl_exprlist_correct sp le e m al e1 m1 vl ->
+ eval_exprlist ge sp le e m al t e1 m1 vl ->
+ transl_exprlist_correct sp le e m al t e1 m1 vl ->
eval_operation ge sp op vl = Some v ->
- transl_expr_correct sp le e m (Eop op al) e1 m1 v.
+ transl_expr_correct sp le e m (Eop op al) t e1 m1 v.
Proof.
intros until v. intros EEL TEL EOP. red; intros.
simpl in TE. monadInv TE. intro EQ1.
@@ -323,9 +338,9 @@ Proof.
apply exec_instrs_incr with s1. eauto with rtlg.
apply exec_one; eapply exec_Iop. eauto with rtlg.
subst vl.
- rewrite (@eval_operation_preserved Cminor.function RTL.function ge tge).
+ rewrite (@eval_operation_preserved Cminor.fundef RTL.fundef ge tge).
eexact EOP.
- exact symbols_preserved.
+ exact symbols_preserved. traceEq.
(* Match-env *)
split. inversion TRG. eauto with rtlg.
(* Result reg *)
@@ -344,11 +359,11 @@ Qed.
Lemma transl_expr_Eassign_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (id : ident) (a : expr) (e1 : env) (m1 : mem)
+ (id : ident) (a : expr) (t: trace) (e1 : env) (m1 : mem)
(v : val),
- eval_expr ge sp le e m a e1 m1 v ->
- transl_expr_correct sp le e m a e1 m1 v ->
- transl_expr_correct sp le e m (Eassign id a) (PTree.set id v e1) m1 v.
+ eval_expr ge sp le e m a t e1 m1 v ->
+ transl_expr_correct sp le e m a t e1 m1 v ->
+ transl_expr_correct sp le e m (Eassign id a) t (PTree.set id v e1) m1 v.
Proof.
intros; red; intros.
simpl in TE. monadInv TE. intro EQ1.
@@ -366,7 +381,7 @@ Proof.
(* Exec *)
split. eapply exec_trans. eexact EX1.
apply exec_instrs_incr with s1. eauto with rtlg.
- exact EX2.
+ eexact EX2. traceEq.
(* Match-env *)
split.
apply match_env_update_var with rs1 r s s0; auto.
@@ -387,13 +402,13 @@ Qed.
Lemma transl_expr_Eload_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
(chunk : memory_chunk) (addr : addressing)
- (al : exprlist) (e1 : env) (m1 : mem) (v : val)
+ (al : exprlist) (t: trace) (e1 : env) (m1 : mem) (v : val)
(vl : list val) (a: val),
- eval_exprlist ge sp le e m al e1 m1 vl ->
- transl_exprlist_correct sp le e m al e1 m1 vl ->
+ eval_exprlist ge sp le e m al t e1 m1 vl ->
+ transl_exprlist_correct sp le e m al t e1 m1 vl ->
eval_addressing ge sp addr vl = Some a ->
Mem.loadv chunk m1 a = Some v ->
- transl_expr_correct sp le e m (Eload chunk addr al) e1 m1 v.
+ transl_expr_correct sp le e m (Eload chunk addr al) t e1 m1 v.
Proof.
intros; red; intros. simpl in TE. monadInv TE. intro EQ1. clear TE.
simpl in MUT.
@@ -407,7 +422,7 @@ Proof.
apply exec_instrs_incr with s1. eauto with rtlg.
apply exec_one. eapply exec_Iload. eauto with rtlg.
rewrite RES1. rewrite (@eval_addressing_preserved _ _ ge tge).
- eexact H1. exact symbols_preserved. assumption.
+ eexact H1. exact symbols_preserved. assumption. traceEq.
(* Match-env *)
split. eapply match_env_update_temp. assumption. inversion TRG. assumption.
(* Result *)
@@ -425,16 +440,17 @@ Qed.
Lemma transl_expr_Estore_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
(chunk : memory_chunk) (addr : addressing)
- (al : exprlist) (b : expr) (e1 : env) (m1 : mem)
- (vl : list val) (e2 : env) (m2 m3 : mem)
+ (al : exprlist) (b : expr) (t t1: trace) (e1 : env) (m1 : mem)
+ (vl : list val) (t2: trace) (e2 : env) (m2 m3 : mem)
(v : val) (a: val),
- eval_exprlist ge sp le e m al e1 m1 vl ->
- transl_exprlist_correct sp le e m al e1 m1 vl ->
- eval_expr ge sp le e1 m1 b e2 m2 v ->
- transl_expr_correct sp le e1 m1 b e2 m2 v ->
+ eval_exprlist ge sp le e m al t1 e1 m1 vl ->
+ transl_exprlist_correct sp le e m al t1 e1 m1 vl ->
+ eval_expr ge sp le e1 m1 b t2 e2 m2 v ->
+ transl_expr_correct sp le e1 m1 b t2 e2 m2 v ->
eval_addressing ge sp addr vl = Some a ->
Mem.storev chunk m2 a v = Some m3 ->
- transl_expr_correct sp le e m (Estore chunk addr al b) e2 m3 v.
+ t = t1 ** t2 ->
+ transl_expr_correct sp le e m (Estore chunk addr al b) t e2 m3 v.
Proof.
intros; red; intros. simpl in TE; monadInv TE. intro EQ2; clear TE.
simpl in MUT.
@@ -467,10 +483,11 @@ Proof.
tauto. right. apply sym_not_equal.
apply valid_fresh_different with s. inversion TRG; assumption.
assumption.
- rewrite H5; rewrite RES1.
+ rewrite H6; rewrite RES1.
rewrite (@eval_addressing_preserved _ _ ge tge).
eexact H3. exact symbols_preserved.
rewrite RES2. assumption.
+ reflexivity. traceEq.
(* Match-env *)
split. assumption.
(* Result *)
@@ -488,18 +505,20 @@ Qed.
Lemma transl_expr_Ecall_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (sig : signature) (a : expr) (bl : exprlist)
- (e1 e2 : env) (m1 m2 m3 : mem) (vf : val)
- (vargs : list val) (vres : val) (f : Cminor.function),
- eval_expr ge sp le e m a e1 m1 vf ->
- transl_expr_correct sp le e m a e1 m1 vf ->
- eval_exprlist ge sp le e1 m1 bl e2 m2 vargs ->
- transl_exprlist_correct sp le e1 m1 bl e2 m2 vargs ->
+ (sig : signature) (a : expr) (bl : exprlist) (t t1: trace)
+ (e1: env) (m1: mem) (t2: trace) (e2 : env) (m2 : mem)
+ (t3: trace) (m3: mem) (vf : val)
+ (vargs : list val) (vres : val) (f : Cminor.fundef),
+ eval_expr ge sp le e m a t1 e1 m1 vf ->
+ transl_expr_correct sp le e m a t1 e1 m1 vf ->
+ eval_exprlist ge sp le e1 m1 bl t2 e2 m2 vargs ->
+ transl_exprlist_correct sp le e1 m1 bl t2 e2 m2 vargs ->
Genv.find_funct ge vf = Some f ->
- Cminor.fn_sig f = sig ->
- eval_funcall ge m2 f vargs m3 vres ->
- transl_function_correct m2 f vargs m3 vres ->
- transl_expr_correct sp le e m (Ecall sig a bl) e2 m3 vres.
+ Cminor.funsig f = sig ->
+ eval_funcall ge m2 f vargs t3 m3 vres ->
+ transl_function_correct m2 f vargs t3 m3 vres ->
+ t = t1 ** t2 ** t3 ->
+ transl_expr_correct sp le e m (Ecall sig a bl) t e2 m3 vres.
Proof.
intros. red; simpl; intros.
monadInv TE. intro EQ3. clear TE.
@@ -530,7 +549,7 @@ Proof.
generalize (H6 tf TF). intro EXF.
assert (EX3: exec_instrs tge (st_code s2) sp n rs2 m2
- nd (rs2#rd <- vres) m3).
+ t3 nd (rs2#rd <- vres) m3).
apply exec_one. eapply exec_Icall.
eauto with rtlg. simpl. replace (rs2#r) with vf. eexact TFIND.
rewrite <- RES1. symmetry. apply OTHER2.
@@ -543,10 +562,7 @@ Proof.
tauto. byContradiction. apply valid_fresh_absurd with r s0.
eauto with rtlg. assumption.
tauto.
- generalize TF. unfold transl_function.
- destruct (transl_fun f init_state).
- intro; discriminate. destruct p. intros. injection TF0. intro.
- rewrite <- H7; simpl. auto.
+ generalize (sig_transl_function _ _ TF). congruence.
rewrite RES2. assumption.
exists (rs2#rd <- vres).
@@ -555,7 +571,7 @@ Proof.
apply exec_instrs_incr with s3. eauto with rtlg.
eapply exec_trans. eexact EX2.
apply exec_instrs_incr with s2. eauto with rtlg.
- exact EX3.
+ eexact EX3. reflexivity. traceEq.
(* Match env *)
split. apply match_env_update_temp. assumption.
inversion TRG. assumption.
@@ -591,13 +607,14 @@ Qed.
Lemma transl_expr_Econdition_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (a : condexpr) (b c : expr) (e1 : env) (m1 : mem)
- (v1 : bool) (e2 : env) (m2 : mem) (v2 : val),
- eval_condexpr ge sp le e m a e1 m1 v1 ->
- transl_condition_correct sp le e m a e1 m1 v1 ->
- eval_expr ge sp le e1 m1 (if v1 then b else c) e2 m2 v2 ->
- transl_expr_correct sp le e1 m1 (if v1 then b else c) e2 m2 v2 ->
- transl_expr_correct sp le e m (Econdition a b c) e2 m2 v2.
+ (a : condexpr) (b c : expr) (t t1: trace) (e1 : env) (m1 : mem)
+ (v1 : bool) (t2: trace) (e2 : env) (m2 : mem) (v2 : val),
+ eval_condexpr ge sp le e m a t1 e1 m1 v1 ->
+ transl_condition_correct sp le e m a t1 e1 m1 v1 ->
+ eval_expr ge sp le e1 m1 (if v1 then b else c) t2 e2 m2 v2 ->
+ transl_expr_correct sp le e1 m1 (if v1 then b else c) t2 e2 m2 v2 ->
+ t = t1 ** t2 ->
+ transl_expr_correct sp le e m (Econdition a b c) t e2 m2 v2.
Proof.
intros; red; intros. simpl in TE; monadInv TE. intro EQ1; clear TE.
simpl in MUT.
@@ -619,7 +636,7 @@ Proof.
(* Exec *)
split. eapply exec_trans. eexact EX1.
apply exec_instrs_incr with s1. eauto with rtlg.
- exact EX2.
+ eexact EX2. auto.
(* Match-env *)
split. assumption.
(* Result value *)
@@ -639,7 +656,7 @@ Proof.
(* Exec *)
split. eapply exec_trans. eexact EX1.
apply exec_instrs_incr with s0. eauto with rtlg.
- exact EX2.
+ eexact EX2. auto.
(* Match-env *)
split. assumption.
(* Result value *)
@@ -653,13 +670,14 @@ Qed.
Lemma transl_expr_Elet_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (a b : expr) (e1 : env) (m1 : mem) (v1 : val)
- (e2 : env) (m2 : mem) (v2 : val),
- eval_expr ge sp le e m a e1 m1 v1 ->
- transl_expr_correct sp le e m a e1 m1 v1 ->
- eval_expr ge sp (v1 :: le) e1 m1 b e2 m2 v2 ->
- transl_expr_correct sp (v1 :: le) e1 m1 b e2 m2 v2 ->
- transl_expr_correct sp le e m (Elet a b) e2 m2 v2.
+ (a b : expr) (t t1: trace) (e1 : env) (m1 : mem) (v1 : val)
+ (t2: trace) (e2 : env) (m2 : mem) (v2 : val),
+ eval_expr ge sp le e m a t1 e1 m1 v1 ->
+ transl_expr_correct sp le e m a t1 e1 m1 v1 ->
+ eval_expr ge sp (v1 :: le) e1 m1 b t2 e2 m2 v2 ->
+ transl_expr_correct sp (v1 :: le) e1 m1 b t2 e2 m2 v2 ->
+ t = t1 ** t2 ->
+ transl_expr_correct sp le e m (Elet a b) t e2 m2 v2.
Proof.
intros; red; intros.
simpl in TE; monadInv TE. intro EQ1.
@@ -678,17 +696,17 @@ Proof.
assert (TRG2: target_reg_ok s0 (add_letvar map r) mut b rd).
inversion TRG. apply target_reg_other.
unfold reg_in_map, add_letvar; simpl. red; intro.
- elim H10; intro. apply H3. left. assumption.
- elim H11; intro. apply valid_fresh_absurd with rd s.
- assumption. rewrite <- H12. eauto with rtlg.
- apply H3. right. assumption.
+ elim H11; intro. apply H4. left. assumption.
+ elim H12; intro. apply valid_fresh_absurd with rd s.
+ assumption. rewrite <- H13. eauto with rtlg.
+ apply H4. right. assumption.
eauto with rtlg.
generalize (H2 _ _ _ _ _ _ _ _ MWF2 EQ0 ME2 MUT2 TRG2).
intros [rs2 [EX2 [ME3 [RES2 OTHER2]]]].
exists rs2.
(* Exec *)
split. eapply exec_trans. eexact EX1.
- apply exec_instrs_incr with s1. eauto with rtlg. exact EX2.
+ apply exec_instrs_incr with s1. eauto with rtlg. eexact EX2. auto.
(* Match-env *)
split. apply mk_match_env. exact (me_vars _ _ _ _ ME3).
generalize (me_letvars _ _ _ _ ME3).
@@ -699,21 +717,21 @@ Proof.
intros. transitivity (rs1#r0).
apply OTHER2. eauto with rtlg.
unfold mutated_reg. unfold add_letvar; simpl. assumption.
- elim H5; intro. left.
+ elim H6; intro. left.
unfold reg_in_map, add_letvar; simpl.
- unfold reg_in_map in H6; tauto.
+ unfold reg_in_map in H7; tauto.
tauto.
apply OTHER1. eauto with rtlg.
assumption.
right. red; intro. apply valid_fresh_absurd with r0 s.
- assumption. rewrite H6. eauto with rtlg.
+ assumption. rewrite H7. eauto with rtlg.
Qed.
Lemma transl_expr_Eletvar_correct:
forall (sp: val) (le : list val) (e : env)
(m : mem) (n : nat) (v : val),
nth_error le n = Some v ->
- transl_expr_correct sp le e m (Eletvar n) e m v.
+ transl_expr_correct sp le e m (Eletvar n) E0 e m v.
Proof.
intros; red; intros.
simpl in TE; monadInv TE. intro EQ1.
@@ -746,9 +764,46 @@ Proof.
intro; monadSimpl.
Qed.
+Lemma transl_expr_Ealloc_correct:
+ forall (sp: val) (le : letenv) (e : env) (m : mem)
+ (a : expr) (t: trace) (e1 : env) (m1 : mem) (n: int)
+ (m2: mem) (b: block),
+ eval_expr ge sp le e m a t e1 m1 (Vint n) ->
+ transl_expr_correct sp le e m a t e1 m1 (Vint n) ->
+ Mem.alloc m1 0 (Int.signed n) = (m2, b) ->
+ transl_expr_correct sp le e m (Ealloc a) t e1 m2 (Vptr b Int.zero).
+Proof.
+ intros until b; intros EE TEC ALLOC; red; intros.
+ simpl in TE. monadInv TE. intro EQ1.
+ simpl in MUT.
+ assert (TRG': target_reg_ok s1 map mut a r); eauto with rtlg.
+ assert (MWF': map_wf map s1). eauto with rtlg.
+ generalize (TEC _ _ _ _ _ _ _ _ MWF' EQ1 ME MUT TRG').
+ intros [rs1 [EX1 [ME1 [RR1 RO1]]]].
+ exists (rs1#rd <- (Vptr b Int.zero)).
+(* Exec *)
+ split. eapply exec_trans. eexact EX1.
+ apply exec_instrs_incr with s1. eauto with rtlg.
+ apply exec_one; eapply exec_Ialloc. eauto with rtlg.
+ eexact RR1. assumption. traceEq.
+(* Match-env *)
+ split. inversion TRG. eauto with rtlg.
+(* Result *)
+ split. apply Regmap.gss.
+(* Other regs *)
+ intros. rewrite Regmap.gso.
+ apply RO1. eauto with rtlg. assumption.
+ case (Reg.eq r0 r); intro.
+ subst r0. left. elim (alloc_reg_fresh_or_in_map _ _ _ _ _ _ MWF EQ); intro.
+ auto. byContradiction; eauto with rtlg.
+ right; assumption.
+ elim H1; intro. red; intro. subst r0. inversion TRG. contradiction.
+ auto.
+Qed.
+
Lemma transl_condition_CEtrue_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem),
- transl_condition_correct sp le e m CEtrue e m true.
+ transl_condition_correct sp le e m CEtrue E0 e m true.
Proof.
intros; red; intros. simpl in TE; monadInv TE. subst ns.
exists rs. split. apply exec_refl. split. auto. auto.
@@ -756,7 +811,7 @@ Qed.
Lemma transl_condition_CEfalse_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem),
- transl_condition_correct sp le e m CEfalse e m false.
+ transl_condition_correct sp le e m CEfalse E0 e m false.
Proof.
intros; red; intros. simpl in TE; monadInv TE. subst ns.
exists rs. split. apply exec_refl. split. auto. auto.
@@ -764,12 +819,12 @@ Qed.
Lemma transl_condition_CEcond_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (cond : condition) (al : exprlist) (e1 : env)
+ (cond : condition) (al : exprlist) (t1: trace) (e1 : env)
(m1 : mem) (vl : list val) (b : bool),
- eval_exprlist ge sp le e m al e1 m1 vl ->
- transl_exprlist_correct sp le e m al e1 m1 vl ->
+ eval_exprlist ge sp le e m al t1 e1 m1 vl ->
+ transl_exprlist_correct sp le e m al t1 e1 m1 vl ->
eval_condition cond vl = Some b ->
- transl_condition_correct sp le e m (CEcond cond al) e1 m1 b.
+ transl_condition_correct sp le e m (CEcond cond al) t1 e1 m1 b.
Proof.
intros; red; intros. simpl in TE; monadInv TE. intro EQ1; clear TE.
assert (MWF1: map_wf map s1). eauto with rtlg.
@@ -788,6 +843,7 @@ Proof.
rewrite RES1. assumption.
eapply exec_Icond_false. eauto with rtlg.
rewrite RES1. assumption.
+ traceEq.
(* Match-env *)
split. assumption.
(* Regs *)
@@ -800,13 +856,14 @@ Qed.
Lemma transl_condition_CEcondition_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (a b c : condexpr) (e1 : env) (m1 : mem)
- (vb1 : bool) (e2 : env) (m2 : mem) (vb2 : bool),
- eval_condexpr ge sp le e m a e1 m1 vb1 ->
- transl_condition_correct sp le e m a e1 m1 vb1 ->
- eval_condexpr ge sp le e1 m1 (if vb1 then b else c) e2 m2 vb2 ->
- transl_condition_correct sp le e1 m1 (if vb1 then b else c) e2 m2 vb2 ->
- transl_condition_correct sp le e m (CEcondition a b c) e2 m2 vb2.
+ (a b c : condexpr) (t t1: trace) (e1 : env) (m1 : mem)
+ (vb1 : bool) (t2: trace) (e2 : env) (m2 : mem) (vb2 : bool),
+ eval_condexpr ge sp le e m a t1 e1 m1 vb1 ->
+ transl_condition_correct sp le e m a t1 e1 m1 vb1 ->
+ eval_condexpr ge sp le e1 m1 (if vb1 then b else c) t2 e2 m2 vb2 ->
+ transl_condition_correct sp le e1 m1 (if vb1 then b else c) t2 e2 m2 vb2 ->
+ t = t1 ** t2 ->
+ transl_condition_correct sp le e m (CEcondition a b c) t e2 m2 vb2.
Proof.
intros; red; intros. simpl in TE; monadInv TE. intro EQ1; clear TE.
simpl in MUT.
@@ -823,7 +880,7 @@ Proof.
exists rs2.
split. eapply exec_trans. eexact EX1.
apply exec_instrs_incr with s1. eauto with rtlg.
- exact EX2.
+ eexact EX2. auto.
split. assumption.
intros. transitivity (rs1#r).
apply OTHER2; eauto with rtlg.
@@ -835,7 +892,7 @@ Proof.
exists rs2.
split. eapply exec_trans. eexact EX1.
apply exec_instrs_incr with s0. eauto with rtlg.
- exact EX2.
+ eexact EX2. auto.
split. assumption.
intros. transitivity (rs1#r).
apply OTHER2; eauto with rtlg.
@@ -844,7 +901,7 @@ Qed.
Lemma transl_exprlist_Enil_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem),
- transl_exprlist_correct sp le e m Enil e m nil.
+ transl_exprlist_correct sp le e m Enil E0 e m nil.
Proof.
intros; red; intros.
generalize TE. simpl. destruct rl; monadSimpl.
@@ -857,17 +914,18 @@ Qed.
Lemma transl_exprlist_Econs_correct:
forall (sp: val) (le : letenv) (e : env) (m : mem)
- (a : expr) (bl : exprlist) (e1 : env) (m1 : mem)
- (v : val) (e2 : env) (m2 : mem) (vl : list val),
- eval_expr ge sp le e m a e1 m1 v ->
- transl_expr_correct sp le e m a e1 m1 v ->
- eval_exprlist ge sp le e1 m1 bl e2 m2 vl ->
- transl_exprlist_correct sp le e1 m1 bl e2 m2 vl ->
- transl_exprlist_correct sp le e m (Econs a bl) e2 m2 (v :: vl).
+ (a : expr) (bl : exprlist) (t t1: trace) (e1 : env) (m1 : mem)
+ (v : val) (t2: trace) (e2 : env) (m2 : mem) (vl : list val),
+ eval_expr ge sp le e m a t1 e1 m1 v ->
+ transl_expr_correct sp le e m a t1 e1 m1 v ->
+ eval_exprlist ge sp le e1 m1 bl t2 e2 m2 vl ->
+ transl_exprlist_correct sp le e1 m1 bl t2 e2 m2 vl ->
+ t = t1 ** t2 ->
+ transl_exprlist_correct sp le e m (Econs a bl) t e2 m2 (v :: vl).
Proof.
intros. red. intros.
inversion TRG.
- rewrite <- H10 in TE. simpl in TE. monadInv TE. intro EQ1.
+ rewrite <- H11 in TE. simpl in TE. monadInv TE. intro EQ1.
simpl in MUT.
assert (MUT1: incl (mutated_expr a) mut); eauto with coqlib.
assert (MUT2: incl (mutated_exprlist bl) mut); eauto with coqlib.
@@ -875,12 +933,13 @@ Proof.
assert (TRG1: target_reg_ok s1 map mut a r); eauto with rtlg.
generalize (H0 _ _ _ _ _ _ _ _ MWF1 EQ1 ME MUT1 TRG1).
intros [rs1 [EX1 [ME1 [RES1 OTHER1]]]].
- generalize (H2 _ _ _ _ _ _ _ _ MWF EQ ME1 MUT2 H11).
+ generalize (H2 _ _ _ _ _ _ _ _ MWF EQ ME1 MUT2 H12).
intros [rs2 [EX2 [ME2 [RES2 OTHER2]]]].
exists rs2.
(* Exec *)
split. eapply exec_trans. eexact EX1.
- apply exec_instrs_incr with s1. eauto with rtlg. assumption.
+ apply exec_instrs_incr with s1. eauto with rtlg.
+ eexact EX2. auto.
(* Match-env *)
split. assumption.
(* Results *)
@@ -894,22 +953,23 @@ Proof.
transitivity (rs1#r0).
apply OTHER2; auto. tauto.
apply OTHER1; auto. eauto with rtlg.
- elim H14; intro. left; assumption. right; apply sym_not_equal; tauto.
+ elim H15; intro. left; assumption. right; apply sym_not_equal; tauto.
Qed.
-Lemma transl_funcall_correct:
- forall (m : mem) (f : Cminor.function) (vargs : list val)
- (m1 : mem) (sp : block) (e e2 : env) (m2 : mem)
- (out : outcome) (vres : val),
+Lemma transl_funcall_internal_correct:
+ forall (m : mem) (f : Cminor.function)
+ (vargs : list val) (m1 : mem) (sp : block) (e : env) (t : trace)
+ (e2 : env) (m2 : mem) (out : outcome) (vres : val),
Mem.alloc m 0 (fn_stackspace f) = (m1, sp) ->
- set_locals (Cminor.fn_vars f) (set_params vargs (Cminor.fn_params f)) = e ->
- exec_stmt ge (Vptr sp Int.zero) e m1 (fn_body f) e2 m2 out ->
- transl_stmt_correct (Vptr sp Int.zero) e m1 (fn_body f) e2 m2 out ->
+ set_locals (fn_vars f) (set_params vargs (Cminor.fn_params f)) = e ->
+ exec_stmt ge (Vptr sp Int.zero) e m1 (fn_body f) t e2 m2 out ->
+ transl_stmt_correct (Vptr sp Int.zero) e m1 (fn_body f) t e2 m2 out ->
outcome_result_value out f.(Cminor.fn_sig).(sig_res) vres ->
- transl_function_correct m f vargs (Mem.free m2 sp) vres.
+ transl_function_correct m (Internal f)
+ vargs t (Mem.free m2 sp) vres.
Proof.
intros; red; intros.
- generalize TE. unfold transl_function.
+ generalize TE. unfold transl_fundef, transl_function; simpl.
caseEq (transl_fun f init_state).
intros; discriminate.
intros ns s. unfold transl_fun. monadSimpl.
@@ -934,30 +994,41 @@ Proof.
apply map_wf_incr with s1. apply state_incr_trans with s2; eauto with rtlg.
assumption.
- assert (RRG: return_reg_ok s3 y0 (ret_reg (Cminor.fn_sig f) r)).
+ set (rr := ret_reg (Cminor.fn_sig f) r) in *.
+ assert (RRG: return_reg_ok s3 y0 rr).
apply return_reg_ok_incr with s2. eauto with rtlg.
- apply new_reg_return_ok with s1; assumption.
+ unfold rr; apply new_reg_return_ok with s1; assumption.
generalize (H2 _ _ _ _ _ _ _ _ _ rs MWF EQ3 ME OUT RRG).
intros [rs1 [EX1 [ME1 MR1]]].
- apply exec_funct with m1 n rs1 (ret_reg (Cminor.fn_sig f) r).
- rewrite <- TF; simpl; assumption.
- rewrite <- TF; simpl. exact EX1.
- rewrite <- TF; simpl. apply instr_at_incr with s3.
+ rewrite <- TF. apply exec_funct_internal with m1 n rs1 rr; simpl.
+ assumption.
+ exact EX1.
+ apply instr_at_incr with s3.
eauto with rtlg. discriminate. eauto with rtlg.
generalize MR1. unfold match_return_outcome.
generalize H3. unfold outcome_result_value.
- unfold ret_reg; destruct (sig_res (Cminor.fn_sig f)).
+ unfold rr, ret_reg; destruct (sig_res (Cminor.fn_sig f)).
unfold regmap_optget. destruct out; try contradiction.
destruct o; try contradiction. intros; congruence.
unfold regmap_optget. destruct out; contradiction||auto.
destruct o; contradiction||auto.
Qed.
+Lemma transl_funcall_external_correct:
+ forall (ef : external_function) (m : mem) (args : list val) (t: trace) (res : val),
+ event_match ef args t res ->
+ transl_function_correct m (External ef) args t m res.
+Proof.
+ intros; red; intros. unfold transl_function in TE; simpl in TE.
+ inversion TE; subst tf.
+ apply exec_funct_external. auto.
+Qed.
+
Lemma transl_stmt_Sskip_correct:
forall (sp: val) (e : env) (m : mem),
- transl_stmt_correct sp e m Sskip e m Out_normal.
+ transl_stmt_correct sp e m Sskip E0 e m Out_normal.
Proof.
intros; red; intros. simpl in TE. monadInv TE.
subst s'; subst ns.
@@ -966,13 +1037,15 @@ Proof.
Qed.
Lemma transl_stmt_Sseq_continue_correct:
- forall (sp : val) (e : env) (m : mem) (s1 : stmt) (e1 : env)
- (m1 : mem) (s2 : stmt) (e2 : env) (m2 : mem) (out : outcome),
- exec_stmt ge sp e m s1 e1 m1 Out_normal ->
- transl_stmt_correct sp e m s1 e1 m1 Out_normal ->
- exec_stmt ge sp e1 m1 s2 e2 m2 out ->
- transl_stmt_correct sp e1 m1 s2 e2 m2 out ->
- transl_stmt_correct sp e m (Sseq s1 s2) e2 m2 out.
+ forall (sp : val) (e : env) (m : mem) (t: trace) (s1 : stmt)
+ (t1: trace) (e1 : env) (m1 : mem) (s2 : stmt) (t2: trace)
+ (e2 : env) (m2 : mem) (out : outcome),
+ exec_stmt ge sp e m s1 t1 e1 m1 Out_normal ->
+ transl_stmt_correct sp e m s1 t1 e1 m1 Out_normal ->
+ exec_stmt ge sp e1 m1 s2 t2 e2 m2 out ->
+ transl_stmt_correct sp e1 m1 s2 t2 e2 m2 out ->
+ t = t1 ** t2 ->
+ transl_stmt_correct sp e m (Sseq s1 s2) t e2 m2 out.
Proof.
intros; red; intros. simpl in TE; monadInv TE. intro EQ0.
assert (MWF1: map_wf map s0). eauto with rtlg.
@@ -987,18 +1060,18 @@ Proof.
(* Exec *)
split. eapply exec_trans. eexact EX1.
apply exec_instrs_incr with s0. eauto with rtlg.
- exact EX2.
+ eexact EX2. auto.
(* Match-env + return *)
tauto.
Qed.
Lemma transl_stmt_Sseq_stop_correct:
- forall (sp : val) (e : env) (m : mem) (s1 s2 : stmt) (e1 : env)
+ forall (sp : val) (e : env) (m : mem) (t: trace) (s1 s2 : stmt) (e1 : env)
(m1 : mem) (out : outcome),
- exec_stmt ge sp e m s1 e1 m1 out ->
- transl_stmt_correct sp e m s1 e1 m1 out ->
+ exec_stmt ge sp e m s1 t e1 m1 out ->
+ transl_stmt_correct sp e m s1 t e1 m1 out ->
out <> Out_normal ->
- transl_stmt_correct sp e m (Sseq s1 s2) e1 m1 out.
+ transl_stmt_correct sp e m (Sseq s1 s2) t e1 m1 out.
Proof.
intros; red; intros.
simpl in TE; monadInv TE. intro EQ0; clear TE.
@@ -1010,11 +1083,11 @@ Proof.
Qed.
Lemma transl_stmt_Sexpr_correct:
- forall (sp: val) (e : env) (m : mem) (a : expr)
+ forall (sp: val) (e : env) (m : mem) (a : expr) (t: trace)
(e1 : env) (m1 : mem) (v : val),
- eval_expr ge sp nil e m a e1 m1 v ->
- transl_expr_correct sp nil e m a e1 m1 v ->
- transl_stmt_correct sp e m (Sexpr a) e1 m1 Out_normal.
+ eval_expr ge sp nil e m a t e1 m1 v ->
+ transl_expr_correct sp nil e m a t e1 m1 v ->
+ transl_stmt_correct sp e m (Sexpr a) t e1 m1 Out_normal.
Proof.
intros; red; intros.
simpl in OUT. subst nd.
@@ -1028,17 +1101,18 @@ Qed.
Lemma transl_stmt_Sifthenelse_correct:
forall (sp: val) (e : env) (m : mem) (a : condexpr)
- (sl1 sl2 : stmt) (e1 : env) (m1 : mem)
- (v1 : bool) (e2 : env) (m2 : mem) (out : outcome),
- eval_condexpr ge sp nil e m a e1 m1 v1 ->
- transl_condition_correct sp nil e m a e1 m1 v1 ->
- exec_stmt ge sp e1 m1 (if v1 then sl1 else sl2) e2 m2 out ->
- transl_stmt_correct sp e1 m1 (if v1 then sl1 else sl2) e2 m2 out ->
- transl_stmt_correct sp e m (Sifthenelse a sl1 sl2) e2 m2 out.
+ (s1 s2 : stmt) (t t1: trace) (e1 : env) (m1 : mem)
+ (v1 : bool) (t2: trace) (e2 : env) (m2 : mem) (out : outcome),
+ eval_condexpr ge sp nil e m a t1 e1 m1 v1 ->
+ transl_condition_correct sp nil e m a t1 e1 m1 v1 ->
+ exec_stmt ge sp e1 m1 (if v1 then s1 else s2) t2 e2 m2 out ->
+ transl_stmt_correct sp e1 m1 (if v1 then s1 else s2) t2 e2 m2 out ->
+ t = t1 ** t2 ->
+ transl_stmt_correct sp e m (Sifthenelse a s1 s2) t e2 m2 out.
Proof.
intros; red; intros until rs; intro MWF.
- simpl. case (more_likely a sl1 sl2); monadSimpl; intro EQ2; intros.
- assert (MWF1: map_wf map s1). eauto with rtlg.
+ simpl. case (more_likely a s1 s2); monadSimpl; intro EQ2; intros.
+ assert (MWF1: map_wf map s3). eauto with rtlg.
generalize (H0 _ _ _ _ _ _ _ rs MWF1 EQ2 ME (incl_refl _)).
intros [rs1 [EX1 [ME1 OTHER1]]].
destruct v1.
@@ -1047,17 +1121,17 @@ Proof.
generalize (H2 _ _ _ _ _ _ _ _ _ _ MWF0 EQ0 ME1 OUT RRG0).
intros [rs2 [EX2 [ME2 MRE2]]].
exists rs2. split.
- eapply exec_trans. eexact EX1. apply exec_instrs_incr with s1.
- eauto with rtlg. exact EX2.
+ eapply exec_trans. eexact EX1. apply exec_instrs_incr with s3.
+ eauto with rtlg. eexact EX2. auto.
tauto.
generalize (H2 _ _ _ _ _ _ _ _ _ _ MWF EQ ME1 OUT RRG).
intros [rs2 [EX2 [ME2 MRE2]]].
exists rs2. split.
eapply exec_trans. eexact EX1. apply exec_instrs_incr with s0.
- eauto with rtlg. exact EX2.
+ eauto with rtlg. eexact EX2. auto.
tauto.
- assert (MWF1: map_wf map s1). eauto with rtlg.
+ assert (MWF1: map_wf map s3). eauto with rtlg.
generalize (H0 _ _ _ _ _ _ _ rs MWF1 EQ2 ME (incl_refl _)).
intros [rs1 [EX1 [ME1 OTHER1]]].
destruct v1.
@@ -1065,27 +1139,28 @@ Proof.
intros [rs2 [EX2 [ME2 MRE2]]].
exists rs2. split.
eapply exec_trans. eexact EX1. apply exec_instrs_incr with s0.
- eauto with rtlg. exact EX2.
+ eauto with rtlg. eexact EX2. auto.
tauto.
assert (MWF0: map_wf map s0). eauto with rtlg.
assert (RRG0: return_reg_ok s0 map rret). eauto with rtlg.
generalize (H2 _ _ _ _ _ _ _ _ _ _ MWF0 EQ0 ME1 OUT RRG0).
intros [rs2 [EX2 [ME2 MRE2]]].
exists rs2. split.
- eapply exec_trans. eexact EX1. apply exec_instrs_incr with s1.
- eauto with rtlg. exact EX2.
+ eapply exec_trans. eexact EX1. apply exec_instrs_incr with s3.
+ eauto with rtlg. eexact EX2. auto.
tauto.
Qed.
Lemma transl_stmt_Sloop_loop_correct:
- forall (sp: val) (e : env) (m : mem) (sl : stmt)
- (e1 : env) (m1 : mem) (e2 : env) (m2 : mem)
+ forall (sp: val) (e : env) (m : mem) (sl : stmt) (t t1: trace)
+ (e1 : env) (m1 : mem) (t2: trace) (e2 : env) (m2 : mem)
(out : outcome),
- exec_stmt ge sp e m sl e1 m1 Out_normal ->
- transl_stmt_correct sp e m sl e1 m1 Out_normal ->
- exec_stmt ge sp e1 m1 (Sloop sl) e2 m2 out ->
- transl_stmt_correct sp e1 m1 (Sloop sl) e2 m2 out ->
- transl_stmt_correct sp e m (Sloop sl) e2 m2 out.
+ exec_stmt ge sp e m sl t1 e1 m1 Out_normal ->
+ transl_stmt_correct sp e m sl t1 e1 m1 Out_normal ->
+ exec_stmt ge sp e1 m1 (Sloop sl) t2 e2 m2 out ->
+ transl_stmt_correct sp e1 m1 (Sloop sl) t2 e2 m2 out ->
+ t = t1 ** t2 ->
+ transl_stmt_correct sp e m (Sloop sl) t e2 m2 out.
Proof.
intros; red; intros.
generalize TE. simpl. monadSimpl. subst s2; subst n0. intros.
@@ -1107,22 +1182,23 @@ Proof.
apply exec_instrs_extends with s1.
eapply update_instr_extends.
eexact EQ. eauto with rtlg. eexact EQ1. eexact EX1.
- apply exec_trans with ns rs1 m1.
+ apply exec_trans with E0 ns rs1 m1 t2.
apply exec_one. apply exec_Inop.
generalize EQ1. unfold update_instr.
case (plt n (st_nextnode s1)); intro; monadSimpl.
- rewrite <- H4. simpl. apply PTree.gss.
+ rewrite <- H5. simpl. apply PTree.gss.
exact EX2.
+ reflexivity. traceEq.
tauto.
Qed.
Lemma transl_stmt_Sloop_stop_correct:
- forall (sp: val) (e : env) (m : mem) (sl : stmt)
+ forall (sp: val) (e : env) (m : mem) (t: trace) (sl : stmt)
(e1 : env) (m1 : mem) (out : outcome),
- exec_stmt ge sp e m sl e1 m1 out ->
- transl_stmt_correct sp e m sl e1 m1 out ->
+ exec_stmt ge sp e m sl t e1 m1 out ->
+ transl_stmt_correct sp e m sl t e1 m1 out ->
out <> Out_normal ->
- transl_stmt_correct sp e m (Sloop sl) e1 m1 out.
+ transl_stmt_correct sp e m (Sloop sl) t e1 m1 out.
Proof.
intros; red; intros.
simpl in TE. monadInv TE. subst s2; subst n0.
@@ -1145,11 +1221,11 @@ Proof.
Qed.
Lemma transl_stmt_Sblock_correct:
- forall (sp: val) (e : env) (m : mem) (sl : stmt)
+ forall (sp: val) (e : env) (m : mem) (sl : stmt) (t: trace)
(e1 : env) (m1 : mem) (out : outcome),
- exec_stmt ge sp e m sl e1 m1 out ->
- transl_stmt_correct sp e m sl e1 m1 out ->
- transl_stmt_correct sp e m (Sblock sl) e1 m1 (outcome_block out).
+ exec_stmt ge sp e m sl t e1 m1 out ->
+ transl_stmt_correct sp e m sl t e1 m1 out ->
+ transl_stmt_correct sp e m (Sblock sl) t e1 m1 (outcome_block out).
Proof.
intros; red; intros. simpl in TE.
assert (OUT': outcome_node out ncont (ncont :: nexits) nret nd).
@@ -1180,7 +1256,7 @@ Qed.
Lemma transl_stmt_Sexit_correct:
forall (sp: val) (e : env) (m : mem) (n : nat),
- transl_stmt_correct sp e m (Sexit n) e m (Out_exit n).
+ transl_stmt_correct sp e m (Sexit n) E0 e m (Out_exit n).
Proof.
intros; red; intros.
simpl in TE. simpl in OUT.
@@ -1194,7 +1270,7 @@ Lemma transl_switch_correct:
transl_switch r nexits cases default s = OK ns s' ->
nth_error nexits (switch_target i default cases) = Some nd ->
rs#r = Vint i ->
- exec_instrs tge s'.(st_code) sp ns rs m nd rs m.
+ exec_instrs tge s'.(st_code) sp ns rs m E0 nd rs m.
Proof.
induction cases; simpl; intros.
generalize (transl_exit_correct _ _ _ _ _ H). intros.
@@ -1202,23 +1278,23 @@ Proof.
destruct a as [key1 exit1]. monadInv H. clear H. intro EQ1.
caseEq (Int.eq i key1); intro IEQ; rewrite IEQ in H0.
(* i = key1 *)
- apply exec_trans with n0 rs m. apply exec_one.
+ apply exec_trans with E0 n0 rs m E0. apply exec_one.
eapply exec_Icond_true. eauto with rtlg. simpl. rewrite H1. congruence.
generalize (transl_exit_correct _ _ _ _ _ EQ0); intro.
- assert (n0 = nd). congruence. subst n0. apply exec_refl.
+ assert (n0 = nd). congruence. subst n0. apply exec_refl. traceEq.
(* i <> key1 *)
- apply exec_trans with n rs m. apply exec_one.
+ apply exec_trans with E0 n rs m E0. apply exec_one.
eapply exec_Icond_false. eauto with rtlg. simpl. rewrite H1. congruence.
- apply exec_instrs_incr with s0; eauto with rtlg.
+ apply exec_instrs_incr with s0; eauto with rtlg. traceEq.
Qed.
Lemma transl_stmt_Sswitch_correct:
forall (sp : val) (e : env) (m : mem) (a : expr)
- (cases : list (int * nat)) (default : nat) (e1 : env) (m1 : mem)
- (n : int),
- eval_expr ge sp nil e m a e1 m1 (Vint n) ->
- transl_expr_correct sp nil e m a e1 m1 (Vint n) ->
- transl_stmt_correct sp e m (Sswitch a cases default) e1 m1
+ (cases : list (int * nat)) (default : nat)
+ (t1 : trace) (e1 : env) (m1 : mem) (n : int),
+ eval_expr ge sp nil e m a t1 e1 m1 (Vint n) ->
+ transl_expr_correct sp nil e m a t1 e1 m1 (Vint n) ->
+ transl_stmt_correct sp e m (Sswitch a cases default) t1 e1 m1
(Out_exit (switch_target n default cases)).
Proof.
intros; red; intros. monadInv TE. clear TE; intros EQ1.
@@ -1234,14 +1310,14 @@ Proof.
(* execution *)
split. eapply exec_trans. eexact EXEC1.
apply exec_instrs_incr with s1. eauto with rtlg.
- eapply transl_switch_correct; eauto.
+ eapply transl_switch_correct; eauto. traceEq.
(* match_env & match_reg *)
tauto.
Qed.
Lemma transl_stmt_Sreturn_none_correct:
forall (sp: val) (e : env) (m : mem),
- transl_stmt_correct sp e m (Sreturn None) e m (Out_return None).
+ transl_stmt_correct sp e m (Sreturn None) E0 e m (Out_return None).
Proof.
intros; red; intros. generalize TE. simpl.
destruct rret; monadSimpl.
@@ -1250,11 +1326,11 @@ Proof.
Qed.
Lemma transl_stmt_Sreturn_some_correct:
- forall (sp: val) (e : env) (m : mem) (a : expr)
+ forall (sp: val) (e : env) (m : mem) (a : expr) (t: trace)
(e1 : env) (m1 : mem) (v : val),
- eval_expr ge sp nil e m a e1 m1 v ->
- transl_expr_correct sp nil e m a e1 m1 v ->
- transl_stmt_correct sp e m (Sreturn (Some a)) e1 m1 (Out_return (Some v)).
+ eval_expr ge sp nil e m a t e1 m1 v ->
+ transl_expr_correct sp nil e m a t e1 m1 v ->
+ transl_stmt_correct sp e m (Sreturn (Some a)) t e1 m1 (Out_return (Some v)).
Proof.
intros; red; intros. generalize TE; simpl.
destruct rret. intro EQ.
@@ -1278,9 +1354,9 @@ Scheme eval_expr_ind_5 := Minimality for eval_expr Sort Prop
with exec_stmt_ind_5 := Minimality for exec_stmt Sort Prop.
Theorem transl_function_correctness:
- forall m f vargs m' vres,
- eval_funcall ge m f vargs m' vres ->
- transl_function_correct m f vargs m' vres.
+ forall m f vargs t m' vres,
+ eval_funcall ge m f vargs t m' vres ->
+ transl_function_correct m f vargs t m' vres.
Proof
(eval_funcall_ind_5 ge
transl_expr_correct
@@ -1298,13 +1374,15 @@ Proof
transl_expr_Econdition_correct
transl_expr_Elet_correct
transl_expr_Eletvar_correct
+ transl_expr_Ealloc_correct
transl_condition_CEtrue_correct
transl_condition_CEfalse_correct
transl_condition_CEcond_correct
transl_condition_CEcondition_correct
transl_exprlist_Enil_correct
transl_exprlist_Econs_correct
- transl_funcall_correct
+ transl_funcall_internal_correct
+ transl_funcall_external_correct
transl_stmt_Sskip_correct
transl_stmt_Sexpr_correct
transl_stmt_Sifthenelse_correct
@@ -1319,26 +1397,23 @@ Proof
transl_stmt_Sreturn_some_correct).
Theorem transl_program_correct:
- forall (r: val),
- Cminor.exec_program prog r ->
- RTL.exec_program tprog r.
+ forall (t: trace) (r: val),
+ Cminor.exec_program prog t r ->
+ RTL.exec_program tprog t r.
Proof.
- intros r [b [f [m [SYMB [FUNC [SIG EVAL]]]]]].
+ intros t r [b [f [m [SYMB [FUNC [SIG EVAL]]]]]].
generalize (function_ptr_translated _ _ FUNC).
intros [tf [TFIND TRANSLF]].
red; exists b; exists tf; exists m.
split. rewrite symbols_preserved.
replace (prog_main tprog) with (prog_main prog).
assumption.
- symmetry; apply transform_partial_program_main with transl_function.
+ symmetry; apply transform_partial_program_main with transl_fundef.
exact TRANSL.
split. exact TFIND.
- split. generalize TRANSLF. unfold transl_function.
- destruct (transl_fun f init_state).
- intro; discriminate. destruct p; intro EQ; injection EQ; intro EQ1.
- rewrite <- EQ1. simpl. congruence.
- rewrite (Genv.init_mem_transf_partial transl_function prog TRANSL).
- exact (transl_function_correctness _ _ _ _ _ EVAL _ TRANSLF).
+ split. generalize (sig_transl_function _ _ TRANSLF). congruence.
+ unfold fundef; rewrite (Genv.init_mem_transf_partial transl_fundef prog TRANSL).
+ exact (transl_function_correctness _ _ _ _ _ _ EVAL _ TRANSLF).
Qed.
End CORRECTNESS.
diff --git a/backend/RTLgenproof1.v b/backend/RTLgenproof1.v
index 85d420e0..8b149015 100644
--- a/backend/RTLgenproof1.v
+++ b/backend/RTLgenproof1.v
@@ -5,6 +5,7 @@ Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Values.
+Require Import Events.
Require Import Mem.
Require Import Globalenvs.
Require Import Op.
@@ -200,28 +201,28 @@ Variable s1 s2: state.
Hypothesis EXT: state_extends s1 s2.
Lemma exec_instr_not_halt:
- forall ge c sp pc rs m pc' rs' m',
- exec_instr ge c sp pc rs m pc' rs' m' -> c!pc <> None.
+ forall ge c sp pc rs m t pc' rs' m',
+ exec_instr ge c sp pc rs m t pc' rs' m' -> c!pc <> None.
Proof.
induction 1; rewrite H; discriminate.
Qed.
Lemma exec_instr_in_s2:
- forall ge sp pc rs m pc' rs' m',
- exec_instr ge s1.(st_code) sp pc rs m pc' rs' m' ->
+ forall ge sp pc rs m t pc' rs' m',
+ exec_instr ge s1.(st_code) sp pc rs m t pc' rs' m' ->
s2.(st_code)!pc = s1.(st_code)!pc.
Proof.
intros.
elim (EXT pc); intro.
- elim (exec_instr_not_halt _ _ _ _ _ _ _ _ _ H H0).
+ elim (exec_instr_not_halt _ _ _ _ _ _ _ _ _ _ H H0).
assumption.
Qed.
Lemma exec_instr_extends_rec:
- forall ge c sp pc rs m pc' rs' m',
- exec_instr ge c sp pc rs m pc' rs' m' ->
+ forall ge c sp pc rs m t pc' rs' m',
+ exec_instr ge c sp pc rs m t pc' rs' m' ->
forall c', c!pc = c'!pc ->
- exec_instr ge c' sp pc rs m pc' rs' m'.
+ exec_instr ge c' sp pc rs m t pc' rs' m'.
Proof.
induction 1; intros.
apply exec_Inop. congruence.
@@ -230,14 +231,15 @@ Proof.
apply exec_Istore with chunk addr args src a.
congruence. auto. auto.
apply exec_Icall with sig ros args f; auto. congruence.
+ apply exec_Ialloc with arg sz; auto. congruence.
apply exec_Icond_true with cond args ifnot; auto. congruence.
apply exec_Icond_false with cond args ifso; auto. congruence.
Qed.
Lemma exec_instr_extends:
- forall ge sp pc rs m pc' rs' m',
- exec_instr ge s1.(st_code) sp pc rs m pc' rs' m' ->
- exec_instr ge s2.(st_code) sp pc rs m pc' rs' m'.
+ forall ge sp pc rs m t pc' rs' m',
+ exec_instr ge s1.(st_code) sp pc rs m t pc' rs' m' ->
+ exec_instr ge s2.(st_code) sp pc rs m t pc' rs' m'.
Proof.
intros.
apply exec_instr_extends_rec with (st_code s1).
@@ -246,21 +248,21 @@ Proof.
Qed.
Lemma exec_instrs_extends_rec:
- forall ge c sp pc rs m pc' rs' m',
- exec_instrs ge c sp pc rs m pc' rs' m' ->
+ forall ge c sp pc rs m t pc' rs' m',
+ exec_instrs ge c sp pc rs m t pc' rs' m' ->
c = s1.(st_code) ->
- exec_instrs ge s2.(st_code) sp pc rs m pc' rs' m'.
+ exec_instrs ge s2.(st_code) sp pc rs m t pc' rs' m'.
Proof.
induction 1; intros.
apply exec_refl.
apply exec_one. apply exec_instr_extends; auto. rewrite <- H0; auto.
- apply exec_trans with pc2 rs2 m2; auto.
+ apply exec_trans with t1 pc2 rs2 m2 t2; auto.
Qed.
Lemma exec_instrs_extends:
- forall ge sp pc rs m pc' rs' m',
- exec_instrs ge s1.(st_code) sp pc rs m pc' rs' m' ->
- exec_instrs ge s2.(st_code) sp pc rs m pc' rs' m'.
+ forall ge sp pc rs m t pc' rs' m',
+ exec_instrs ge s1.(st_code) sp pc rs m t pc' rs' m' ->
+ exec_instrs ge s2.(st_code) sp pc rs m t pc' rs' m'.
Proof.
intros.
apply exec_instrs_extends_rec with (st_code s1); auto.
@@ -281,9 +283,9 @@ Variable s1 s2: state.
Hypothesis INCR: state_incr s1 s2.
Lemma exec_instr_incr:
- forall ge sp pc rs m pc' rs' m',
- exec_instr ge s1.(st_code) sp pc rs m pc' rs' m' ->
- exec_instr ge s2.(st_code) sp pc rs m pc' rs' m'.
+ forall ge sp pc rs m t pc' rs' m',
+ exec_instr ge s1.(st_code) sp pc rs m t pc' rs' m' ->
+ exec_instr ge s2.(st_code) sp pc rs m t pc' rs' m'.
Proof.
intros.
apply exec_instr_extends with s1.
@@ -292,9 +294,9 @@ Proof.
Qed.
Lemma exec_instrs_incr:
- forall ge sp pc rs m pc' rs' m',
- exec_instrs ge s1.(st_code) sp pc rs m pc' rs' m' ->
- exec_instrs ge s2.(st_code) sp pc rs m pc' rs' m'.
+ forall ge sp pc rs m t pc' rs' m',
+ exec_instrs ge s1.(st_code) sp pc rs m t pc' rs' m' ->
+ exec_instrs ge s2.(st_code) sp pc rs m t pc' rs' m'.
Proof.
intros.
apply exec_instrs_extends with s1.
@@ -1318,6 +1320,7 @@ forall (P : expr -> Prop) (P0 : condexpr -> Prop)
P e -> forall e0 : expr, P e0 -> P (Econdition c e e0)) ->
(forall e : expr, P e -> forall e0 : expr, P e0 -> P (Elet e e0)) ->
(forall n : nat, P (Eletvar n)) ->
+ (forall e : expr, P e -> P (Ealloc e)) ->
P0 CEtrue ->
P0 CEfalse ->
(forall (c : condition) (e : exprlist), P1 e -> P0 (CEcond c e)) ->
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index d15dbb88..33338d37 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -6,7 +6,7 @@ Require Import AST.
Require Import Op.
Require Import Registers.
Require Import RTL.
-Require Import union_find.
+Require Conventions.
(** * The type system *)
@@ -38,7 +38,7 @@ Definition regenv := reg -> typ.
Section WT_INSTR.
Variable env: regenv.
-Variable funct: function.
+Variable funsig: signature.
Inductive wt_instr : instruction -> Prop :=
| wt_Inop:
@@ -72,13 +72,17 @@ Inductive wt_instr : instruction -> Prop :=
List.map env args = sig.(sig_args) ->
env res = match sig.(sig_res) with None => Tint | Some ty => ty end ->
wt_instr (Icall sig ros args res s)
+ | wt_Ialloc:
+ forall arg res s,
+ env arg = Tint -> env res = Tint ->
+ wt_instr (Ialloc arg res s)
| wt_Icond:
forall cond args s1 s2,
List.map env args = type_of_condition cond ->
wt_instr (Icond cond args s1 s2)
| wt_Ireturn:
forall optres,
- option_map env optres = funct.(fn_sig).(sig_res) ->
+ option_map env optres = funsig.(sig_res) ->
wt_instr (Ireturn optres).
End WT_INSTR.
@@ -88,15 +92,27 @@ End WT_INSTR.
parameters agree in types with the function signature, and
names of parameters are pairwise distinct. *)
-Record wt_function (env: regenv) (f: function) : Prop :=
+Record wt_function (f: function) (env: regenv): Prop :=
mk_wt_function {
wt_params:
List.map env f.(fn_params) = f.(fn_sig).(sig_args);
wt_norepet:
list_norepet f.(fn_params);
wt_instrs:
- forall pc instr, f.(fn_code)!pc = Some instr -> wt_instr env f instr
- }.
+ forall pc instr,
+ f.(fn_code)!pc = Some instr -> wt_instr env f.(fn_sig) instr
+}.
+
+Inductive wt_fundef: fundef -> Prop :=
+ | wt_fundef_external: forall ef,
+ Conventions.sig_external_ok ef.(ef_sig) ->
+ wt_fundef (External ef)
+ | wt_function_internal: forall f env,
+ wt_function f env ->
+ wt_fundef (Internal f).
+
+Definition wt_program (p: program): Prop :=
+ forall i f, In (i, f) (prog_funct p) -> wt_fundef f.
(** * Type inference *)
@@ -109,1022 +125,221 @@ Record wt_function (env: regenv) (f: function) : Prop :=
each pseudo-register from its uses in the code. We follow the second
approach.
- The algorithm and its correctness proof in this section were
- contributed by Damien Doligez. *)
+ We delegate the task of determining the type of each pseudo-register
+ to an external ``oracle'': a function written in Caml and not
+ proved correct. We verify the returned type environment using
+ the following Coq code, which we will prove correct. *)
-(** ** Type inference algorithm *)
+Parameter infer_type_environment:
+ function -> list (node * instruction) -> option regenv.
-Set Implicit Arguments.
+(** ** Algorithm to check the correctness of a type environment *)
-(** Type inference for RTL is similar to that for simply-typed
- lambda-calculus: we use type variables to represent the types
- of registers that have not yet been determined to be [Tint] or [Tfloat]
- based on their uses. We need exactly one type variable per pseudo-register,
- therefore type variables can be conveniently equated with registers.
- The type of a register during inference is therefore either
- [tTy t] (with [t = Tint] or [t = Tfloat]) for a known type,
- or [tReg r] to mean ``the same type as that of register [r]''. *)
+Section TYPECHECKING.
-Inductive myT : Set :=
- | tTy : typ -> myT
- | tReg : reg -> myT.
+Variable funct: function.
+Variable env: regenv.
-(** The algorithm proceeds by unification of the currently inferred
- type for a pseudo-register with the type dictated by its uses.
- Unification builds on a ``union-find'' data structure representing
- equivalence classes of types (see module [Union_find]).
-*)
+Lemma typ_eq: forall (t1 t2: typ), {t1=t2} + {t1<>t2}.
+Proof. decide equality. Qed.
-Module myreg.
- Definition T := myT.
- Definition eq : forall (x y : T), {x=y}+{x<>y}.
- Proof.
- destruct x; destruct y; auto;
- try (apply right; discriminate); try (apply left; discriminate).
- destruct t; destruct t0;
- try (apply right; congruence); try (apply left; congruence).
- elim (peq r r0); intros.
- rewrite a; apply left; apply refl_equal.
- apply right; congruence.
- Defined.
-End myreg.
-
-Module mymap.
- Definition elt := myreg.T.
- Definition encode (t : myreg.T) : positive :=
- match t with
- | tTy Tint => xH
- | tTy Tfloat => xI xH
- | tReg r => xO r
- end.
- Definition decode (p : positive) : elt :=
- match p with
- | xH => tTy Tint
- | xI _ => tTy Tfloat
- | xO r => tReg r
- end.
-
- Lemma encode_decode : forall x : myreg.T, decode (encode x) = x.
- Proof.
- destruct x; try destruct t; simpl; auto.
- Qed.
-
- Lemma encode_injective :
- forall (x y : myreg.T), encode x = encode y -> x = y.
- Proof.
- intros.
- unfold encode in H. destruct x; destruct y; try congruence;
- try destruct t; try destruct t0; congruence.
- Qed.
-
- Definition T := PTree.t positive.
- Definition empty := PTree.empty positive.
- Definition get (adr : elt) (t : T) :=
- option_map decode (PTree.get (encode adr) t).
- Definition add (adr dat : elt) (t : T) :=
- PTree.set (encode adr) (encode dat) t.
-
- Theorem get_empty : forall (x : elt), get x empty = None.
- Proof.
- intro.
- unfold get. unfold empty.
- rewrite PTree.gempty.
- simpl; auto.
- Qed.
- Theorem get_add_1 :
- forall (x y : elt) (m : T), get x (add x y m) = Some y.
- Proof.
- intros.
- unfold add. unfold get.
- rewrite PTree.gss.
- simpl; rewrite encode_decode; auto.
- Qed.
- Theorem get_add_2 :
- forall (x y z : elt) (m : T), z <> x -> get z (add x y m) = get z m.
- Proof.
- intros.
- unfold get. unfold add.
- rewrite PTree.gso; auto.
- intro. apply H. apply encode_injective. auto.
- Qed.
-End mymap.
-
-Module Uf := Unionfind (myreg) (mymap).
-
-Definition error := Uf.identify Uf.empty (tTy Tint) (tTy Tfloat).
-
-Fixpoint fold2 (A B : Set) (f : Uf.T -> A -> B -> Uf.T)
- (init : Uf.T) (la : list A) (lb : list B) {struct la}
- : Uf.T :=
- match la, lb with
- | ha::ta, hb::tb => fold2 f (f init ha hb) ta tb
- | nil, nil => init
- | _, _ => error
- end.
+Definition check_reg (r: reg) (ty: typ): bool :=
+ if typ_eq (env r) ty then true else false.
-Definition option_fold2 (A B : Set) (f : Uf.T -> A -> B -> Uf.T)
- (init : Uf.T) (oa : option A) (ob : option B)
- : Uf.T :=
- match oa, ob with
- | Some a, Some b => f init a b
- | None, None => init
- | _, _ => error
- end.
-
-Definition teq (ty1 ty2 : typ) : bool :=
- match ty1, ty2 with
- | Tint, Tint => true
- | Tfloat, Tfloat => true
+Fixpoint check_regs (rl: list reg) (tyl: list typ) {struct rl}: bool :=
+ match rl, tyl with
+ | nil, nil => true
+ | r1::rs, ty::tys => check_reg r1 ty && check_regs rs tys
| _, _ => false
end.
-Definition type_rtl_arg (u : Uf.T) (r : reg) (t : typ) :=
- Uf.identify u (tReg r) (tTy t).
-
-Definition type_rtl_ros (u : Uf.T) (ros : reg+ident) :=
- match ros with
- | inl r => Uf.identify u (tReg r) (tTy Tint)
- | inr s => u
- end.
+Definition check_op (op: operation) (args: list reg) (res: reg): bool :=
+ let (targs, tres) := type_of_operation op in
+ check_regs args targs && check_reg res tres.
-Definition type_of_sig_res (sig : signature) :=
- match sig.(sig_res) with None => Tint | Some ty => ty end.
-
-(** This is the core type inference function. The [u] argument is
- the current substitution / equivalence classes between types.
- An updated set of equivalence classes, reflecting unifications
- possibly performed during the type-checking of [i], is returned.
- Note that [type_rtl_instr] never fails explicitly. However,
- in case of type error (e.g. applying the [Oadd] integer operation
- to float registers), the equivalence relation returned will
- put [tTy Tint] and [tTy Tfloat] in the same equivalence class.
- This fact will propagate through type inference for other instructions,
- and be detected at the end of type inference, indicating a typing failure.
-*)
-
-Definition type_rtl_instr (rtyp : option typ)
- (u : Uf.T) (_ : positive) (i : instruction) :=
+Definition check_instr (i: instruction) : bool :=
match i with
- | Inop _ => u
- | Iop Omove (r1 :: nil) r0 _ => Uf.identify u (tReg r0) (tReg r1)
- | Iop Omove _ _ _ => error
- | Iop Oundef nil _ _ => u
- | Iop Oundef _ _ _ => error
- | Iop op args r0 _ =>
- let (argtyp, restyp) := type_of_operation op in
- let u1 := type_rtl_arg u r0 restyp in
- fold2 type_rtl_arg u1 args argtyp
- | Iload chunk addr args r0 _ =>
- let u1 := type_rtl_arg u r0 (type_of_chunk chunk) in
- fold2 type_rtl_arg u1 args (type_of_addressing addr)
- | Istore chunk addr args r0 _ =>
- let u1 := type_rtl_arg u r0 (type_of_chunk chunk) in
- fold2 type_rtl_arg u1 args (type_of_addressing addr)
- | Icall sign ros args r0 _ =>
- let u1 := type_rtl_ros u ros in
- let u2 := type_rtl_arg u1 r0 (type_of_sig_res sign) in
- fold2 type_rtl_arg u2 args sign.(sig_args)
+ | Inop _ =>
+ true
+ | Iop Omove (arg::nil) res _ =>
+ if typ_eq (env arg) (env res) then true else false
+ | Iop Omove args res _ =>
+ false
+ | Iop Oundef nil res _ =>
+ true
+ | Iop Oundef args res _ =>
+ false
+ | Iop op args res _ =>
+ check_op op args res
+ | Iload chunk addr args dst _ =>
+ check_regs args (type_of_addressing addr)
+ && check_reg dst (type_of_chunk chunk)
+ | Istore chunk addr args src _ =>
+ check_regs args (type_of_addressing addr)
+ && check_reg src (type_of_chunk chunk)
+ | Icall sig ros args res _ =>
+ match ros with inl r => check_reg r Tint | inr s => true end
+ && check_regs args sig.(sig_args)
+ && check_reg res (match sig.(sig_res) with None => Tint | Some ty => ty end)
+ | Ialloc arg res _ =>
+ check_reg arg Tint && check_reg res Tint
| Icond cond args _ _ =>
- fold2 type_rtl_arg u args (type_of_condition cond)
- | Ireturn o => option_fold2 type_rtl_arg u o rtyp
+ check_regs args (type_of_condition cond)
+ | Ireturn optres =>
+ match optres, funct.(fn_sig).(sig_res) with
+ | None, None => true
+ | Some r, Some t => check_reg r t
+ | _, _ => false
+ end
end.
-Definition mk_env (u : Uf.T) (r : reg) :=
- if myreg.eq (Uf.repr u (tReg r)) (Uf.repr u (tTy Tfloat))
- then Tfloat
- else Tint.
+Definition check_params_norepet (params: list reg): bool :=
+ if list_norepet_dec Reg.eq params then true else false.
-Fixpoint member (x : reg) (l : list reg) {struct l} : bool :=
- match l with
- | nil => false
- | y :: rest => if peq x y then true else member x rest
+Fixpoint check_instrs (instrs: list (node * instruction)) : bool :=
+ match instrs with
+ | nil => true
+ | (pc, i) :: rem => check_instr i && check_instrs rem
end.
-Fixpoint repet (l : list reg) : bool :=
- match l with
- | nil => false
- | x :: rest => member x rest || repet rest
- end.
+(** ** Correctness of the type-checking algorithm *)
-Definition type_rtl_function (f : function) :=
- let u1 := PTree.fold (type_rtl_instr f.(fn_sig).(sig_res))
- f.(fn_code) Uf.empty in
- let u2 := fold2 type_rtl_arg u1 f.(fn_params) f.(fn_sig).(sig_args) in
- if repet f.(fn_params) then
- None
- else
- if myreg.eq (Uf.repr u2 (tTy Tint)) (Uf.repr u2 (tTy Tfloat))
- then None
- else Some (mk_env u2).
-
-Unset Implicit Arguments.
-
-(** ** Correctness proof for type inference *)
-
-(** General properties of the type equivalence relation. *)
-
-Definition consistent (u : Uf.T) :=
- Uf.repr u (tTy Tint) <> Uf.repr u (tTy Tfloat).
-
-Lemma consistent_not_eq : forall (u : Uf.T) (A : Type) (x y : A),
- consistent u ->
- (if myreg.eq (Uf.repr u (tTy Tint)) (Uf.repr u (tTy Tfloat)) then x else y)
- = y.
- Proof.
- intros.
- unfold consistent in H.
- destruct (myreg.eq (Uf.repr u (tTy Tint)) (Uf.repr u (tTy Tfloat)));
- congruence.
- Qed.
-
-Lemma equal_eq : forall (t : myT) (A : Type) (x y : A),
- (if myreg.eq t t then x else y) = x.
- Proof.
- intros.
- destruct (myreg.eq t t); congruence.
- Qed.
-
-Lemma error_inconsistent : forall (A : Prop), consistent error -> A.
- Proof.
- intros.
- absurd (consistent error); auto.
- intro.
- unfold error in H. unfold consistent in H.
- rewrite Uf.sameclass_identify_1 in H.
- congruence.
- Qed.
-
-Lemma teq_correct : forall (t1 t2 : typ), teq t1 t2 = true -> t1 = t2.
- Proof.
- intros; destruct t1; destruct t2; try simpl in H; congruence.
- Qed.
-
-Definition included (u1 u2 : Uf.T) : Prop :=
- forall (e1 e2: myT),
- Uf.repr u1 e1 = Uf.repr u1 e2 -> Uf.repr u2 e1 = Uf.repr u2 e2.
-
-Lemma included_refl :
- forall (e : Uf.T), included e e.
- Proof.
- unfold included. auto.
- Qed.
-
-Lemma included_trans :
- forall (e1 e2 e3 : Uf.T),
- included e3 e2 -> included e2 e1 -> included e3 e1.
- Proof.
- unfold included. auto.
- Qed.
-
-Lemma included_consistent :
- forall (u1 u2 : Uf.T),
- included u1 u2 -> consistent u2 -> consistent u1.
- Proof.
- unfold consistent. unfold included.
- intros.
- intro. apply H0. apply H.
- auto.
- Qed.
-
-Lemma included_identify :
- forall (u : Uf.T) (t1 t2 : myT), included u (Uf.identify u t1 t2).
- Proof.
- unfold included.
- intros.
- apply Uf.sameclass_identify_2; auto.
- Qed.
-
-Lemma type_arg_correct_1 :
- forall (u : Uf.T) (r : reg) (t : typ),
- consistent (type_rtl_arg u r t)
- -> Uf.repr (type_rtl_arg u r t) (tReg r)
- = Uf.repr (type_rtl_arg u r t) (tTy t).
- Proof.
- intros.
- unfold type_rtl_arg.
- rewrite Uf.sameclass_identify_1.
- auto.
- Qed.
-
-Lemma type_arg_correct :
- forall (u : Uf.T) (r : reg) (t : typ),
- consistent (type_rtl_arg u r t) -> mk_env (type_rtl_arg u r t) r = t.
- Proof.
- intros.
- unfold mk_env.
- rewrite type_arg_correct_1.
- destruct t.
- apply consistent_not_eq; auto.
- destruct (myreg.eq (Uf.repr (type_rtl_arg u r Tfloat) (tTy Tfloat)));
- congruence.
- auto.
- Qed.
-
-Lemma type_arg_included :
- forall (u : Uf.T) (r : reg) (t : typ), included u (type_rtl_arg u r t).
- Proof.
- intros.
- unfold type_rtl_arg.
- apply included_identify.
- Qed.
-
-Lemma type_arg_extends :
- forall (u : Uf.T) (r : reg) (t : typ),
- consistent (type_rtl_arg u r t) -> consistent u.
- Proof.
- intros.
- apply included_consistent with (u2 := type_rtl_arg u r t).
- apply type_arg_included.
- auto.
- Qed.
-
-Lemma type_args_included :
- forall (l1 : list reg) (l2 : list typ) (u : Uf.T),
- consistent (fold2 type_rtl_arg u l1 l2)
- -> included u (fold2 type_rtl_arg u l1 l2).
- Proof.
- induction l1; intros; destruct l2.
- simpl in H; simpl; apply included_refl.
- simpl in H. apply error_inconsistent. auto.
- simpl in H. apply error_inconsistent. auto.
- simpl.
- simpl in H.
- apply included_trans with (e2 := type_rtl_arg u a t).
- apply type_arg_included.
- apply IHl1.
- auto.
- Qed.
-
-Lemma type_args_extends :
- forall (l1 : list reg) (l2 : list typ) (u : Uf.T),
- consistent (fold2 type_rtl_arg u l1 l2) -> consistent u.
- Proof.
- intros.
- apply (included_consistent _ _ (type_args_included l1 l2 u H)).
- auto.
- Qed.
-
-Lemma type_args_correct :
- forall (l1 : list reg) (l2 : list typ) (u : Uf.T),
- consistent (fold2 type_rtl_arg u l1 l2)
- -> map (mk_env (fold2 type_rtl_arg u l1 l2)) l1 = l2.
- Proof.
- induction l1.
- intros.
- destruct l2.
- unfold map; simpl; auto.
- simpl in H; apply error_inconsistent; auto.
- intros.
- destruct l2.
- simpl in H; apply error_inconsistent; auto.
- simpl.
- simpl in H.
- rewrite (IHl1 l2 (type_rtl_arg u a t) H).
- unfold mk_env.
- destruct t.
- rewrite (type_args_included _ _ _ H (tReg a) (tTy Tint)).
- rewrite consistent_not_eq; auto.
- apply type_arg_correct_1.
- apply type_args_extends with (l1 := l1) (l2 := l2); auto.
- rewrite (type_args_included _ _ _ H (tReg a) (tTy Tfloat)).
- rewrite equal_eq; auto.
- apply type_arg_correct_1.
- apply type_args_extends with (l1 := l1) (l2 := l2); auto.
- Qed.
-
-(** Correctness of [wt_params]. *)
-
-Lemma type_rtl_function_params :
- forall (f: function) (env: regenv),
- type_rtl_function f = Some env
- -> List.map env f.(fn_params) = f.(fn_sig).(sig_args).
- Proof.
- destruct f; unfold type_rtl_function; simpl.
- destruct (repet fn_params); simpl; intros; try congruence.
- pose (u := PTree.fold (type_rtl_instr (sig_res fn_sig)) fn_code Uf.empty).
- fold u in H.
- cut (consistent (fold2 type_rtl_arg u fn_params (sig_args fn_sig))).
- intro.
- pose (R := Uf.repr (fold2 type_rtl_arg u fn_params (sig_args fn_sig))).
- fold R in H.
- destruct (myreg.eq (R (tTy Tint)) (R (tTy Tfloat))).
- congruence.
- injection H.
- intro.
- rewrite <- H1.
- apply type_args_correct.
- auto.
- intro.
- rewrite H0 in H.
- rewrite equal_eq in H.
- congruence.
- Qed.
-
-(** Correctness of [wt_norepet]. *)
-
-Lemma member_correct :
- forall (l : list reg) (a : reg), member a l = false -> ~In a l.
- Proof.
- induction l; simpl; intros; try tauto.
- destruct (peq a0 a); simpl; try congruence.
- intro. destruct H0; try congruence.
- generalize H0; apply IHl; auto.
- Qed.
-
-Lemma repet_correct :
- forall (l : list reg), repet l = false -> list_norepet l.
- Proof.
- induction l; simpl; intros.
- exact (list_norepet_nil reg).
- elim (orb_false_elim (member a l) (repet l) H); intros.
- apply list_norepet_cons.
- apply member_correct; auto.
- apply IHl; auto.
- Qed.
-
-Lemma type_rtl_function_norepet :
- forall (f: function) (env: regenv),
- type_rtl_function f = Some env
- -> list_norepet f.(fn_params).
- Proof.
- destruct f; unfold type_rtl_function; simpl.
- intros. cut (repet fn_params = false).
- intro. apply repet_correct. auto.
- destruct (repet fn_params); congruence.
- Qed.
-
-(** Correctness of [wt_instrs]. *)
-
-Lemma step1 :
- forall (f : function) (env : regenv),
- type_rtl_function f = Some env
- -> exists u2 : Uf.T,
- included (PTree.fold (type_rtl_instr f.(fn_sig).(sig_res))
- f.(fn_code) Uf.empty)
- u2
- /\ env = mk_env u2
- /\ consistent u2.
- Proof.
- intros f env.
- pose (u1 := (PTree.fold (type_rtl_instr f.(fn_sig).(sig_res))
- f.(fn_code) Uf.empty)).
- fold u1.
- unfold type_rtl_function.
- intros.
- destruct (repet f.(fn_params)).
- congruence.
- fold u1 in H.
- pose (u2 := (fold2 type_rtl_arg u1 f.(fn_params) f.(fn_sig).(sig_args))).
- fold u2 in H.
- exists u2.
- caseEq (myreg.eq (Uf.repr u2 (tTy Tint)) (Uf.repr u2 (tTy Tfloat))).
- intros.
- rewrite e in H.
- rewrite equal_eq in H.
- congruence.
- intros.
- rewrite consistent_not_eq in H.
- apply conj.
- unfold u2.
- apply type_args_included.
- auto.
- apply conj; auto.
- congruence.
- auto.
- Qed.
-
-Lemma let_fold_args_res :
- forall (u : Uf.T) (l : list reg) (r : reg) (e : list typ * typ),
- (let (argtyp, restyp) := e in
- fold2 type_rtl_arg (type_rtl_arg u r restyp) l argtyp)
- = fold2 type_rtl_arg (type_rtl_arg u r (snd e)) l (fst e).
- Proof.
- intros. rewrite (surjective_pairing e). simpl. auto.
- Qed.
-
-Lemma type_args_res_included :
- forall (l1 : list reg) (l2 : list typ) (u : Uf.T) (r : reg) (t : typ),
- consistent (fold2 type_rtl_arg (type_rtl_arg u r t) l1 l2)
- -> included u (fold2 type_rtl_arg (type_rtl_arg u r t) l1 l2).
- Proof.
- intros.
- apply included_trans with (e2 := type_rtl_arg u r t).
- apply type_arg_included.
- apply type_args_included; auto.
- Qed.
-
-Lemma type_args_res_ros_included :
- forall (l1 : list reg) (l2 : list typ) (u : Uf.T) (r : reg) (t : typ)
- (ros : reg+ident),
- consistent (fold2 type_rtl_arg (type_rtl_arg (type_rtl_ros u ros) r t) l1 l2)
- -> included u (fold2 type_rtl_arg (type_rtl_arg (type_rtl_ros u ros) r t) l1 l2).
+Lemma check_reg_correct:
+ forall r ty, check_reg r ty = true -> env r = ty.
Proof.
- intros.
- apply included_trans with (e2 := type_rtl_ros u ros).
- unfold type_rtl_ros; destruct ros.
- apply included_identify.
- apply included_refl.
- apply type_args_res_included; auto.
+ unfold check_reg; intros.
+ destruct (typ_eq (env r) ty). auto. discriminate.
Qed.
-Lemma type_instr_included :
- forall (p : positive) (i : instruction) (u : Uf.T) (res_ty : option typ),
- consistent (type_rtl_instr res_ty u p i)
- -> included u (type_rtl_instr res_ty u p i).
- Proof.
- intros.
- destruct i; simpl; simpl in H; try apply type_args_res_included; auto.
- apply included_refl; auto.
- destruct o; simpl; simpl in H; try apply type_args_res_included; auto.
- destruct l; simpl; simpl in H; auto.
- apply error_inconsistent; auto.
- destruct l; simpl; simpl in H; auto.
- apply included_identify.
- apply error_inconsistent; auto.
- destruct l; simpl; simpl in H; auto.
- apply included_refl.
- apply error_inconsistent; auto.
- apply type_args_res_ros_included; auto.
- apply type_args_included; auto.
- destruct res_ty; destruct o; simpl; simpl in H;
- try (apply error_inconsistent; auto; fail).
- apply type_arg_included.
- apply included_refl.
- Qed.
-
-Lemma type_instrs_extends :
- forall (l : list (positive * instruction)) (u : Uf.T) (res_ty : option typ),
- consistent
- (fold_left (fun v p => type_rtl_instr res_ty v (fst p) (snd p)) l u)
- -> consistent u.
+Lemma check_regs_correct:
+ forall rl tyl, check_regs rl tyl = true -> List.map env rl = tyl.
Proof.
- induction l; simpl; intros.
- auto.
- apply included_consistent
- with (u2 := (type_rtl_instr res_ty u (fst a) (snd a))).
- apply type_instr_included.
- apply IHl with (res_ty := res_ty); auto.
- apply IHl with (res_ty := res_ty); auto.
+ induction rl; destruct tyl; simpl; intros.
+ auto. discriminate. discriminate.
+ elim (andb_prop _ _ H); intros.
+ rewrite (check_reg_correct _ _ H0). rewrite (IHrl tyl H1). auto.
Qed.
-Lemma type_instrs_included :
- forall (l : list (positive * instruction)) (u : Uf.T) (res_ty : option typ),
- consistent
- (fold_left (fun v p => type_rtl_instr res_ty v (fst p) (snd p)) l u)
- -> included u
- (fold_left (fun v p => type_rtl_instr res_ty v (fst p) (snd p)) l u).
- Proof.
- induction l; simpl; intros.
- apply included_refl; auto.
- apply included_trans with (e2 := (type_rtl_instr res_ty u (fst a) (snd a))).
- apply type_instr_included.
- apply type_instrs_extends with (res_ty := res_ty) (l := l); auto.
- apply IHl; auto.
- Qed.
-
-Lemma step2 :
- forall (res_ty : option typ) (c : code) (u0 : Uf.T),
- consistent (PTree.fold (type_rtl_instr res_ty) c u0) ->
- forall (pc : positive) (i : instruction),
- c!pc = Some i
- -> exists u : Uf.T,
- consistent (type_rtl_instr res_ty u pc i)
- /\ included (type_rtl_instr res_ty u pc i)
- (PTree.fold (type_rtl_instr res_ty) c u0).
- Proof.
- intros.
- rewrite PTree.fold_spec.
- rewrite PTree.fold_spec in H.
- pose (H1 := PTree.elements_correct _ _ H0).
- generalize H. clear H.
- generalize u0. clear u0.
- generalize H1. clear H1.
- induction (PTree.elements c).
- intros.
- absurd (In (pc, i) nil).
- apply in_nil.
- auto.
- intros.
- simpl in H.
- elim H1.
- intro.
- rewrite H2 in H.
- simpl in H.
- rewrite H2. simpl.
- exists u0.
- apply conj.
- apply type_instrs_extends with (res_ty := res_ty) (l := l).
- auto.
- apply type_instrs_included.
- auto.
- intro.
- simpl.
- apply IHl.
- auto.
- auto.
- Qed.
-
-Definition mapped (u : Uf.T) (r : reg) :=
- Uf.repr u (tReg r) = Uf.repr u (tTy Tfloat)
- \/ Uf.repr u (tReg r) = Uf.repr u (tTy Tint).
-
-Definition definite (u : Uf.T) (i : instruction) :=
- match i with
- | Inop _ => True
- | Iop Omove (r1 :: nil) r0 _ => Uf.repr u (tReg r1) = Uf.repr u (tReg r0)
- | Iop Oundef _ _ _ => True
- | Iop _ args r0 _ =>
- mapped u r0 /\ forall r : reg, In r args -> mapped u r
- | Iload _ _ args r0 _ =>
- mapped u r0 /\ forall r : reg, In r args -> mapped u r
- | Istore _ _ args r0 _ =>
- mapped u r0 /\ forall r : reg, In r args -> mapped u r
- | Icall _ ros args r0 _ =>
- match ros with inl r => mapped u r | _ => True end
- /\ mapped u r0 /\ forall r : reg, In r args -> mapped u r
- | Icond _ args _ _ =>
- forall r : reg, In r args -> mapped u r
- | Ireturn None => True
- | Ireturn (Some r) => mapped u r
- end.
-
-Lemma type_arg_complete :
- forall (u : Uf.T) (r : reg) (t : typ),
- mapped (type_rtl_arg u r t) r.
-Proof.
- intros.
- unfold type_rtl_arg.
- unfold mapped.
- destruct t.
- right; apply Uf.sameclass_identify_1.
- left; apply Uf.sameclass_identify_1.
-Qed.
-
-Lemma type_arg_mapped :
- forall (u : Uf.T) (r r0 : reg) (t : typ),
- mapped u r0 -> mapped (type_rtl_arg u r t) r0.
-Proof.
- unfold mapped.
- unfold type_rtl_arg.
- intros.
- elim H; intros.
- left; apply Uf.sameclass_identify_2; auto.
- right; apply Uf.sameclass_identify_2; auto.
-Qed.
-
-Lemma type_args_mapped :
- forall (lr : list reg) (lt : list typ) (u : Uf.T) (r : reg),
- consistent (fold2 type_rtl_arg u lr lt) ->
- mapped u r ->
- mapped (fold2 type_rtl_arg u lr lt) r.
+Lemma check_op_correct:
+ forall op args res,
+ check_op op args res = true ->
+ (List.map env args, env res) = type_of_operation op.
Proof.
- induction lr; simpl; intros.
- destruct lt; simpl; auto; try (apply error_inconsistent; auto; fail).
- destruct lt; simpl; auto; try (apply error_inconsistent; auto; fail).
- apply IHlr.
+ unfold check_op; intros.
+ destruct (type_of_operation op) as [targs tres].
+ elim (andb_prop _ _ H); intros.
+ rewrite (check_regs_correct _ _ H0).
+ rewrite (check_reg_correct _ _ H1).
auto.
- apply type_arg_mapped; auto.
-Qed.
-
-Lemma type_args_complete :
- forall (lr : list reg) (lt : list typ) (u : Uf.T),
- consistent (fold2 type_rtl_arg u lr lt)
- -> forall r, (In r lr -> mapped (fold2 type_rtl_arg u lr lt) r).
-Proof.
- induction lr; simpl; intros.
- destruct lt; simpl; try tauto.
- destruct lt; simpl.
- apply error_inconsistent; auto.
- elim H0; intros.
- rewrite H1.
- rewrite H1 in H.
- apply type_args_mapped; auto.
- apply type_arg_complete.
- apply IHlr; auto.
-Qed.
-
-Lemma type_res_complete :
- forall (u : Uf.T) (lr : list reg) (lt : list typ) (r : reg) (t : typ),
- consistent (fold2 type_rtl_arg (type_rtl_arg u r t) lr lt)
- -> mapped (fold2 type_rtl_arg (type_rtl_arg u r t) lr lt) r.
-Proof.
- intros.
- apply type_args_mapped; auto.
- apply type_arg_complete.
Qed.
-Lemma type_args_res_complete :
- forall (u : Uf.T) (lr : list reg) (lt : list typ) (r : reg) (t : typ),
- consistent (fold2 type_rtl_arg (type_rtl_arg u r t) lr lt)
- -> mapped (fold2 type_rtl_arg (type_rtl_arg u r t) lr lt) r
- /\ forall rr, (In rr lr -> mapped (fold2 type_rtl_arg (type_rtl_arg u r t)
- lr lt)
- rr).
+Lemma check_instr_correct:
+ forall i, check_instr i = true -> wt_instr env funct.(fn_sig) i.
Proof.
- intros.
- apply conj.
- apply type_res_complete; auto.
- apply type_args_complete; auto.
+ unfold check_instr; intros; destruct i.
+ (* nop *)
+ constructor.
+ (* op *)
+ destruct o;
+ try (apply wt_Iop; [congruence|congruence|apply check_op_correct;auto]).
+ destruct l; try discriminate. destruct l; try discriminate.
+ destruct (typ_eq (env r0) (env r)); try discriminate.
+ apply wt_Iopmove; auto.
+ destruct l; try discriminate.
+ apply wt_Iopundef.
+ (* load *)
+ elim (andb_prop _ _ H); intros.
+ constructor. apply check_regs_correct; auto. apply check_reg_correct; auto.
+ (* store *)
+ elim (andb_prop _ _ H); intros.
+ constructor. apply check_regs_correct; auto. apply check_reg_correct; auto.
+ (* call *)
+ elim (andb_prop _ _ H); clear H; intros.
+ elim (andb_prop _ _ H); clear H; intros.
+ constructor.
+ destruct s0; auto. apply check_reg_correct; auto.
+ apply check_regs_correct; auto.
+ apply check_reg_correct; auto.
+ (* alloc *)
+ elim (andb_prop _ _ H); intros.
+ constructor; apply check_reg_correct; auto.
+ (* cond *)
+ constructor. apply check_regs_correct; auto.
+ (* return *)
+ constructor.
+ destruct o; simpl; destruct funct.(fn_sig).(sig_res); try discriminate.
+ rewrite (check_reg_correct _ _ H); auto.
+ auto.
Qed.
-Lemma type_ros_complete :
- forall (u : Uf.T) (lr : list reg) (lt : list typ) (r r1 : reg) (t : typ),
- consistent (fold2 type_rtl_arg (type_rtl_arg
- (type_rtl_ros u (inl ident r1)) r t) lr lt)
- ->
- mapped (fold2 type_rtl_arg (type_rtl_arg
- (type_rtl_ros u (inl ident r1)) r t) lr lt) r1.
+Lemma check_instrs_correct:
+ forall instrs,
+ check_instrs instrs = true ->
+ forall pc i, In (pc, i) instrs -> wt_instr env funct.(fn_sig) i.
Proof.
- intros.
- apply type_args_mapped; auto.
- apply type_arg_mapped.
- unfold type_rtl_ros.
- unfold mapped.
- right.
- apply Uf.sameclass_identify_1; auto.
+ induction instrs; simpl; intros.
+ elim H0.
+ destruct a as [pc' i']. elim (andb_prop _ _ H); clear H; intros.
+ elim H0; intro.
+ inversion H2; subst pc' i'. apply check_instr_correct; auto.
+ eauto.
Qed.
-Lemma type_res_correct :
- forall (u : Uf.T) (lr : list reg) (lt : list typ) (r : reg) (t : typ),
- consistent (fold2 type_rtl_arg (type_rtl_arg u r t) lr lt)
- -> mk_env (fold2 type_rtl_arg (type_rtl_arg u r t) lr lt) r = t.
-Proof.
- intros.
- unfold mk_env.
- rewrite (type_args_included _ _ _ H (tReg r) (tTy t)).
- destruct t.
- apply consistent_not_eq; auto.
- apply equal_eq; auto.
- unfold type_rtl_arg; apply Uf.sameclass_identify_1; auto.
-Qed.
+End TYPECHECKING.
-Lemma type_ros_correct :
- forall (u : Uf.T) (lr : list reg) (lt : list typ) (r r1 : reg) (t : typ),
- consistent (fold2 type_rtl_arg (type_rtl_arg
- (type_rtl_ros u (inl ident r1)) r t) lr lt)
- ->
- mk_env (fold2 type_rtl_arg (type_rtl_arg
- (type_rtl_ros u (inl ident r1)) r t) lr lt) r1
- = Tint.
-Proof.
- intros.
- unfold mk_env.
- rewrite (type_args_included _ _ _ H (tReg r1) (tTy Tint)).
- apply consistent_not_eq; auto.
- rewrite (type_arg_included (type_rtl_ros u (inl ident r1)) r t (tReg r1) (tTy Tint)).
- auto.
- simpl.
- apply Uf.sameclass_identify_1; auto.
-Qed.
-
-Lemma step3 :
- forall (u : Uf.T) (f : function) (c : code) (i : instruction) (pc : positive),
- c!pc = Some i ->
- consistent (type_rtl_instr f.(fn_sig).(sig_res) u pc i)
- -> wt_instr (mk_env (type_rtl_instr f.(fn_sig).(sig_res) u pc i)) f i
- /\ definite (type_rtl_instr f.(fn_sig).(sig_res) u pc i) i.
- Proof.
- Opaque type_rtl_arg.
- intros.
- destruct i; simpl in H0; simpl.
- (* Inop *)
- apply conj; auto. apply wt_Inop.
- (* Iop *)
- destruct o;
- try (apply conj; [
- apply wt_Iop; try congruence; simpl;
- rewrite (type_args_correct _ _ _ H0);
- rewrite (type_res_correct _ _ _ _ _ H0);
- auto
- |apply (type_args_res_complete _ _ _ _ _ H0)]).
- (* Omove *)
- destruct l; [apply error_inconsistent; auto | idtac].
- destruct l; [idtac | apply error_inconsistent; auto].
- apply conj.
- apply wt_Iopmove.
- simpl.
- unfold mk_env.
- rewrite Uf.sameclass_identify_1.
- congruence.
- simpl.
- rewrite Uf.sameclass_identify_1; congruence.
- (* Oundef *)
- destruct l; [idtac | apply error_inconsistent; auto].
- apply conj. apply wt_Iopundef.
- unfold definite. auto.
- (* Iload *)
- apply conj.
- apply wt_Iload.
- rewrite (type_args_correct _ _ _ H0); auto.
- rewrite (type_res_correct _ _ _ _ _ H0); auto.
- simpl; apply (type_args_res_complete _ _ _ _ _ H0).
- (* IStore *)
- apply conj.
- apply wt_Istore.
- rewrite (type_args_correct _ _ _ H0); auto.
- rewrite (type_res_correct _ _ _ _ _ H0); auto.
- simpl; apply (type_args_res_complete _ _ _ _ _ H0).
- (* Icall *)
- apply conj.
- apply wt_Icall.
- destruct s0; auto. apply type_ros_correct. auto.
- apply type_args_correct. auto.
- fold (type_of_sig_res s). apply type_res_correct. auto.
- destruct s0.
- apply conj.
- apply type_ros_complete. auto.
- apply type_args_res_complete. auto.
- apply conj; auto.
- apply type_args_res_complete. auto.
- (* Icond *)
- apply conj.
- apply wt_Icond.
- apply (type_args_correct _ _ _ H0).
- simpl; apply (type_args_complete _ _ _ H0).
- (* Ireturn *)
- destruct o; simpl.
- apply conj.
- apply wt_Ireturn.
- destruct f.(fn_sig).(sig_res); simpl; simpl in H0.
- rewrite type_arg_correct; auto.
- apply error_inconsistent; auto.
- destruct f.(fn_sig).(sig_res); simpl; simpl in H0.
- apply type_arg_complete.
- apply error_inconsistent; auto.
- apply conj; auto. apply wt_Ireturn.
- destruct f.(fn_sig).(sig_res); simpl; simpl in H0.
- apply error_inconsistent; auto.
- congruence.
- Transparent type_rtl_arg.
- Qed.
-
-Lemma mapped_included_consistent :
- forall (u1 u2 : Uf.T) (r : reg),
- mapped u1 r ->
- included u1 u2 ->
- consistent u2 ->
- mk_env u2 r = mk_env u1 r.
-Proof.
- intros.
- unfold mk_env.
- unfold mapped in H.
- elim H; intros; rewrite H2; rewrite (H0 _ _ H2).
- rewrite equal_eq; rewrite equal_eq; auto.
- rewrite (consistent_not_eq u2).
- rewrite (consistent_not_eq u1).
- auto.
- apply included_consistent with (u2 := u2).
- auto.
- auto.
- auto.
-Qed.
+(** ** The type inference function **)
-Lemma mapped_list_included :
- forall (u1 u2 : Uf.T) (lr : list reg),
- (forall r, In r lr -> mapped u1 r) ->
- included u1 u2 ->
- consistent u2 ->
- map (mk_env u2) lr = map (mk_env u1) lr.
-Proof.
- induction lr; simpl; intros.
- auto.
- rewrite (mapped_included_consistent u1 u2 a).
- rewrite IHlr; auto.
- apply (H a); intros.
- left; auto.
- auto.
- auto.
-Qed.
+Definition type_function (f: function): option regenv :=
+ let instrs := PTree.elements f.(fn_code) in
+ match infer_type_environment f instrs with
+ | None => None
+ | Some env =>
+ if check_regs env f.(fn_params) f.(fn_sig).(sig_args)
+ && check_params_norepet f.(fn_params)
+ && check_instrs f env instrs
+ then Some env else None
+ end.
-Lemma included_mapped :
- forall (u1 u2 : Uf.T) (r : reg),
- included u1 u2 ->
- mapped u1 r ->
- mapped u2 r.
-Proof.
- unfold mapped.
- intros.
- elim H0; intros.
- left; rewrite (H _ _ H1); auto.
- right; rewrite (H _ _ H1); auto.
-Qed.
+Definition type_external_function (ef: external_function): bool :=
+ List.fold_right
+ (fun l b => match l with Locations.S _ => false | Locations.R _ => b end)
+ true (Conventions.loc_arguments ef.(ef_sig)).
-Lemma included_mapped_forall :
- forall (u1 u2 : Uf.T) (r : reg) (l : list reg),
- included u1 u2 ->
- mapped u1 r /\ (forall r, In r l -> mapped u1 r) ->
- mapped u2 r /\ (forall r, In r l -> mapped u2 r).
+Lemma type_function_correct:
+ forall f env,
+ type_function f = Some env ->
+ wt_function f env.
Proof.
- intros.
- elim H0; intros.
- apply conj.
- apply (included_mapped _ _ r H); auto.
- intros.
- apply (included_mapped _ _ r0 H).
- apply H2; auto.
+ unfold type_function; intros until env.
+ set (instrs := PTree.elements f.(fn_code)).
+ case (infer_type_environment f instrs).
+ intro env'.
+ caseEq (check_regs env' f.(fn_params) f.(fn_sig).(sig_args)); intro; simpl; try congruence.
+ caseEq (check_params_norepet f.(fn_params)); intro; simpl; try congruence.
+ caseEq (check_instrs f env' instrs); intro; simpl; try congruence.
+ intro EQ; inversion EQ; subst env'.
+ constructor.
+ apply check_regs_correct; auto.
+ unfold check_params_norepet in H0.
+ destruct (list_norepet_dec Reg.eq (fn_params f)). auto. discriminate.
+ intros. eapply check_instrs_correct. eauto.
+ unfold instrs. apply PTree.elements_correct. eauto.
+ congruence.
Qed.
-Lemma definite_included :
- forall (u1 u2 : Uf.T) (i : instruction),
- included u1 u2 -> definite u1 i -> definite u2 i.
+Lemma type_external_function_correct:
+ forall ef,
+ type_external_function ef = true ->
+ Conventions.sig_external_ok ef.(ef_sig).
Proof.
- unfold definite.
- intros.
- destruct i; try apply (included_mapped_forall _ _ _ _ H H0); auto.
- destruct o; try apply (included_mapped_forall _ _ _ _ H H0); auto.
- destruct l; auto.
- apply (included_mapped_forall _ _ _ _ H H0).
- destruct l; auto.
- apply (included_mapped_forall _ _ _ _ H H0).
- destruct s0; auto.
- elim H0; intros.
- apply conj.
- apply (included_mapped _ _ _ H H1).
- apply (included_mapped_forall _ _ _ _ H H2).
- elim H0; intros.
- apply conj; auto.
- apply (included_mapped_forall _ _ _ _ H H2).
- intros.
- apply (included_mapped _ _ _ H (H0 r H1)).
- destruct o; auto.
- apply (included_mapped _ _ _ H H0).
+ intro ef. unfold type_external_function, Conventions.sig_external_ok.
+ generalize (Conventions.loc_arguments (ef_sig ef)).
+ induction l; simpl.
+ tauto.
+ destruct a. intros. firstorder congruence.
+ congruence.
Qed.
-Lemma step4 :
- forall (f : function) (u1 u3 : Uf.T) (i : instruction),
- included u3 u1 ->
- wt_instr (mk_env u3) f i ->
- definite u3 i ->
- consistent u1 ->
- wt_instr (mk_env u1) f i.
- Proof.
- intros f u1 u3 i H1 H H0 X.
- destruct H; try simpl in H0; try (elim H0; intros).
- apply wt_Inop.
- apply wt_Iopmove. unfold mk_env. rewrite (H1 _ _ H0). auto.
- apply wt_Iopundef.
- apply wt_Iop; auto.
- destruct op; try congruence; simpl; simpl in H3;
- simpl in H0; elim H0; intros; rewrite (mapped_included_consistent _ _ _ H4 H1 X);
- rewrite (mapped_list_included _ _ _ H5 H1); auto.
- apply wt_Iload.
- rewrite (mapped_list_included _ _ _ H4 H1); auto.
- rewrite (mapped_included_consistent _ _ _ H3 H1 X). auto.
- apply wt_Istore.
- rewrite (mapped_list_included _ _ _ H4 H1); auto.
- rewrite (mapped_included_consistent _ _ _ H3 H1 X). auto.
- elim H5; intros; destruct ros; apply wt_Icall.
- rewrite (mapped_included_consistent _ _ _ H4 H1 X); auto.
- rewrite (mapped_list_included _ _ _ H7 H1); auto.
- rewrite (mapped_included_consistent _ _ _ H6 H1 X); auto.
- auto.
- rewrite (mapped_list_included _ _ _ H7 H1); auto.
- rewrite (mapped_included_consistent _ _ _ H6 H1 X); auto.
- apply wt_Icond. rewrite (mapped_list_included _ _ _ H0 H1); auto.
- apply wt_Ireturn.
- destruct optres; destruct f.(fn_sig).(sig_res);
- simpl in H; simpl; try congruence.
- rewrite (mapped_included_consistent _ _ _ H0 H1 X); auto.
- Qed.
-
-Lemma type_rtl_function_instrs :
- forall (f: function) (env: regenv),
- type_rtl_function f = Some env
- -> forall pc i, f.(fn_code)!pc = Some i -> wt_instr env f i.
- Proof.
- intros.
- elim (step1 _ _ H).
- intros.
- elim H1.
- intros.
- elim H3.
- intros.
- rewrite H4.
- elim (step2 _ _ _ (included_consistent _ _ H2 H5) _ _ H0).
- intros.
- elim H6. intros.
- elim (step3 x0 f _ _ _ H0); auto. intros.
- apply (step4 f _ _ i H2); auto.
- apply (step4 _ _ _ _ H8 H9 H10).
- apply (included_consistent _ _ H2); auto.
- apply (definite_included _ _ _ H8 H10); auto.
- Qed.
-
-(** Combining the sub-proofs. *)
-
-Theorem type_rtl_function_correct:
- forall (f: function) (env: regenv),
- type_rtl_function f = Some env -> wt_function env f.
- Proof.
- intros.
- exact (mk_wt_function env f (type_rtl_function_params f _ H)
- (type_rtl_function_norepet f _ H)
- (type_rtl_function_instrs f _ H)).
- Qed.
-
-Definition wt_program (p: program) : Prop :=
- forall i f, In (i, f) (prog_funct p) -> type_rtl_function f <> None.
-
(** * Type preservation during evaluation *)
(** The type system for RTL is not sound in that it does not guarantee
@@ -1143,6 +358,7 @@ Require Import Globalenvs.
Require Import Values.
Require Import Mem.
Require Import Integers.
+Require Import Events.
Definition wt_regset (env: regenv) (rs: regset) : Prop :=
forall r, Val.has_type (rs#r) (env r).
@@ -1180,6 +396,14 @@ Proof.
apply wt_regset_assign; auto.
Qed.
+Lemma wt_event_match:
+ forall ef args t res,
+ event_match ef args t res ->
+ Val.has_type res (proj_sig_res ef.(ef_sig)).
+Proof.
+ induction 1. inversion H0; exact I.
+Qed.
+
Section SUBJECT_REDUCTION.
Variable p: program.
@@ -1190,26 +414,24 @@ Let ge := Genv.globalenv p.
Definition exec_instr_subject_reduction
(c: code) (sp: val)
- (pc: node) (rs: regset) (m: mem)
+ (pc: node) (rs: regset) (m: mem) (t: trace)
(pc': node) (rs': regset) (m': mem) : Prop :=
forall env f
(CODE: c = fn_code f)
- (WT_FN: wt_function env f)
+ (WT_FN: wt_function f env)
(WT_RS: wt_regset env rs),
wt_regset env rs'.
Definition exec_function_subject_reduction
- (f: function) (args: list val) (m: mem) (res: val) (m': mem) : Prop :=
- forall env,
- wt_function env f ->
- Val.has_type_list args f.(fn_sig).(sig_args) ->
- Val.has_type res
- (match f.(fn_sig).(sig_res) with None => Tint | Some ty => ty end).
+ (f: fundef) (args: list val) (m: mem) (t: trace) (res: val) (m': mem) : Prop :=
+ wt_fundef f ->
+ Val.has_type_list args (sig_args (funsig f)) ->
+ Val.has_type res (proj_sig_res (funsig f)).
Lemma subject_reduction:
- forall f args m res m',
- exec_function ge f args m res m' ->
- exec_function_subject_reduction f args m res m'.
+ forall f args m t res m',
+ exec_function ge f args m t res m' ->
+ exec_function_subject_reduction f args m t res m'.
Proof.
apply (exec_function_ind_3 ge
exec_instr_subject_reduction
@@ -1217,7 +439,7 @@ Proof.
exec_function_subject_reduction);
intros; red; intros;
try (rewrite CODE in H;
- generalize (wt_instrs env _ WT_FN pc _ H);
+ generalize (wt_instrs _ _ WT_FN pc _ H);
intro WT_INSTR;
inversion WT_INSTR).
@@ -1233,7 +455,7 @@ Proof.
apply wt_regset_assign. auto.
replace (env res) with (snd (type_of_operation op)).
- apply type_of_operation_sound with function ge rs##args sp; auto.
+ apply type_of_operation_sound with fundef ge rs##args sp; auto.
rewrite <- H7. reflexivity.
apply wt_regset_assign. auto. rewrite H8.
@@ -1241,37 +463,39 @@ Proof.
assumption.
- apply wt_regset_assign. auto. rewrite H11. rewrite H1.
- assert (type_rtl_function f <> None).
+ apply wt_regset_assign. auto. rewrite H11. rewrite <- H1.
+ assert (wt_fundef f).
destruct ros; simpl in H0.
- pattern f. apply Genv.find_funct_prop with function p (rs#r).
+ pattern f. apply Genv.find_funct_prop with fundef p (rs#r).
exact wt_p. exact H0.
caseEq (Genv.find_symbol ge i); intros; rewrite H12 in H0.
- pattern f. apply Genv.find_funct_ptr_prop with function p b.
+ pattern f. apply Genv.find_funct_ptr_prop with fundef p b.
exact wt_p. exact H0.
discriminate.
- assert (exists env1, wt_function env1 f).
- caseEq (type_rtl_function f); intros; try congruence.
- exists t. apply type_rtl_function_correct; auto.
- elim H13; intros env1 WT_FN1.
- eapply H3. eexact WT_FN1. rewrite <- H1. rewrite <- H10.
+ eapply H3. auto. rewrite H1. rewrite <- H10.
apply wt_regset_list; auto.
+ apply wt_regset_assign. auto. rewrite H6; exact I.
+
assumption.
assumption.
assumption.
eauto.
eauto.
+ inversion H4; subst f0.
assert (WT_INIT: wt_regset env (init_regs args (fn_params f))).
- apply wt_init_regs. rewrite (wt_params env f H4). assumption.
- generalize (H1 env f (refl_equal (fn_code f)) H4 WT_INIT).
+ apply wt_init_regs. rewrite (wt_params _ _ H7). assumption.
+ generalize (H1 env f (refl_equal (fn_code f)) H7 WT_INIT).
intro WT_RS.
- generalize (wt_instrs env _ H4 pc _ H2).
+ generalize (wt_instrs _ _ H7 pc _ H2).
intro WT_INSTR; inversion WT_INSTR.
- destruct or; simpl in H3; simpl in H7; rewrite <- H7.
+ unfold proj_sig_res; simpl.
+ destruct or; simpl in H3; simpl in H8; rewrite <- H8.
rewrite H3. apply WT_RS.
rewrite H3. simpl; auto.
+
+ simpl. eapply wt_event_match; eauto.
Qed.
End SUBJECT_REDUCTION.
diff --git a/backend/Stacking.v b/backend/Stacking.v
index 1f0c4542..85ac9335 100644
--- a/backend/Stacking.v
+++ b/backend/Stacking.v
@@ -186,6 +186,8 @@ Definition transl_instr
Mstore chunk (transl_addr fe addr) args src :: k
| Lcall sig ros =>
Mcall sig ros :: k
+ | Lalloc =>
+ Malloc :: k
| Llabel lbl =>
Mlabel lbl :: k
| Lgoto lbl =>
@@ -222,5 +224,8 @@ Definition transf_function (f: Linear.function) : option Mach.function :=
f.(Linear.fn_stacksize)
fe.(fe_size)).
+Definition transf_fundef (f: Linear.fundef) : option Mach.fundef :=
+ AST.transf_partial_fundef transf_function f.
+
Definition transf_program (p: Linear.program) : option Mach.program :=
- transform_partial_program transf_function p.
+ transform_partial_program transf_fundef p.
diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v
index 002ca8d5..96926707 100644
--- a/backend/Stackingproof.v
+++ b/backend/Stackingproof.v
@@ -15,6 +15,7 @@ Require Import Integers.
Require Import Values.
Require Import Op.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Locations.
Require Import Mach.
@@ -299,7 +300,7 @@ Lemma exec_Mgetstack':
get_slot fr ty (offset_of_index fe idx) v ->
Machabstr.exec_instrs tge tf sp parent
(Mgetstack (Int.repr (offset_of_index fe idx)) ty dst :: c) rs fr m
- c (rs#dst <- v) fr m.
+ E0 c (rs#dst <- v) fr m.
Proof.
intros. apply Machabstr.exec_one. apply Machabstr.exec_Mgetstack.
rewrite offset_of_index_no_overflow. auto. auto.
@@ -311,7 +312,7 @@ Lemma exec_Msetstack':
set_slot fr ty (offset_of_index fe idx) (rs src) fr' ->
Machabstr.exec_instrs tge tf sp parent
(Msetstack src (Int.repr (offset_of_index fe idx)) ty :: c) rs fr m
- c rs fr' m.
+ E0 c rs fr' m.
Proof.
intros. apply Machabstr.exec_one. apply Machabstr.exec_Msetstack.
rewrite offset_of_index_no_overflow. auto. auto.
@@ -632,7 +633,7 @@ Lemma save_int_callee_save_correct_rec:
exists fr',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (save_int_callee_save fe) k l) rs fr m
- k rs fr' m
+ E0 k rs fr' m
/\ fr'.(high) = 0
/\ fr'.(low) = -fe.(fe_size)
/\ (forall r,
@@ -670,7 +671,7 @@ Proof.
intros [fr' [A [B [C [D E]]]]].
exists fr'.
split. eapply Machabstr.exec_trans. apply exec_Msetstack'; eauto with stacking.
- exact A.
+ eexact A. traceEq.
split. auto.
split. auto.
split. intros. elim H3; intros. subst r.
@@ -701,7 +702,7 @@ Lemma save_int_callee_save_correct:
exists fr',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (save_int_callee_save fe) k int_callee_save_regs) rs fr m
- k rs fr' m
+ E0 k rs fr' m
/\ fr'.(high) = 0
/\ fr'.(low) = -fe.(fe_size)
/\ (forall r,
@@ -733,7 +734,7 @@ Lemma save_float_callee_save_correct_rec:
exists fr',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (save_float_callee_save fe) k l) rs fr m
- k rs fr' m
+ E0 k rs fr' m
/\ fr'.(high) = 0
/\ fr'.(low) = -fe.(fe_size)
/\ (forall r,
@@ -771,7 +772,7 @@ Proof.
intros [fr' [A [B [C [D E]]]]].
exists fr'.
split. eapply Machabstr.exec_trans. apply exec_Msetstack'; eauto with stacking.
- exact A.
+ eexact A. traceEq.
split. auto.
split. auto.
split. intros. elim H3; intros. subst r.
@@ -802,7 +803,7 @@ Lemma save_float_callee_save_correct:
exists fr',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (save_float_callee_save fe) k float_callee_save_regs) rs fr m
- k rs fr' m
+ E0 k rs fr' m
/\ fr'.(high) = 0
/\ fr'.(low) = -fe.(fe_size)
/\ (forall r,
@@ -846,7 +847,7 @@ Lemma save_callee_save_correct:
exists fr',
Machabstr.exec_instrs tge tf sp parent
(save_callee_save fe k) rs fr m
- k rs fr' m
+ E0 k rs fr' m
/\ agree (LTL.call_regs ls) rs fr' parent rs.
Proof.
intros. unfold save_callee_save.
@@ -857,7 +858,7 @@ Proof.
generalize (save_float_callee_save_correct k sp parent rs fr1 m B1 C1).
intros [fr2 [A2 [B2 [C2 [D2 E2]]]]].
exists fr2.
- split. eapply Machabstr.exec_trans. eexact A1. eexact A2.
+ split. eapply Machabstr.exec_trans. eexact A1. eexact A2. traceEq.
constructor; unfold LTL.call_regs; auto.
(* agree_local *)
intros. rewrite E2; auto with stacking.
@@ -886,7 +887,7 @@ Lemma restore_int_callee_save_correct_rec:
exists ls', exists rs',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (restore_int_callee_save fe) k l) rs fr m
- k rs' fr m
+ E0 k rs' fr m
/\ (forall r, In r l -> rs' r = rs0 r)
/\ (forall r, ~(In r l) -> rs' r = rs r)
/\ agree ls' rs' fr parent rs0.
@@ -916,11 +917,11 @@ Proof.
generalize (IHl ls1 rs1 R1 R2 R3).
intros [ls' [rs' [A [B [C D]]]]].
exists ls'. exists rs'.
- split. apply Machabstr.exec_trans with k1 rs1 fr m.
+ split. apply Machabstr.exec_trans with E0 k1 rs1 fr m E0.
unfold rs1; apply exec_Mgetstack'; eauto with stacking.
apply get_slot_index; eauto with stacking.
symmetry. eauto with stacking.
- eauto with stacking.
+ eauto with stacking. traceEq.
split. intros. elim H2; intros.
subst r. rewrite C. unfold rs1. apply Regmap.gss. inversion H0; auto.
auto.
@@ -945,7 +946,7 @@ Lemma restore_int_callee_save_correct:
exists ls', exists rs',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (restore_int_callee_save fe) k int_callee_save_regs) rs fr m
- k rs' fr m
+ E0 k rs' fr m
/\ (forall r, In r int_callee_save_regs -> rs' r = rs0 r)
/\ (forall r, ~(In r int_callee_save_regs) -> rs' r = rs r)
/\ agree ls' rs' fr parent rs0.
@@ -962,7 +963,7 @@ Lemma restore_float_callee_save_correct_rec:
exists ls', exists rs',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (restore_float_callee_save fe) k l) rs fr m
- k rs' fr m
+ E0 k rs' fr m
/\ (forall r, In r l -> rs' r = rs0 r)
/\ (forall r, ~(In r l) -> rs' r = rs r)
/\ agree ls' rs' fr parent rs0.
@@ -992,11 +993,11 @@ Proof.
generalize (IHl ls1 rs1 R1 R2 R3).
intros [ls' [rs' [A [B [C D]]]]].
exists ls'. exists rs'.
- split. apply Machabstr.exec_trans with k1 rs1 fr m.
+ split. apply Machabstr.exec_trans with E0 k1 rs1 fr m E0.
unfold rs1; apply exec_Mgetstack'; eauto with stacking.
apply get_slot_index; eauto with stacking.
symmetry. eauto with stacking.
- exact A.
+ exact A. traceEq.
split. intros. elim H2; intros.
subst r. rewrite C. unfold rs1. apply Regmap.gss. inversion H0; auto.
auto.
@@ -1021,7 +1022,7 @@ Lemma restore_float_callee_save_correct:
exists ls', exists rs',
Machabstr.exec_instrs tge tf sp parent
(List.fold_right (restore_float_callee_save fe) k float_callee_save_regs) rs fr m
- k rs' fr m
+ E0 k rs' fr m
/\ (forall r, In r float_callee_save_regs -> rs' r = rs0 r)
/\ (forall r, ~(In r float_callee_save_regs) -> rs' r = rs r)
/\ agree ls' rs' fr parent rs0.
@@ -1036,7 +1037,7 @@ Lemma restore_callee_save_correct:
exists rs',
Machabstr.exec_instrs tge tf sp parent
(restore_callee_save fe k) rs fr m
- k rs' fr m
+ E0 k rs' fr m
/\ (forall r,
In r int_callee_save_regs \/ In r float_callee_save_regs ->
rs' r = rs0 r)
@@ -1053,7 +1054,7 @@ Proof.
generalize (restore_float_callee_save_correct sp parent
k fr m rs0 ls1 rs1 D).
intros [ls2 [rs2 [P [Q [R S]]]]].
- exists rs2. split. eapply Machabstr.exec_trans. eexact A. exact P.
+ exists rs2. split. eapply Machabstr.exec_trans. eexact A. eexact P. traceEq.
split. intros. elim H0; intros.
rewrite R. apply B. auto. apply list_disjoint_notin with int_callee_save_regs.
apply int_float_callee_save_disjoint. auto.
@@ -1148,8 +1149,8 @@ Proof.
Qed.
Lemma exec_instr_incl:
- forall f sp c1 ls1 m1 c2 ls2 m2,
- Linear.exec_instr ge f sp c1 ls1 m1 c2 ls2 m2 ->
+ forall f sp c1 ls1 m1 t c2 ls2 m2,
+ Linear.exec_instr ge f sp c1 ls1 m1 t c2 ls2 m2 ->
incl c1 f.(fn_code) ->
incl c2 f.(fn_code).
Proof.
@@ -1159,8 +1160,8 @@ Proof.
Qed.
Lemma exec_instrs_incl:
- forall f sp c1 ls1 m1 c2 ls2 m2,
- Linear.exec_instrs ge f sp c1 ls1 m1 c2 ls2 m2 ->
+ forall f sp c1 ls1 m1 t c2 ls2 m2,
+ Linear.exec_instrs ge f sp c1 ls1 m1 t c2 ls2 m2 ->
incl c1 f.(fn_code) ->
incl c2 f.(fn_code).
Proof.
@@ -1174,7 +1175,7 @@ Lemma symbols_preserved:
forall id, Genv.find_symbol tge id = Genv.find_symbol ge id.
Proof.
intros. unfold ge, tge.
- apply Genv.find_symbol_transf_partial with transf_function.
+ apply Genv.find_symbol_transf_partial with transf_fundef.
exact TRANSF.
Qed.
@@ -1182,11 +1183,11 @@ Lemma functions_translated:
forall f v,
Genv.find_funct ge v = Some f ->
exists tf,
- Genv.find_funct tge v = Some tf /\ transf_function f = Some tf.
+ Genv.find_funct tge v = Some tf /\ transf_fundef f = Some tf.
Proof.
intros.
- generalize (Genv.find_funct_transf_partial transf_function TRANSF H).
- case (transf_function f).
+ generalize (Genv.find_funct_transf_partial transf_fundef TRANSF H).
+ case (transf_fundef f).
intros tf [A B]. exists tf. tauto.
intros. tauto.
Qed.
@@ -1195,15 +1196,26 @@ Lemma function_ptr_translated:
forall f v,
Genv.find_funct_ptr ge v = Some f ->
exists tf,
- Genv.find_funct_ptr tge v = Some tf /\ transf_function f = Some tf.
+ Genv.find_funct_ptr tge v = Some tf /\ transf_fundef f = Some tf.
Proof.
intros.
- generalize (Genv.find_funct_ptr_transf_partial transf_function TRANSF H).
- case (transf_function f).
+ generalize (Genv.find_funct_ptr_transf_partial transf_fundef TRANSF H).
+ case (transf_fundef f).
intros tf [A B]. exists tf. tauto.
intros. tauto.
Qed.
+Lemma sig_preserved:
+ forall f tf, transf_fundef f = Some tf -> Mach.funsig tf = Linear.funsig f.
+Proof.
+ intros until tf; unfold transf_fundef, transf_partial_fundef.
+ destruct f. unfold transf_function.
+ destruct (zlt (fn_stacksize f) 0). congruence.
+ destruct (zlt (- Int.min_signed) (fe_size (make_env (function_bounds f)))). congruence.
+ intros. inversion H; reflexivity.
+ intro. inversion H. reflexivity.
+Qed.
+
(** Correctness of stack pointer relocation in operations and
addressing modes. *)
@@ -1287,7 +1299,7 @@ Qed.
Definition exec_instr_prop
(f: function) (sp: val)
- (c: code) (ls: locset) (m: mem)
+ (c: code) (ls: locset) (m: mem) (t: trace)
(c': code) (ls': locset) (m': mem) :=
forall tf rs fr parent rs0
(TRANSL: transf_function f = Some tf)
@@ -1297,7 +1309,7 @@ Definition exec_instr_prop
exists rs', exists fr',
Machabstr.exec_instrs tge tf (shift_sp tf sp) parent
(transl_code (make_env (function_bounds f)) c) rs fr m
- (transl_code (make_env (function_bounds f)) c') rs' fr' m'
+ t (transl_code (make_env (function_bounds f)) c') rs' fr' m'
/\ agree f ls' rs' fr' parent rs0.
(** The simulation property for function calls has different preconditions
@@ -1306,19 +1318,19 @@ Definition exec_instr_prop
postconditions (preservation of callee-save registers). *)
Definition exec_function_prop
- (f: function)
- (ls: locset) (m: mem)
+ (f: fundef)
+ (ls: locset) (m: mem) (t: trace)
(ls': locset) (m': mem) :=
forall tf rs parent
- (TRANSL: transf_function f = Some tf)
- (WTF: wt_function f)
+ (TRANSL: transf_fundef f = Some tf)
+ (WTF: wt_fundef f)
(AG1: forall r, rs r = ls (R r))
(AG2: forall ofs ty,
6 <= ofs ->
- ofs + typesize ty <= size_arguments f.(fn_sig) ->
+ ofs + typesize ty <= size_arguments (funsig f) ->
get_slot parent ty (Int.signed (Int.repr (4 * ofs))) (ls (S (Outgoing ofs ty)))),
exists rs',
- Machabstr.exec_function tge tf parent rs m rs' m'
+ Machabstr.exec_function tge tf parent rs m t rs' m'
/\ (forall r,
In (R r) temporaries \/ In (R r) destroyed_at_call ->
rs' r = ls' (R r))
@@ -1329,9 +1341,9 @@ Definition exec_function_prop
Hypothesis wt_prog: wt_program prog.
Lemma transf_function_correct:
- forall f ls m ls' m',
- Linear.exec_function ge f ls m ls' m' ->
- exec_function_prop f ls m ls' m'.
+ forall f ls m t ls' m',
+ Linear.exec_function ge f ls m t ls' m' ->
+ exec_function_prop f ls m t ls' m'.
Proof.
assert (RED: forall f i c,
transl_code (make_env (function_bounds f)) (i :: c) =
@@ -1412,13 +1424,13 @@ Proof.
(* Lcall *)
inversion WTI.
- assert (WTF': wt_function f').
+ assert (WTF': wt_fundef f').
destruct ros; simpl in H.
- apply (Genv.find_funct_prop wt_function wt_prog H).
+ apply (Genv.find_funct_prop wt_fundef wt_prog H).
destruct (Genv.find_symbol ge i); try discriminate.
- apply (Genv.find_funct_ptr_prop wt_function wt_prog H).
+ apply (Genv.find_funct_ptr_prop wt_fundef wt_prog H).
assert (TR: exists tf', Mach.find_function tge ros rs0 = Some tf'
- /\ transf_function f' = Some tf').
+ /\ transf_fundef f' = Some tf').
destruct ros; simpl in H; simpl.
eapply functions_translated.
rewrite (agree_eval_reg _ _ _ _ _ _ m0 AG). auto.
@@ -1430,7 +1442,7 @@ Proof.
intro. symmetry. eapply agree_reg; eauto.
assert (AG2: forall ofs ty,
6 <= ofs ->
- ofs + typesize ty <= size_arguments f'.(fn_sig) ->
+ ofs + typesize ty <= size_arguments (funsig f') ->
get_slot fr ty (Int.signed (Int.repr (4 * ofs))) (rs (S (Outgoing ofs ty)))).
intros.
assert (slot_bounded f (Outgoing ofs ty)).
@@ -1446,6 +1458,13 @@ Proof.
apply Machabstr.exec_one; apply Machabstr.exec_Mcall with tf'; auto.
apply agree_return_regs with rs0; auto.
+ (* Lalloc *)
+ exists (rs0#loc_alloc_result <- (Vptr blk Int.zero)); exists fr. split.
+ apply Machabstr.exec_one; eapply Machabstr.exec_Malloc; eauto.
+ rewrite (agree_eval_reg _ _ _ _ _ _ loc_alloc_argument AG). auto.
+ apply agree_set_reg; auto.
+ red. simpl. generalize (max_over_regs_of_funct_pos f int_callee_save). omega.
+
(* Llabel *)
exists rs0; exists fr. split.
apply Machabstr.exec_one; apply Machabstr.exec_Mlabel.
@@ -1483,25 +1502,30 @@ Proof.
generalize (H2 tf rs' fr' parent rs0 TRANSL WTF B INCL').
intros [rs'' [fr'' [C D]]].
exists rs''; exists fr''. split.
- eapply Machabstr.exec_trans. eexact A. eexact C.
+ eapply Machabstr.exec_trans. eexact A. eexact C. auto.
auto.
(* function *)
+ generalize TRANSL; clear TRANSL.
+ unfold transf_fundef, transf_partial_fundef.
+ caseEq (transf_function f); try congruence.
+ intros tfn TRANSL EQ. inversion EQ; clear EQ; subst tf.
+ inversion WTF as [|f' WTFN]. subst f'.
assert (X: forall link ra,
exists rs' : regset,
- Machabstr.exec_function_body tge tf parent link ra rs0 m rs' (free m2 stk) /\
+ Machabstr.exec_function_body tge tfn parent link ra rs0 m t rs' (free m2 stk) /\
(forall r : mreg,
In (R r) temporaries \/ In (R r) destroyed_at_call -> rs' r = rs2 (R r)) /\
(forall r : mreg,
In r int_callee_save_regs \/ In r float_callee_save_regs -> rs' r = rs0 r)).
intros.
set (sp := Vptr stk Int.zero) in *.
- set (tsp := shift_sp tf sp).
+ set (tsp := shift_sp tfn sp).
set (fe := make_env (function_bounds f)).
- assert (low (init_frame tf) = -fe.(fe_size)).
+ assert (low (init_frame tfn) = -fe.(fe_size)).
simpl low. rewrite (unfold_transf_function _ _ TRANSL).
reflexivity.
- assert (exists fr1, set_slot (init_frame tf) Tint 0 link fr1).
+ assert (exists fr1, set_slot (init_frame tfn) Tint 0 link fr1).
apply set_slot_ok. reflexivity. omega. rewrite H2. generalize (size_pos f).
unfold fe. simpl typesize. omega.
elim H3; intros fr1 SET1; clear H3.
@@ -1526,28 +1550,29 @@ Proof.
apply slot_gi. omega. omega.
simpl typesize. omega. simpl typesize. omega.
inversion H8. symmetry. exact H11.
- generalize (save_callee_save_correct f tf TRANSL
+ generalize (save_callee_save_correct f tfn TRANSL
tsp parent
(transl_code (make_env (function_bounds f)) f.(fn_code))
rs0 fr2 m1 rs AG1 AG2 H5 H6 UNDEF).
intros [fr [EXP AG]].
- generalize (H1 tf rs0 fr parent rs0 TRANSL WTF AG (incl_refl _)).
+ generalize (H1 tfn rs0 fr parent rs0 TRANSL WTFN AG (incl_refl _)).
intros [rs' [fr' [EXB AG']]].
- generalize (restore_callee_save_correct f tf TRANSL tsp parent
+ generalize (restore_callee_save_correct f tfn TRANSL tsp parent
(Mreturn :: transl_code (make_env (function_bounds f)) b)
fr' m2 rs0 rs2 rs' AG').
intros [rs'' [EXX [REGS1 REGS2]]].
exists rs''. split. eapply Machabstr.exec_funct_body.
- rewrite (unfold_transf_function f tf TRANSL); eexact H.
+ rewrite (unfold_transf_function f tfn TRANSL); eexact H.
eexact SET1. eexact SET2.
- replace (Mach.fn_code tf) with
+ replace (Mach.fn_code tfn) with
(transl_body f (make_env (function_bounds f))).
- replace (Vptr stk (Int.repr (- fn_framesize tf))) with tsp.
+ replace (Vptr stk (Int.repr (- fn_framesize tfn))) with tsp.
eapply Machabstr.exec_trans. eexact EXP.
- eapply Machabstr.exec_trans. eexact EXB. eexact EXX.
+ eapply Machabstr.exec_trans. eexact EXB. eexact EXX.
+ reflexivity. traceEq.
unfold tsp, shift_sp, sp. unfold Val.add.
rewrite Int.add_commut. rewrite Int.add_zero. auto.
- rewrite (unfold_transf_function f tf TRANSL). simpl. auto.
+ rewrite (unfold_transf_function f tfn TRANSL). simpl. auto.
split. intros. rewrite REGS2. symmetry; eapply agree_reg; eauto.
apply int_callee_save_not_destroyed; auto.
apply float_callee_save_not_destroyed; auto.
@@ -1563,21 +1588,36 @@ Proof.
rewrite REGS2'. apply REGS2. auto. auto.
rewrite H4. auto.
split; auto.
+
+ (* external function *)
+ simpl in TRANSL. inversion TRANSL; subst tf.
+ inversion WTF. subst ef0. set (sg := ef_sig ef) in *.
+ exists (rs#(loc_result sg) <- res).
+ split. econstructor. eauto.
+ fold sg. rewrite H0. rewrite Conventions.loc_external_arguments_loc_arguments; auto.
+ rewrite list_map_compose. apply list_map_exten; intros. auto.
+ reflexivity.
+ split; intros. rewrite H1.
+ unfold Regmap.set. case (RegEq.eq r (loc_result sg)); intro.
+ rewrite e. rewrite Locmap.gss; auto. rewrite Locmap.gso; auto.
+ red; auto.
+ apply Regmap.gso. red; intro; subst r.
+ elim (Conventions.loc_result_not_callee_save _ H2).
Qed.
End PRESERVATION.
Theorem transl_program_correct:
- forall (p: Linear.program) (tp: Mach.program) (r: val),
+ forall (p: Linear.program) (tp: Mach.program) (t: trace) (r: val),
wt_program p ->
transf_program p = Some tp ->
- Linear.exec_program p r ->
- Machabstr.exec_program tp r.
+ Linear.exec_program p t r ->
+ Machabstr.exec_program tp t r.
Proof.
- intros p tp r WTP TRANSF
+ intros p tp t r WTP TRANSF
[fptr [f [ls' [m [FINDSYMB [FINDFUNC [SIG [EXEC RES]]]]]]]].
- assert (WTF: wt_function f).
- apply (Genv.find_funct_ptr_prop wt_function WTP FINDFUNC).
+ assert (WTF: wt_fundef f).
+ apply (Genv.find_funct_ptr_prop wt_fundef WTP FINDFUNC).
set (ls := Locmap.init Vundef) in *.
set (rs := Regmap.init Vundef).
set (fr := empty_frame).
@@ -1585,26 +1625,25 @@ Proof.
intros; reflexivity.
assert (AG2: forall ofs ty,
6 <= ofs ->
- ofs + typesize ty <= size_arguments f.(fn_sig) ->
+ ofs + typesize ty <= size_arguments (funsig f) ->
get_slot fr ty (Int.signed (Int.repr (4 * ofs))) (ls (S (Outgoing ofs ty)))).
rewrite SIG. unfold size_arguments, sig_args, size_arguments_rec.
intros. generalize (typesize_pos ty).
intro. omegaContradiction.
generalize (function_ptr_translated p tp TRANSF f _ FINDFUNC).
intros [tf [TFIND TRANSL]].
- generalize (transf_function_correct p tp TRANSF WTP _ _ _ _ _ EXEC
+ generalize (transf_function_correct p tp TRANSF WTP _ _ _ _ _ _ EXEC
tf rs fr TRANSL WTF AG1 AG2).
intros [rs' [A [B C]]].
red. exists fptr; exists tf; exists rs'; exists m.
split. rewrite (symbols_preserved p tp TRANSF).
- rewrite (transform_partial_program_main transf_function p TRANSF).
+ rewrite (transform_partial_program_main transf_fundef p TRANSF).
assumption.
split. assumption.
- split. rewrite (unfold_transf_function f tf TRANSL); simpl.
- assumption.
split. replace (Genv.init_mem tp) with (Genv.init_mem p).
- exact A. symmetry. apply Genv.init_mem_transf_partial with transf_function.
+ exact A. symmetry. apply Genv.init_mem_transf_partial with transf_fundef.
exact TRANSF.
- rewrite <- RES. rewrite (unfold_transf_function f tf TRANSL); simpl.
+ rewrite <- RES. replace R3 with (loc_result (funsig f)).
apply B. right. apply loc_result_acceptable.
+ rewrite SIG; reflexivity.
Qed.
diff --git a/backend/Stackingtyping.v b/backend/Stackingtyping.v
index 85d19229..996ada4c 100644
--- a/backend/Stackingtyping.v
+++ b/backend/Stackingtyping.v
@@ -166,6 +166,8 @@ Proof.
(* call *)
apply wt_instrs_cons; auto.
constructor; auto.
+ (* alloc *)
+ apply wt_instrs_cons; auto. constructor.
(* label *)
apply wt_instrs_cons; auto.
constructor.
@@ -208,6 +210,20 @@ Proof.
rewrite H2. eapply size_no_overflow; eauto.
Qed.
+Lemma wt_transf_fundef:
+ forall f tf,
+ Lineartyping.wt_fundef f ->
+ transf_fundef f = Some tf ->
+ wt_fundef tf.
+Proof.
+ intros f tf WT. inversion WT; subst.
+ simpl; intros; inversion H0. constructor; auto.
+ unfold transf_fundef, transf_partial_fundef.
+ caseEq (transf_function f0); try congruence.
+ intros tfn TRANSF EQ. inversion EQ; subst tf.
+ constructor; eapply wt_transf_function; eauto.
+Qed.
+
Lemma program_typing_preserved:
forall (p: Linear.program) (tp: Mach.program),
transf_program p = Some tp ->
@@ -215,8 +231,8 @@ Lemma program_typing_preserved:
Machtyping.wt_program tp.
Proof.
intros; red; intros.
- generalize (transform_partial_program_function transf_function p i f H H1).
+ generalize (transform_partial_program_function transf_fundef p i f H H1).
intros [f0 [IN TRANSF]].
- apply wt_transf_function with f0; auto.
+ apply wt_transf_fundef with f0; auto.
eapply H0; eauto.
Qed.
diff --git a/backend/Tunneling.v b/backend/Tunneling.v
index 9c3e82c4..4fbdc9fd 100644
--- a/backend/Tunneling.v
+++ b/backend/Tunneling.v
@@ -100,6 +100,8 @@ Fixpoint tunnel_block (f: LTL.function) (b: block) {struct b} : block :=
Bstore chunk addr args src (tunnel_block f b)
| Bcall sig ros b =>
Bcall sig ros (tunnel_block f b)
+ | Balloc b =>
+ Balloc (tunnel_block f b)
| Bgoto s =>
Bgoto (branch_target f s)
| Bcond cond args s1 s2 =>
@@ -126,6 +128,9 @@ Definition tunnel_function (f: LTL.function) : LTL.function :=
(fn_entrypoint f)
(wf_tunneled_code f).
+Definition tunnel_fundef (f: LTL.fundef) : LTL.fundef :=
+ transf_fundef tunnel_function f.
+
Definition tunnel_program (p: LTL.program) : LTL.program :=
- transform_program tunnel_function p.
+ transform_program tunnel_fundef p.
diff --git a/backend/Tunnelingproof.v b/backend/Tunnelingproof.v
index 111d1d83..88547e76 100644
--- a/backend/Tunnelingproof.v
+++ b/backend/Tunnelingproof.v
@@ -5,6 +5,7 @@ Require Import Maps.
Require Import AST.
Require Import Values.
Require Import Mem.
+Require Import Events.
Require Import Globalenvs.
Require Import Op.
Require Import Locations.
@@ -82,68 +83,75 @@ Let tge := Genv.globalenv tp.
Lemma functions_translated:
forall v f,
Genv.find_funct ge v = Some f ->
- Genv.find_funct tge v = Some (tunnel_function f).
-Proof (@Genv.find_funct_transf _ _ tunnel_function p).
+ Genv.find_funct tge v = Some (tunnel_fundef f).
+Proof (@Genv.find_funct_transf _ _ tunnel_fundef p).
Lemma function_ptr_translated:
forall v f,
Genv.find_funct_ptr ge v = Some f ->
- Genv.find_funct_ptr tge v = Some (tunnel_function f).
-Proof (@Genv.find_funct_ptr_transf _ _ tunnel_function p).
+ Genv.find_funct_ptr tge v = Some (tunnel_fundef f).
+Proof (@Genv.find_funct_ptr_transf _ _ tunnel_fundef p).
Lemma symbols_preserved:
forall id,
Genv.find_symbol tge id = Genv.find_symbol ge id.
-Proof (@Genv.find_symbol_transf _ _ tunnel_function p).
+Proof (@Genv.find_symbol_transf _ _ tunnel_fundef p).
+
+Lemma sig_preserved:
+ forall f, funsig (tunnel_fundef f) = funsig f.
+Proof.
+ destruct f; reflexivity.
+Qed.
(** These are inversion lemmas, characterizing what happens in the LTL
semantics when executing [Bgoto] instructions or basic blocks. *)
Lemma exec_instrs_Bgoto_inv:
- forall sp b1 ls1 m1 b2 ls2 m2,
- exec_instrs ge sp b1 ls1 m1 b2 ls2 m2 ->
+ forall sp b1 ls1 m1 t b2 ls2 m2,
+ exec_instrs ge sp b1 ls1 m1 t b2 ls2 m2 ->
forall s1,
- b1 = Bgoto s1 -> b2 = b1 /\ ls2 = ls1 /\ m2 = m1.
+ b1 = Bgoto s1 -> t = E0 /\ b2 = b1 /\ ls2 = ls1 /\ m2 = m1.
Proof.
induction 1.
intros. tauto.
- intros. subst b1. inversion H.
- intros. generalize (IHexec_instrs1 s1 H1). intros [A [B C]].
- subst b2; subst rs2; subst m2. eauto.
+ intros. subst b1. inversion H.
+ intros. generalize (IHexec_instrs1 s1 H2). intros [A [B [C D]]].
+ subst t1 b2 rs2 m2.
+ generalize (IHexec_instrs2 s1 H2). intros [A [B [C D]]].
+ subst t2 b3 rs3 m3. intuition. traceEq.
Qed.
Lemma exec_block_Bgoto_inv:
- forall sp s ls1 m1 out ls2 m2,
- exec_block ge sp (Bgoto s) ls1 m1 out ls2 m2 ->
- out = Cont s /\ ls2 = ls1 /\ m2 = m1.
+ forall sp s ls1 m1 t out ls2 m2,
+ exec_block ge sp (Bgoto s) ls1 m1 t out ls2 m2 ->
+ t = E0 /\ out = Cont s /\ ls2 = ls1 /\ m2 = m1.
Proof.
intros. inversion H;
- generalize (exec_instrs_Bgoto_inv _ _ _ _ _ _ _ H0 s (refl_equal _));
- intros [A [B C]].
- split. congruence. tauto.
- discriminate.
- discriminate.
- discriminate.
+ generalize (exec_instrs_Bgoto_inv _ _ _ _ _ _ _ _ H0 s (refl_equal _));
+ intros [A [B [C D]]];
+ try discriminate.
+ intuition congruence.
Qed.
Lemma exec_blocks_Bgoto_inv:
- forall c sp pc1 ls1 m1 out ls2 m2 s,
- exec_blocks ge c sp pc1 ls1 m1 out ls2 m2 ->
+ forall c sp pc1 ls1 m1 t out ls2 m2 s,
+ exec_blocks ge c sp pc1 ls1 m1 t out ls2 m2 ->
c!pc1 = Some (Bgoto s) ->
- (out = Cont pc1 /\ ls2 = ls1 /\ m2 = m1)
- \/ exec_blocks ge c sp s ls1 m1 out ls2 m2.
+ (t = E0 /\ out = Cont pc1 /\ ls2 = ls1 /\ m2 = m1)
+ \/ exec_blocks ge c sp s ls1 m1 t out ls2 m2.
Proof.
induction 1; intros.
left; tauto.
assert (b = Bgoto s). congruence. subst b.
- generalize (exec_block_Bgoto_inv _ _ _ _ _ _ _ H0).
- intros [A [B C]]. subst out; subst rs'; subst m'.
+ generalize (exec_block_Bgoto_inv _ _ _ _ _ _ _ _ H0).
+ intros [A [B [C D]]]. subst t out rs' m'.
right. apply exec_blocks_refl.
- elim (IHexec_blocks1 H1).
- intros [A [B C]].
- assert (pc2 = pc1). congruence. subst rs2; subst m2; subst pc2.
- apply IHexec_blocks2; auto.
- intro. right. apply exec_blocks_trans with pc2 rs2 m2; auto.
+ elim (IHexec_blocks1 H2).
+ intros [A [B [C D]]].
+ assert (pc2 = pc1). congruence. subst t1 rs2 m2 pc2.
+ replace t with t2. apply IHexec_blocks2; auto. traceEq.
+ intro. right.
+ eapply exec_blocks_trans; eauto.
Qed.
(** The following [exec_*_prop] predicates state the correctness
@@ -163,43 +171,43 @@ Definition tunnel_outcome (f: function) (out: outcome) :=
end.
Definition exec_instr_prop
- (sp: val) (b1: block) (ls1: locset) (m1: mem)
+ (sp: val) (b1: block) (ls1: locset) (m1: mem) (t: trace)
(b2: block) (ls2: locset) (m2: mem) : Prop :=
forall f,
exec_instr tge sp (tunnel_block f b1) ls1 m1
- (tunnel_block f b2) ls2 m2.
+ t (tunnel_block f b2) ls2 m2.
Definition exec_instrs_prop
- (sp: val) (b1: block) (ls1: locset) (m1: mem)
+ (sp: val) (b1: block) (ls1: locset) (m1: mem) (t: trace)
(b2: block) (ls2: locset) (m2: mem) : Prop :=
forall f,
exec_instrs tge sp (tunnel_block f b1) ls1 m1
- (tunnel_block f b2) ls2 m2.
+ t (tunnel_block f b2) ls2 m2.
Definition exec_block_prop
- (sp: val) (b1: block) (ls1: locset) (m1: mem)
+ (sp: val) (b1: block) (ls1: locset) (m1: mem) (t: trace)
(out: outcome) (ls2: locset) (m2: mem) : Prop :=
forall f,
exec_block tge sp (tunnel_block f b1) ls1 m1
- (tunnel_outcome f out) ls2 m2.
+ t (tunnel_outcome f out) ls2 m2.
Definition tunneled_code (f: function) :=
PTree.map (fun pc b => tunnel_block f b) (fn_code f).
Definition exec_blocks_prop
(c: code) (sp: val)
- (pc: node) (ls1: locset) (m1: mem)
+ (pc: node) (ls1: locset) (m1: mem) (t: trace)
(out: outcome) (ls2: locset) (m2: mem) : Prop :=
forall f,
f.(fn_code) = c ->
exec_blocks tge (tunneled_code f) sp
(branch_target f pc) ls1 m1
- (tunnel_outcome f out) ls2 m2.
+ t (tunnel_outcome f out) ls2 m2.
Definition exec_function_prop
- (f: function) (ls1: locset) (m1: mem)
- (ls2: locset) (m2: mem) : Prop :=
- exec_function tge (tunnel_function f) ls1 m1 ls2 m2.
+ (f: fundef) (ls1: locset) (m1: mem) (t: trace)
+ (ls2: locset) (m2: mem) : Prop :=
+ exec_function tge (tunnel_fundef f) ls1 m1 t ls2 m2.
Scheme exec_instr_ind5 := Minimality for LTL.exec_instr Sort Prop
with exec_instrs_ind5 := Minimality for LTL.exec_instrs Sort Prop
@@ -212,9 +220,9 @@ Scheme exec_instr_ind5 := Minimality for LTL.exec_instr Sort Prop
using the [exec_*_prop] predicates above as induction hypotheses. *)
Lemma tunnel_function_correct:
- forall f ls1 m1 ls2 m2,
- exec_function ge f ls1 m1 ls2 m2 ->
- exec_function_prop f ls1 m1 ls2 m2.
+ forall f ls1 m1 t ls2 m2,
+ exec_function ge f ls1 m1 t ls2 m2 ->
+ exec_function_prop f ls1 m1 t ls2 m2.
Proof.
apply (exec_function_ind5 ge
exec_instr_prop
@@ -239,19 +247,21 @@ Proof.
apply eval_addressing_preserved. exact symbols_preserved.
auto.
(* call *)
- apply exec_Bcall with (tunnel_function f).
+ apply exec_Bcall with (tunnel_fundef f).
generalize H; unfold find_function; destruct ros.
intro. apply functions_translated; auto.
rewrite symbols_preserved. destruct (Genv.find_symbol ge i).
intro. apply function_ptr_translated; auto. congruence.
- rewrite H0; reflexivity.
+ generalize (sig_preserved f). congruence.
apply H2.
+ (* alloc *)
+ eapply exec_Balloc; eauto.
(* instr_refl *)
apply exec_refl.
(* instr_one *)
apply exec_one. apply H0.
(* instr_trans *)
- apply exec_trans with (tunnel_block f b2) rs2 m2; auto.
+ apply exec_trans with t1 (tunnel_block f b2) rs2 m2 t2; auto.
(* goto *)
apply exec_Bgoto. red in H0. simpl in H0. apply H0.
(* cond, true *)
@@ -270,13 +280,13 @@ Proof.
reflexivity. apply H1.
intros [pc' [ATpc BTS]].
assert (b = Bgoto pc'). congruence. subst b.
- generalize (exec_block_Bgoto_inv _ _ _ _ _ _ _ H0).
- intros [A [B C]]. subst out; subst rs'; subst m'.
+ generalize (exec_block_Bgoto_inv _ _ _ _ _ _ _ _ H0).
+ intros [A [B [C D]]]. subst t out rs' m'.
simpl. rewrite BTS. apply exec_blocks_refl.
(* blocks_trans *)
- apply exec_blocks_trans with (branch_target f pc2) rs2 m2.
- exact (H0 f H3). exact (H2 f H3).
- (* function *)
+ apply exec_blocks_trans with t1 (branch_target f pc2) rs2 m2 t2.
+ exact (H0 f H4). exact (H2 f H4). auto.
+ (* internal function *)
econstructor. eexact H.
change (fn_code (tunnel_function f)) with (tunneled_code f).
simpl.
@@ -284,28 +294,30 @@ Proof.
intro BT. rewrite <- BT. exact (H1 f (refl_equal _)).
intros [pc [ATpc BT]].
apply exec_blocks_trans with
- (branch_target f (fn_entrypoint f)) (call_regs rs) m1.
+ E0 (branch_target f (fn_entrypoint f)) (call_regs rs) m1 t.
eapply exec_blocks_one.
unfold tunneled_code. rewrite PTree.gmap. rewrite ATpc.
simpl. reflexivity.
apply exec_Bgoto. rewrite BT. apply exec_refl.
- exact (H1 f (refl_equal _)).
+ exact (H1 f (refl_equal _)). traceEq.
+ (* external function *)
+ econstructor; eauto.
Qed.
End PRESERVATION.
Theorem transf_program_correct:
- forall (p: program) (r: val),
- exec_program p r ->
- exec_program (tunnel_program p) r.
+ forall (p: program) (t: trace) (r: val),
+ exec_program p t r ->
+ exec_program (tunnel_program p) t r.
Proof.
- intros p r [b [f [ls [m [FIND1 [FIND2 [SIG [EX RES]]]]]]]].
- red. exists b; exists (tunnel_function f); exists ls; exists m.
+ intros p t r [b [f [ls [m [FIND1 [FIND2 [SIG [EX RES]]]]]]]].
+ red. exists b; exists (tunnel_fundef f); exists ls; exists m.
split. change (prog_main (tunnel_program p)) with (prog_main p).
rewrite <- FIND1. apply symbols_preserved.
split. apply function_ptr_translated. assumption.
- split. rewrite <- SIG. reflexivity.
+ split. generalize (sig_preserved f). congruence.
split. apply tunnel_function_correct.
unfold tunnel_program. rewrite Genv.init_mem_transf. auto.
- exact RES.
+ rewrite sig_preserved. exact RES.
Qed.
diff --git a/backend/Tunnelingtyping.v b/backend/Tunnelingtyping.v
index 29b74f12..6281afa1 100644
--- a/backend/Tunnelingtyping.v
+++ b/backend/Tunnelingtyping.v
@@ -33,12 +33,19 @@ Proof.
intros; discriminate.
Qed.
+Lemma wt_tunnel_fundef:
+ forall f, wt_fundef f -> wt_fundef (tunnel_fundef f).
+Proof.
+ intros. inversion H; simpl. constructor; auto.
+ constructor. apply wt_tunnel_function; auto.
+Qed.
+
Lemma program_typing_preserved:
forall (p: LTL.program),
wt_program p -> wt_program (tunnel_program p).
Proof.
intros; red; intros.
- generalize (transform_program_function tunnel_function p i f H0).
+ generalize (transform_program_function tunnel_fundef p i f H0).
intros [f0 [IN TRANSF]].
- subst f. apply wt_tunnel_function. eauto.
+ subst f. apply wt_tunnel_fundef. eauto.
Qed.