aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2019-03-29 15:01:33 +0100
committerCyril SIX <cyril.six@kalray.eu>2019-03-29 15:01:33 +0100
commit5f38e0401747a59a821852fdefb7b588a818bdcf (patch)
tree345f4a16f2d450b05602c0be71fa13a54ce90d4b
parente9a05f4ca3157a88a03f71ab31ef59bd96650142 (diff)
parent7633cb38e0440160acf3f60f7795a19224850eec (diff)
downloadcompcert-kvx-5f38e0401747a59a821852fdefb7b588a818bdcf.tar.gz
compcert-kvx-5f38e0401747a59a821852fdefb7b588a818bdcf.zip
Merge remote-tracking branch 'origin/mppa_postpass' into mppa-jumptable
Conflicts: mppa_k1c/Asmblock.v mppa_k1c/Asmblockdeps.v
-rwxr-xr-xconfigure2
-rw-r--r--mppa_k1c/Asm.v6
-rw-r--r--mppa_k1c/Asmblock.v21
-rw-r--r--mppa_k1c/Asmblockdeps.v569
-rw-r--r--mppa_k1c/Asmvliw.v346
-rw-r--r--mppa_k1c/PostpassScheduling.v7
-rw-r--r--mppa_k1c/PostpassSchedulingproof.v176
-rw-r--r--mppa_k1c/abstractbb/Parallelizability.v18
8 files changed, 1121 insertions, 24 deletions
diff --git a/configure b/configure
index a3a2ea8c..5852205e 100755
--- a/configure
+++ b/configure
@@ -804,7 +804,7 @@ if [ "$arch" = "mppa_k1c" ]; then
cat >> Makefile.config <<EOF
ARCHDIRS=$arch $arch/lib $arch/abstractbb $arch/abstractbb/Impure
BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\
- Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v\\
+ Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v\\
ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\
Asmblockdeps.v\\
AbstractBasicBlocksDef.v DepTreeTheory.v ImpDep.v Parallelizability.v\\
diff --git a/mppa_k1c/Asm.v b/mppa_k1c/Asm.v
index 421f3d2c..2d708b79 100644
--- a/mppa_k1c/Asm.v
+++ b/mppa_k1c/Asm.v
@@ -30,7 +30,7 @@ Require Import Smallstep.
Require Import Locations.
Require Stacklayout.
Require Import Conventions.
-Require Import Asmblock.
+Require Import Asmvliw.
Require Import Linking.
Require Import Errors.
@@ -477,7 +477,7 @@ Definition program_proj (p: program) : Asmblock.program :=
End RELSEM.
-Definition semantics (p: program) := Asmblock.semantics (program_proj p).
+Definition semantics (p: program) := Asmvliw.semantics (program_proj p).
(** Determinacy of the [Asm] semantics. *)
@@ -608,7 +608,7 @@ Proof (Genv.senv_match TRANSF).
Theorem transf_program_correct:
- forward_simulation (Asmblock.semantics prog) (semantics tprog).
+ forward_simulation (Asmvliw.semantics prog) (semantics tprog).
Proof.
pose proof (match_program_transf prog tprog TRANSF) as TR.
subst. unfold semantics. rewrite transf_program_proj.
diff --git a/mppa_k1c/Asmblock.v b/mppa_k1c/Asmblock.v
index d630e767..b4cf57ae 100644
--- a/mppa_k1c/Asmblock.v
+++ b/mppa_k1c/Asmblock.v
@@ -178,7 +178,8 @@ Inductive ex_instruction : Type :=
| Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *) *)
| Pbuiltin: external_function -> list (builtin_arg preg)
- -> builtin_res preg -> ex_instruction (**r built-in function (pseudo) *).
+ -> builtin_res preg -> ex_instruction (**r built-in function (pseudo) *)
+.
(** FIXME: comment not up to date !
@@ -842,10 +843,10 @@ Section RELSEM.
set and memory state after execution of the instruction at [rs#PC],
or [Stuck] if the processor is stuck. *)
-Inductive outcome {rgset}: Type :=
- | Next (rs:rgset) (m:mem)
+Inductive outcome: Type :=
+ | Next (rs:regset) (m:mem)
| Stuck.
-Arguments outcome: clear implicits.
+(* Arguments outcome: clear implicits. *)
(** ** Arithmetic Expressions (including comparisons) *)
@@ -1307,7 +1308,7 @@ Definition store_chunk n :=
(** * basic instructions *)
-Definition exec_basic_instr (bi: basic) (rs: regset) (m: mem) : outcome regset :=
+Definition exec_basic_instr (bi: basic) (rs: regset) (m: mem) : outcome :=
match bi with
| PArith ai => Next (exec_arith_instr ai rs) m
@@ -1349,7 +1350,7 @@ Definition exec_basic_instr (bi: basic) (rs: regset) (m: mem) : outcome regset :
| Pnop => Next rs m
end.
-Fixpoint exec_body (body: list basic) (rs: regset) (m: mem): outcome regset :=
+Fixpoint exec_body (body: list basic) (rs: regset) (m: mem): outcome :=
match body with
| nil => Next rs m
| bi::body' =>
@@ -1409,7 +1410,7 @@ Fixpoint label_pos (lbl: label) (pos: Z) (lb: bblocks) {struct lb} : option Z :=
| b :: lb' => if is_label lbl b then Some pos else label_pos lbl (pos + (size b)) lb'
end.
-Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) : outcome regset :=
+Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) : outcome :=
match label_pos lbl 0 (fn_blocks f) with
| None => Stuck
| Some pos =>
@@ -1424,7 +1425,7 @@ Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) : outcome
Warning: in m PC is assumed to be already pointing on the next instruction !
*)
-Definition eval_branch (f: function) (l: label) (rs: regset) (m: mem) (res: option bool) : outcome regset :=
+Definition eval_branch (f: function) (l: label) (rs: regset) (m: mem) (res: option bool) : outcome :=
match res with
| Some true => goto_label f l rs m
| Some false => Next rs m
@@ -1448,7 +1449,7 @@ Definition eval_branch (f: function) (l: label) (rs: regset) (m: mem) (res: opti
we generate cannot use those registers to hold values that must
survive the execution of the pseudo-instruction. *)
-Definition exec_control (f: function) (oc: option control) (rs: regset) (m: mem) : outcome regset :=
+Definition exec_control (f: function) (oc: option control) (rs: regset) (m: mem) : outcome :=
match oc with
| Some ic =>
(** Get/Set system registers *)
@@ -1498,7 +1499,7 @@ Definition exec_control (f: function) (oc: option control) (rs: regset) (m: mem)
| None => Next rs m
end.
-Definition exec_bblock (f: function) (b: bblock) (rs0: regset) (m: mem) : outcome regset :=
+Definition exec_bblock (f: function) (b: bblock) (rs0: regset) (m: mem) : outcome :=
match exec_body (body b) rs0 m with
| Next rs' m' =>
let rs1 := nextblock b rs' in exec_control f (exit b) rs1 m'
diff --git a/mppa_k1c/Asmblockdeps.v b/mppa_k1c/Asmblockdeps.v
index 845a26ed..7f03c66a 100644
--- a/mppa_k1c/Asmblockdeps.v
+++ b/mppa_k1c/Asmblockdeps.v
@@ -609,10 +609,10 @@ Fixpoint trans_body (b: list basic) : list L.macro :=
| b :: lb => (trans_basic b) :: (trans_body lb)
end.
-Definition trans_pcincr (sz: Z) (k: L.macro) := [(#PC, Op (Control (OIncremPC sz)) (Name (#PC) @ Enil)) :: k].
+Definition trans_pcincr (sz: Z) (k: L.macro) := (#PC, Op (Control (OIncremPC sz)) (Name (#PC) @ Enil)) :: k.
Definition trans_block (b: Asmblock.bblock) : L.bblock :=
- trans_body (body b) ++ trans_pcincr (size b) (trans_exit (exit b)).
+ trans_body (body b) ++ (trans_pcincr (size b) (trans_exit (exit b)) :: nil).
Theorem trans_block_noheader_inv: forall bb, trans_block (no_header bb) = trans_block bb.
Proof.
@@ -870,7 +870,7 @@ Lemma forward_simu_control:
exec_control ge fn ex (nextblock b rs) m = Next rs2 m2 ->
match_states (State rs m) s ->
exists s',
- exec Ge (trans_pcincr (size b) (trans_exit ex)) s = Some s'
+ exec Ge (trans_pcincr (size b) (trans_exit ex) :: nil) s = Some s'
/\ match_states (State rs2 m2) s'.
Proof.
intros. destruct ex.
@@ -1066,7 +1066,7 @@ Lemma exec_trans_pcincr_exec:
forall rs m s b,
match_states (State rs m) s ->
exists s',
- exec Ge (trans_pcincr (size b) (trans_exit (exit b))) s = exec Ge [trans_exit (exit b)] s'
+ exec Ge (trans_pcincr (size b) (trans_exit (exit b)) :: nil) s = exec Ge [trans_exit (exit b)] s'
/\ match_states (State (nextblock b rs) m) s'.
Proof.
intros.
@@ -1607,3 +1607,564 @@ Module PChk := ParallelChecks L PosResourceSet.
Definition bblock_para_check (p: Asmblock.bblock) : bool :=
PChk.is_parallelizable (trans_block p).
+
+Require Import Asmvliw Permutation.
+Import PChk.
+
+Section SECT_PAR.
+
+Ltac Simplif :=
+ ((rewrite nextblock_inv by eauto with asmgen)
+ || (rewrite nextblock_inv1 by eauto with asmgen)
+ || (rewrite Pregmap.gss)
+ || (rewrite nextblock_pc)
+ || (rewrite Pregmap.gso by eauto with asmgen)
+ || (rewrite assign_diff by (auto; try discriminate; try (apply ppos_discr; try discriminate; congruence); try (apply ppos_pmem_discr);
+ try (apply not_eq_sym; apply ppos_discr; try discriminate; congruence); try (apply not_eq_sym; apply ppos_pmem_discr); auto))
+ || (rewrite assign_eq)
+ ); auto with asmgen.
+
+Ltac Simpl := repeat Simplif.
+
+Arguments Pos.add: simpl never.
+Arguments ppos: simpl never.
+
+Variable Ge: genv.
+
+Lemma trans_arith_par_correct ge fn rsr mr sr rsw mw sw rsw' mw' i:
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_arith_instr ge i rsr rsw = rsw' -> mw = mw' ->
+ exists sw',
+ macro_prun Ge (trans_arith i) sw sr sr = Some sw'
+ /\ match_states (State rsw' mw') sw'.
+Proof.
+ intros GENV MSR MSW PARARITH MWEQ. subst. inv MSR. inv MSW.
+ unfold parexec_arith_instr. destruct i.
+(* Ploadsymbol *)
+ - destruct i. eexists; split; [| split].
+ * simpl. reflexivity.
+ * Simpl.
+ * simpl. intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRR *)
+ - eexists; split; [| split].
+ * simpl. rewrite (H0 rs). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRI32 *)
+ - eexists; split; [|split].
+ * simpl. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRI64 *)
+ - eexists; split; [|split].
+ * simpl. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRF32 *)
+ - eexists; split; [|split].
+ * simpl. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRF64 *)
+ - eexists; split; [|split].
+ * simpl. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRRR *)
+ - eexists; split; [|split].
+ * simpl. rewrite (H0 rs1). rewrite (H0 rs2). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRRI32 *)
+ - eexists; split; [|split].
+ * simpl. rewrite (H0 rs). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithRRI64 *)
+ - eexists; split; [|split].
+ * simpl. rewrite (H0 rs). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithARRR *)
+ - eexists; split; [|split].
+ * simpl. rewrite (H0 rd). rewrite (H0 rs1). rewrite (H0 rs2). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithARRI32 *)
+ - eexists; split; [|split].
+ * simpl. rewrite (H0 rd). rewrite (H0 rs). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* PArithARRI64 *)
+ - eexists; split; [|split].
+ * simpl. rewrite (H0 rd). rewrite (H0 rs). reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+Qed.
+
+Theorem forward_simu_par_wio_basic ge fn rsr rsw mr mw sr sw bi rsw' mw':
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_basic_instr ge bi rsr rsw mr mw = Next rsw' mw' ->
+ exists sw',
+ macro_prun Ge (trans_basic bi) sw sr sr = Some sw'
+ /\ match_states (State rsw' mw') sw'.
+Proof.
+ intros GENV MSR MSW H.
+ destruct bi.
+(* Arith *)
+ - simpl in H. inversion H. subst rsw' mw'. simpl macro_prun. eapply trans_arith_par_correct; eauto.
+(* Load *)
+ - simpl in H. destruct i.
+ unfold parexec_load in H; destruct (eval_offset _ _) eqn:EVALOFF; try discriminate;
+ destruct (Mem.loadv _ _ _) eqn:MEML; try discriminate; inv H. inv MSR; inv MSW;
+ eexists; split; try split;
+ [ simpl; rewrite EVALOFF; rewrite H; pose (H0 ra); simpl in e; rewrite e; rewrite MEML; reflexivity|
+ Simpl|
+ intros rr; destruct rr; Simpl;
+ destruct (ireg_eq g rd); [
+ subst; Simpl|
+ Simpl; rewrite assign_diff; pose (H1 g); simpl in e; try assumption; Simpl; unfold ppos; apply not_eq_ireg_to_pos; assumption]].
+(* Store *)
+ - simpl in H. destruct i.
+ unfold parexec_store in H; destruct (eval_offset _ _) eqn:EVALOFF; try discriminate.
+ destruct (Mem.storev _ _ _ _) eqn:MEML; try discriminate. inv H; inv MSR; inv MSW.
+ eexists; split; try split.
+ * simpl. rewrite EVALOFF. rewrite H. rewrite (H0 ra). rewrite (H0 rs). rewrite MEML. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+(* Allocframe *)
+ - simpl in H. destruct (Mem.alloc _ _ _) eqn:MEMAL. destruct (Mem.store _ _ _ _) eqn:MEMS; try discriminate.
+ inv H. inv MSR. inv MSW. eexists. split; try split.
+ * simpl. Simpl. rewrite (H0 GPR12). rewrite H. rewrite MEMAL. rewrite MEMS. Simpl.
+ rewrite H. rewrite MEMAL. rewrite MEMS. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g GPR32); [| destruct (ireg_eq g GPR12); [| destruct (ireg_eq g GPR17)]]; subst; Simpl.
+(* Freeframe *)
+ - simpl in H. destruct (Mem.loadv _ _ _) eqn:MLOAD; try discriminate. destruct (rsr GPR12) eqn:SPeq; try discriminate.
+ destruct (Mem.free _ _ _ _) eqn:MFREE; try discriminate. inv H. inv MSR. inv MSW.
+ eexists. split; try split.
+ * simpl. rewrite (H0 GPR12). rewrite H. rewrite SPeq. rewrite MLOAD. rewrite MFREE.
+ Simpl. rewrite (H0 GPR12). rewrite SPeq. rewrite MLOAD. rewrite MFREE. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; Simpl. destruct (ireg_eq g GPR32); [| destruct (ireg_eq g GPR12); [| destruct (ireg_eq g GPR17)]]; subst; Simpl.
+(* Pget *)
+ - simpl in H. destruct rs eqn:rseq; try discriminate. inv H. inv MSR. inv MSW.
+ eexists. split; try split. Simpl. intros rr; destruct rr; Simpl.
+ destruct (ireg_eq g rd); subst; Simpl.
+(* Pset *)
+ - simpl in H. destruct rd eqn:rdeq; try discriminate. inv H. inv MSR; inv MSW.
+ eexists. split; try split. Simpl. intros rr; destruct rr; Simpl.
+(* Pnop *)
+ - simpl in H. inv H. inv MSR. inv MSW. eexists. split; try split. assumption. assumption.
+Qed.
+
+Theorem forward_simu_par_wio_basic_Stuck ge fn rsr rsw mr mw sr sw bi:
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_basic_instr ge bi rsr rsw mr mw = Stuck ->
+ macro_prun Ge (trans_basic bi) sw sr sr = None.
+Proof.
+ intros GENV MSR MSW H0. inv MSR; inv MSW.
+ unfold parexec_basic_instr in H0. destruct bi; try discriminate.
+(* PLoad *)
+ - destruct i; destruct i.
+ all: simpl; rewrite H; rewrite (H1 ra); unfold parexec_load in H0;
+ destruct (eval_offset _ _); auto; destruct (Mem.loadv _ _ _); auto; discriminate.
+(* PStore *)
+ - destruct i; destruct i;
+ simpl; rewrite H; rewrite (H1 ra); rewrite (H1 rs);
+ unfold parexec_store in H0; destruct (eval_offset _ _); auto; destruct (Mem.storev _ _ _); auto; discriminate.
+(* Pallocframe *)
+ - simpl. Simpl. rewrite (H1 SP). rewrite H. destruct (Mem.alloc _ _ _). simpl in H0.
+ destruct (Mem.store _ _ _ _); try discriminate. reflexivity.
+(* Pfreeframe *)
+ - simpl. Simpl. rewrite (H1 SP). rewrite H.
+ destruct (Mem.loadv _ _ _); auto. destruct (rsr GPR12); auto. destruct (Mem.free _ _ _ _); auto.
+ discriminate.
+(* Pget *)
+ - simpl. destruct rs; subst; try discriminate.
+ all: simpl; auto.
+ - simpl. destruct rd; subst; try discriminate.
+ all: simpl; auto.
+Qed.
+
+Theorem forward_simu_par_body:
+ forall bdy ge fn rsr mr sr rsw mw sw rs' m',
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_wio_body ge bdy rsr rsw mr mw = Next rs' m' ->
+ exists s',
+ prun_iw Ge (trans_body bdy) sw sr = Some s'
+ /\ match_states (State rs' m') s'.
+Proof.
+ induction bdy.
+ - intros until m'. intros GENV MSR MSW WIO.
+ simpl in WIO. inv WIO. inv MSR. inv MSW.
+ eexists. split; [| split].
+ * simpl. reflexivity.
+ * assumption.
+ * assumption.
+ - intros until m'. intros GENV MSR MSW WIO.
+ simpl in WIO. destruct (parexec_basic_instr _ _ _ _ _ _) eqn:PARBASIC; try discriminate.
+ exploit forward_simu_par_wio_basic. 4: eapply PARBASIC. all: eauto.
+ intros (sw' & MPRUN & MS'). simpl prun_iw. rewrite MPRUN.
+ eapply IHbdy; eauto.
+Qed.
+
+Theorem forward_simu_par_control ge fn rsr rsw mr mw sr sw sz rs' ex m':
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_control ge fn ex (par_nextblock (Ptrofs.repr sz) rsr) (par_nextblock (Ptrofs.repr sz) rsw) mw = Next rs' m' ->
+ exists s',
+ macro_prun Ge (trans_pcincr sz (trans_exit ex)) sw sr sr = Some s'
+ /\ match_states (State rs' m') s'.
+Proof.
+ intros GENV MSR MSW H0.
+ simpl in *.
+ destruct ex.
+ - destruct c; destruct i; try discriminate.
+ all: try (inv H0; inv MSR; inv MSW; eexists; split; [| split]; [simpl; rewrite (H0 PC); reflexivity | Simpl | intros rr; destruct rr; unfold par_nextblock; Simpl]).
+
+ (* Pj_l *)
+ + simpl in H0. unfold par_goto_label in H0. destruct (label_pos _ _ _) eqn:LPOS; try discriminate. destruct (par_nextblock _ _ _) eqn:NB; try discriminate. inv H0.
+ inv MSR; inv MSW.
+ eexists; split; try split.
+ * simpl. rewrite (H0 PC). unfold goto_label_deps. rewrite LPOS. Simpl.
+ unfold par_nextblock in NB. rewrite Pregmap.gss in NB. rewrite NB. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ (* Pcb *)
+ + simpl in H0. destruct (cmp_for_btest _) eqn:CFB. destruct o; try discriminate. destruct i.
+ ++ unfold par_eval_branch in H0. destruct (Val.cmp_bool _ _ _) eqn:VALCMP; try discriminate. destruct b.
+ +++ unfold par_goto_label in H0. destruct (label_pos _ _ _) eqn:LPOS; try discriminate. destruct (par_nextblock _ _ PC) eqn:NB; try discriminate.
+ inv H0. inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ unfold goto_label_deps. rewrite LPOS.
+ unfold par_nextblock in NB. rewrite Pregmap.gss in NB. rewrite NB. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ +++ inv H0. inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ ++ unfold par_eval_branch in H0. destruct (Val.cmpl_bool _ _ _) eqn:VALCMP; try discriminate. destruct b.
+ +++ unfold par_goto_label in H0. destruct (label_pos _ _ _) eqn:LPOS; try discriminate. destruct (par_nextblock _ _ PC) eqn:NB; try discriminate.
+ inv H0; inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ unfold goto_label_deps. rewrite LPOS.
+ unfold par_nextblock in NB. rewrite Pregmap.gss in NB. rewrite NB. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ +++ inv H0. inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ (* Pcbu *)
+ + simpl in H0. destruct (cmpu_for_btest _) eqn:CFB. destruct o; try discriminate. destruct i.
+ ++ unfold par_eval_branch in H0. destruct (Val_cmpu_bool _ _ _) eqn:VALCMP; try discriminate. destruct b.
+ +++ unfold par_goto_label in H0. destruct (label_pos _ _ _) eqn:LPOS; try discriminate. destruct (par_nextblock _ _ PC) eqn:NB; try discriminate.
+ inv H0. inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ unfold goto_label_deps. rewrite LPOS.
+ unfold par_nextblock in NB. rewrite Pregmap.gss in NB. rewrite NB. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ +++ inv H0. inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ ++ unfold par_eval_branch in H0. destruct (Val_cmplu_bool _ _ _) eqn:VALCMP; try discriminate. destruct b.
+ +++ unfold par_goto_label in H0. destruct (label_pos _ _ _) eqn:LPOS; try discriminate. destruct (par_nextblock _ _ PC) eqn:NB; try discriminate.
+ inv H0; inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ unfold goto_label_deps. rewrite LPOS.
+ unfold par_nextblock in NB. rewrite Pregmap.gss in NB. rewrite NB. reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ +++ inv H0. inv MSR; inv MSW. eexists; split; try split.
+ * simpl. rewrite (H0 PC).
+ rewrite CFB. Simpl. rewrite (H0 r).
+ unfold eval_branch_deps. unfold par_nextblock in VALCMP. rewrite Pregmap.gso in VALCMP; try discriminate. rewrite VALCMP.
+ reflexivity.
+ * Simpl.
+ * intros rr; destruct rr; unfold par_nextblock; Simpl.
+ - simpl in *. inv MSR. inv MSW. inv H0. eexists. split.
+ rewrite (H1 PC). simpl. reflexivity.
+ split. Simpl.
+ intros rr. destruct rr; unfold par_nextblock; Simpl.
+Qed.
+
+Lemma forward_simu_par_control_Stuck ge fn rsr rsw mr mw sr sw sz ex:
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_control ge fn ex (par_nextblock (Ptrofs.repr sz) rsr) (par_nextblock (Ptrofs.repr sz) rsw) mw = Stuck ->
+ macro_prun Ge (trans_pcincr sz (trans_exit ex)) sw sr sr = None.
+Proof.
+ intros GENV MSR MSW H0. inv MSR; inv MSW. destruct ex as [ctl|]; try discriminate.
+ destruct ctl; destruct i; try discriminate; try (simpl; reflexivity).
+(* Pbuiltin *)
+ - simpl in *. rewrite (H1 PC). reflexivity.
+(* Pj_l *)
+ - simpl in *. rewrite (H1 PC). unfold goto_label_deps. unfold par_goto_label in H0.
+ destruct (label_pos _ _ _); auto. simpl in *. unfold par_nextblock in H0. rewrite Pregmap.gss in H0.
+ destruct (Val.offset_ptr _ _); try discriminate; auto.
+(* Pcb *)
+ - simpl in *. destruct (cmp_for_btest bt). destruct i.
+ -- destruct o.
+ + unfold par_eval_branch in H0; unfold eval_branch_deps.
+ rewrite (H1 PC). Simpl. rewrite (H1 r). unfold par_nextblock in H0. rewrite Pregmap.gso in H0; try discriminate.
+ destruct (Val.cmp_bool _ _ _); auto. destruct b; try discriminate. unfold goto_label_deps; unfold par_goto_label in H0.
+ destruct (label_pos _ _ _); auto. rewrite Pregmap.gss in H0. destruct (Val.offset_ptr _ _); auto. discriminate.
+ + rewrite (H1 PC). Simpl. rewrite (H1 r). reflexivity.
+ -- destruct o.
+ + unfold par_eval_branch in H0; unfold eval_branch_deps.
+ rewrite (H1 PC). Simpl. rewrite (H1 r). unfold par_nextblock in H0. rewrite Pregmap.gso in H0; try discriminate.
+ destruct (Val.cmpl_bool _ _ _); auto. destruct b; try discriminate. unfold goto_label_deps; unfold par_goto_label in H0.
+ destruct (label_pos _ _ _); auto. rewrite Pregmap.gss in H0. destruct (Val.offset_ptr _ _); auto. discriminate.
+ + rewrite (H1 PC). Simpl. rewrite (H1 r). reflexivity.
+(* Pcbu *)
+ - simpl in *. destruct (cmpu_for_btest bt). destruct i.
+ -- destruct o.
+ + unfold par_eval_branch in H0; unfold eval_branch_deps.
+ rewrite (H1 PC). Simpl. rewrite (H1 r). unfold par_nextblock in H0. rewrite Pregmap.gso in H0; try discriminate.
+ destruct (Val_cmpu_bool _ _ _); auto. destruct b; try discriminate. unfold goto_label_deps; unfold par_goto_label in H0.
+ destruct (label_pos _ _ _); auto. rewrite Pregmap.gss in H0. destruct (Val.offset_ptr _ _); auto. discriminate.
+ + rewrite (H1 PC). Simpl. rewrite (H1 r). reflexivity.
+ -- destruct o.
+ + unfold par_eval_branch in H0; unfold eval_branch_deps.
+ rewrite (H1 PC). Simpl. rewrite (H1 r). unfold par_nextblock in H0. rewrite Pregmap.gso in H0; try discriminate.
+ destruct (Val_cmplu_bool _ _ _); auto. destruct b; try discriminate. unfold goto_label_deps; unfold par_goto_label in H0.
+ destruct (label_pos _ _ _); auto. rewrite Pregmap.gss in H0. destruct (Val.offset_ptr _ _); auto. discriminate.
+ + rewrite (H1 PC). Simpl. rewrite (H1 r). reflexivity.
+Qed.
+
+Definition trans_block_aux bdy sz ex := (trans_body bdy) ++ (trans_pcincr sz (trans_exit ex) :: nil).
+
+Theorem forward_simu_par_wio_bblock_aux ge fn rsr mr sr rsw mw sw bdy ex sz rs' m':
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_wio_bblock_aux ge fn bdy ex (Ptrofs.repr sz) rsr rsw mr mw = Next rs' m' ->
+ exists s',
+ prun_iw Ge (trans_block_aux bdy sz ex) sw sr = Some s'
+ /\ match_states (State rs' m') s'.
+Proof.
+ intros. unfold parexec_wio_bblock_aux in H2. unfold trans_block_aux.
+ destruct (parexec_wio_body _ _ _ _ _ _) eqn:WIOB; try discriminate.
+ exploit forward_simu_par_body. 4: eapply WIOB. all: eauto.
+ intros (s' & RUNB & MS').
+ exploit forward_simu_par_control. 4: eapply H2. all: eauto.
+ intros (s'' & RUNCTL & MS'').
+ eexists. split.
+ erewrite prun_iw_app_Some; eauto. unfold prun_iw. rewrite RUNCTL. reflexivity. eassumption.
+Qed.
+
+Theorem forward_simu_par_wio_bblock ge fn rsr rsw mr sr sw mw bdy1 bdy2 ex sz rs' m' rs2 m2:
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_wio_bblock_aux ge fn bdy1 ex (Ptrofs.repr sz) rsr rsw mr mw = Next rs' m' ->
+ parexec_wio_body ge bdy2 rsr rs' mr m' = Next rs2 m2 ->
+ exists s2',
+ prun_iw Ge ((trans_block_aux bdy1 sz ex)++(trans_body bdy2)) sw sr = Some s2'
+ /\ match_states (State rs2 m2) s2'.
+Proof.
+ intros. exploit forward_simu_par_wio_bblock_aux. 4: eapply H2. all: eauto.
+ intros (s' & RUNAUX & MS').
+ exploit forward_simu_par_body. 4: eapply H3. all: eauto.
+ intros (s'' & RUNBDY2 & MS'').
+ eexists. split.
+ erewrite prun_iw_app_Some; eauto. eassumption.
+Qed.
+
+Lemma forward_simu_par_body_Stuck bdy: forall ge fn rsr mr sr rsw mw sw,
+ Ge = Genv ge fn ->
+ match_states (State rsr mr) sr ->
+ match_states (State rsw mw) sw ->
+ parexec_wio_body ge bdy rsr rsw mr mw = Stuck ->
+ prun_iw Ge (trans_body bdy) sw sr = None.
+Proof.
+ induction bdy.
+ - intros until sw. intros GENV MSR MSW WIO.
+ simpl in WIO. inv WIO.
+ - intros until sw. intros GENV MSR MSW WIO.
+ simpl in WIO. destruct (parexec_basic_instr _ _ _ _ _ _) eqn:PARBASIC.
+ * exploit forward_simu_par_wio_basic. 4: eapply PARBASIC. all: eauto.
+ intros (sw' & MPRUN & MS'). simpl prun_iw. rewrite MPRUN.
+ eapply IHbdy; eauto.
+ * exploit forward_simu_par_wio_basic_Stuck. 4: eapply PARBASIC. all: eauto.
+ intros X; simpl; rewrite X; auto.
+Qed.
+
+Lemma forward_simu_par_wio_stuck_bdy1 ge fn rs m s1' bdy1 sz ex:
+ Ge = Genv ge fn ->
+ match_states (State rs m) s1' ->
+ parexec_wio_bblock_aux ge fn bdy1 ex (Ptrofs.repr sz) rs rs m m = Stuck ->
+ prun_iw Ge ((trans_block_aux bdy1 sz ex)) s1' s1' = None.
+Proof.
+ unfold parexec_wio_bblock_aux, trans_block_aux; intros.
+ destruct (parexec_wio_body _ _ _ _ _ _) eqn:WIOB.
+ * exploit forward_simu_par_body. 4: eapply WIOB. all: eauto.
+ intros (s' & RUNB & MS').
+ erewrite prun_iw_app_Some; eauto.
+ exploit forward_simu_par_control_Stuck. 4: eauto. all: eauto.
+ simpl. intros X; rewrite X. auto.
+ * apply prun_iw_app_None. eapply forward_simu_par_body_Stuck. 4: eauto. all: eauto.
+Qed.
+
+Lemma forward_simu_par_wio_stuck_bdy2 ge fn rs m s1' bdy1 bdy2 sz ex m' rs':
+ Ge = Genv ge fn ->
+ match_states (State rs m) s1' ->
+ parexec_wio_bblock_aux ge fn bdy1 ex (Ptrofs.repr sz) rs rs m m = Next rs' m' ->
+ parexec_wio_body ge bdy2 rs rs' m m' = Stuck ->
+ prun_iw Ge ((trans_block_aux bdy1 sz ex)++(trans_body bdy2)) s1' s1'=None.
+Proof.
+ intros; exploit forward_simu_par_wio_bblock_aux. 4: eauto. all: eauto.
+ intros (s2' & X1 & X2).
+ erewrite prun_iw_app_Some; eauto.
+ eapply forward_simu_par_body_Stuck. 4: eauto. all: eauto.
+Qed.
+
+Lemma trans_body_perserves_permutation bdy1 bdy2:
+ Permutation bdy1 bdy2 ->
+ Permutation (trans_body bdy1) (trans_body bdy2).
+Proof.
+ induction 1; simpl; econstructor; eauto.
+Qed.
+
+Lemma trans_body_app bdy1: forall bdy2,
+ trans_body (bdy1++bdy2) = (trans_body bdy1) ++ (trans_body bdy2).
+Proof.
+ induction bdy1; simpl; congruence.
+Qed.
+
+Theorem trans_block_perserves_permutation bdy1 bdy2 b:
+ Permutation (bdy1 ++ bdy2) (body b) ->
+ Permutation (trans_block b) ((trans_block_aux bdy1 (size b) (exit b))++(trans_body bdy2)).
+Proof.
+ intro H; unfold trans_block, trans_block_aux.
+ eapply perm_trans.
+ - eapply Permutation_app_tail.
+ apply trans_body_perserves_permutation.
+ apply Permutation_sym; eapply H.
+ - rewrite trans_body_app. rewrite <-! app_assoc.
+ apply Permutation_app_head.
+ apply Permutation_app_comm.
+Qed.
+
+Theorem forward_simu_par:
+ forall rs1 m1 s1' b ge fn rs2 m2,
+ Ge = Genv ge fn ->
+ parexec_bblock ge fn b rs1 m1 (Next rs2 m2) ->
+ match_states (State rs1 m1) s1' ->
+ exists s2',
+ prun Ge (trans_block b) s1' (Some s2')
+ /\ match_states (State rs2 m2) s2'.
+Proof.
+ intros until m2. intros GENV PAREXEC MS.
+ inversion PAREXEC as (bdy1 & bdy2 & PERM & WIO).
+ exploit trans_block_perserves_permutation; eauto.
+ intros Perm.
+ remember (parexec_wio_bblock_aux _ _ _ _ _ _ _ _ _) as pwb.
+ destruct pwb; try congruence.
+ exploit forward_simu_par_wio_bblock; eauto. intros (s2' & PIW & MS').
+ unfold prun; simpl; eexists; split; eauto.
+Qed.
+
+Theorem forward_simu_par_stuck:
+ forall rs1 m1 s1' b ge fn,
+ Ge = Genv ge fn ->
+ parexec_bblock ge fn b rs1 m1 Stuck ->
+ match_states (State rs1 m1) s1' ->
+ prun Ge (trans_block b) s1' None.
+Proof.
+ intros until fn. intros GENV PAREXEC MS.
+ inversion PAREXEC as (bdy1 & bdy2 & PERM & WIO).
+ exploit trans_block_perserves_permutation; eauto.
+ intros Perm.
+ destruct (parexec_wio_bblock_aux _ _ _ _ _ _ _ _) eqn:WIOEXIT.
+ - econstructor; eauto. split. eapply forward_simu_par_wio_stuck_bdy2; eauto. auto.
+ - clear WIO. econstructor; eauto. split; eauto.
+ simpl; apply prun_iw_app_None; auto.
+ eapply forward_simu_par_wio_stuck_bdy1; eauto.
+Qed.
+
+Theorem forward_simu_par_alt:
+ forall rs1 m1 s1' b ge fn o2,
+ Ge = Genv ge fn ->
+ match_states (State rs1 m1) s1' ->
+ parexec_bblock ge fn b rs1 m1 o2 ->
+ exists o2',
+ prun Ge (trans_block b) s1' o2'
+ /\ match_outcome o2 o2'.
+Proof.
+ intros until o2. intros GENV MS PAREXEC. destruct o2 eqn:O2.
+ - exploit forward_simu_par; eauto. intros (s2' & PRUN & MS'). eexists. split. eassumption.
+ unfold match_outcome. eexists; split; auto.
+ - exploit forward_simu_par_stuck; eauto. intros. eexists; split; eauto.
+ constructor; auto.
+Qed.
+
+Lemma bblock_para_check_correct:
+ forall ge fn bb rs m rs' m' o,
+ Ge = Genv ge fn ->
+ exec_bblock ge fn bb rs m = Next rs' m' ->
+ bblock_para_check bb = true ->
+ parexec_bblock ge fn bb rs m o ->
+ o = Next rs' m'.
+Proof.
+ intros. unfold bblock_para_check in H1.
+ exploit forward_simu; eauto. eapply trans_state_match.
+ intros (s2' & EXEC & MS).
+ exploit forward_simu_par_alt. 2: apply (trans_state_match (State rs m)). all: eauto.
+ intros (o2' & PRUN & MO).
+ exploit parallelizable_correct. apply is_para_correct_aux. eassumption.
+ intro. eapply H3 in PRUN. clear H3. destruct o2'.
+ - inv PRUN. inv H3. unfold exec in EXEC. assert (x = s2') by congruence. subst. clear H.
+ assert (m0 = s2') by (apply functional_extensionality; auto). subst. clear H4.
+ destruct o; try discriminate. inv MO. inv H. assert (s2' = x) by congruence. subst.
+ exploit state_equiv. split. eapply MS. eapply H4. intros. inv H. reflexivity.
+ - unfold match_outcome in MO. destruct o.
+ + inv MO. inv H3. discriminate.
+ + clear MO. unfold exec in EXEC. rewrite EXEC in PRUN. discriminate.
+Qed.
+
+End SECT_PAR.
diff --git a/mppa_k1c/Asmvliw.v b/mppa_k1c/Asmvliw.v
new file mode 100644
index 00000000..a6e9f04b
--- /dev/null
+++ b/mppa_k1c/Asmvliw.v
@@ -0,0 +1,346 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* Prashanth Mundkur, SRI International *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* The contributions by Prashanth Mundkur are reused and adapted *)
+(* under the terms of a Contributor License Agreement between *)
+(* SRI International and INRIA. *)
+(* *)
+(* *********************************************************************)
+
+(** Abstract syntax and semantics for K1c assembly language. *)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Values.
+Require Import Memory.
+Require Import Events.
+Require Import Globalenvs.
+Require Import Smallstep.
+Require Import Locations.
+Require Stacklayout.
+Require Import Conventions.
+Require Import Errors.
+Require Export Asmblock.
+Require Import Sorting.Permutation.
+
+Local Open Scope asm.
+
+Section RELSEM.
+
+(** Execution of arith instructions *)
+
+Variable ge: genv.
+
+(* TODO: on pourrait mettre ça dans Asmblock pour factoriser le code
+ en définissant
+ exec_arith_instr ai rs := parexec_arith_instr ai rs rs
+*)
+Definition parexec_arith_instr (ai: ar_instruction) (rsr rsw: regset): regset :=
+ match ai with
+ | PArithR n d => rsw#d <- (arith_eval_r ge n)
+
+ | PArithRR n d s => rsw#d <- (arith_eval_rr n rsr#s)
+
+ | PArithRI32 n d i => rsw#d <- (arith_eval_ri32 n i)
+ | PArithRI64 n d i => rsw#d <- (arith_eval_ri64 n i)
+ | PArithRF32 n d i => rsw#d <- (arith_eval_rf32 n i)
+ | PArithRF64 n d i => rsw#d <- (arith_eval_rf64 n i)
+
+ | PArithRRR n d s1 s2 => rsw#d <- (arith_eval_rrr n rsr#s1 rsr#s2)
+ | PArithRRI32 n d s i => rsw#d <- (arith_eval_rri32 n rsr#s i)
+ | PArithRRI64 n d s i => rsw#d <- (arith_eval_rri64 n rsr#s i)
+
+ | PArithARRR n d s1 s2 => rsw#d <- (arith_eval_arrr n rsr#d rsr#s1 rsr#s2)
+ | PArithARRI32 n d s i => rsw#d <- (arith_eval_arri32 n rsr#d rsr#s i)
+ | PArithARRI64 n d s i => rsw#d <- (arith_eval_arri64 n rsr#d rsr#s i)
+ end.
+
+(** * load/store *)
+
+(* TODO: factoriser ? *)
+Definition parexec_load (chunk: memory_chunk) (rsr rsw: regset) (mr mw: mem)
+ (d: ireg) (a: ireg) (ofs: offset) :=
+ match (eval_offset ge ofs) with
+ | OK ptr =>
+ match Mem.loadv chunk mr (Val.offset_ptr (rsr a) ptr) with
+ | None => Stuck
+ | Some v => Next (rsw#d <- v) mw
+ end
+ | _ => Stuck
+ end.
+
+Definition parexec_store (chunk: memory_chunk) (rsr rsw: regset) (mr mw: mem)
+ (s: ireg) (a: ireg) (ofs: offset) :=
+ match (eval_offset ge ofs) with
+ | OK ptr =>
+ match Mem.storev chunk mr (Val.offset_ptr (rsr a) ptr) (rsr s) with
+ | None => Stuck
+ | Some m' => Next rsw m'
+ end
+ | _ => Stuck
+ end.
+
+(* rem: parexec_store = exec_store *)
+
+(** * basic instructions *)
+
+(* TODO: factoriser ? *)
+Definition parexec_basic_instr (bi: basic) (rsr rsw: regset) (mr mw: mem) :=
+ match bi with
+ | PArith ai => Next (parexec_arith_instr ai rsr rsw) mw
+
+ | PLoadRRO n d a ofs => parexec_load (load_chunk n) rsr rsw mr mw d a ofs
+
+ | PStoreRRO n s a ofs => parexec_store (store_chunk n) rsr rsw mr mw s a ofs
+
+ | Pallocframe sz pos =>
+ let (mw, stk) := Mem.alloc mr 0 sz in
+ let sp := (Vptr stk Ptrofs.zero) in
+ match Mem.storev Mptr mw (Val.offset_ptr sp pos) rsr#SP with
+ | None => Stuck
+ | Some mw => Next (rsw #FP <- (rsr SP) #SP <- sp #RTMP <- Vundef) mw
+ end
+
+ | Pfreeframe sz pos =>
+ match Mem.loadv Mptr mr (Val.offset_ptr rsr#SP pos) with
+ | None => Stuck
+ | Some v =>
+ match rsr SP with
+ | Vptr stk ofs =>
+ match Mem.free mr stk 0 sz with
+ | None => Stuck
+ | Some mw => Next (rsw#SP <- v #RTMP <- Vundef) mw
+ end
+ | _ => Stuck
+ end
+ end
+ | Pget rd ra =>
+ match ra with
+ | RA => Next (rsw#rd <- (rsr#ra)) mw
+ | _ => Stuck
+ end
+ | Pset ra rd =>
+ match ra with
+ | RA => Next (rsw#ra <- (rsr#rd)) mw
+ | _ => Stuck
+ end
+ | Pnop => Next rsw mw
+end.
+
+(* parexec with writes-in-order *)
+Fixpoint parexec_wio_body (body: list basic) (rsr rsw: regset) (mr mw: mem) :=
+ match body with
+ | nil => Next rsw mw
+ | bi::body' =>
+ match parexec_basic_instr bi rsr rsw mr mw with
+ | Next rsw mw => parexec_wio_body body' rsr rsw mr mw
+ | Stuck => Stuck
+ end
+ end.
+
+(** Manipulations over the [PC] register: continuing with the next
+ instruction ([nextblock]) or branching to a label ([goto_label]). *)
+
+(* TODO: factoriser ? *)
+Definition par_nextblock size_b (rs: regset) :=
+ rs#PC <- (Val.offset_ptr rs#PC size_b).
+
+(* TODO: factoriser ? *)
+Definition par_goto_label (f: function) (lbl: label) (rsr rsw: regset) (mw: mem) :=
+ match label_pos lbl 0 (fn_blocks f) with
+ | None => Stuck
+ | Some pos =>
+ match rsr#PC with
+ | Vptr b ofs => Next (rsw#PC <- (Vptr b (Ptrofs.repr pos))) mw
+ | _ => Stuck
+ end
+ end.
+
+(** Evaluating a branch
+
+Warning: in m PC is assumed to be already pointing on the next instruction !
+
+*)
+(* TODO: factoriser ? *)
+Definition par_eval_branch (f: function) (l: label) (rsr rsw: regset) (mw: mem) (res: option bool) :=
+ match res with
+ | Some true => par_goto_label f l rsr rsw mw
+ | Some false => Next (rsw # PC <- (rsr PC)) mw
+ | None => Stuck
+ end.
+
+
+(** Execution of a single control-flow instruction [i] in initial state [rs] and
+ [m]. Return updated state.
+
+ As above: PC is assumed to be incremented on the next block before the control-flow instruction
+
+ For instructions that correspond tobuiltin
+ actual RISC-V instructions, the cases are straightforward
+ transliterations of the informal descriptions given in the RISC-V
+ user-mode specification. For pseudo-instructions, refer to the
+ informal descriptions given above.
+
+ Note that we set to [Vundef] the registers used as temporaries by
+ the expansions of the pseudo-instructions, so that the RISC-V code
+ we generate cannot use those registers to hold values that must
+ survive the execution of the pseudo-instruction. *)
+
+Definition parexec_control (f: function) (oc: option control) (rsr rsw: regset) (mw: mem) :=
+ match oc with
+ | Some ic =>
+(** Get/Set system registers *)
+ match ic with
+
+
+(** Branch Control Unit instructions *)
+ | Pret =>
+ Next (rsw#PC <- (rsr#RA)) mw
+ | Pcall s =>
+ Next (rsw#RA <- (rsr#PC) #PC <- (Genv.symbol_address ge s Ptrofs.zero)) mw
+ | Picall r =>
+ Next (rsw#RA <- (rsr#PC) #PC <- (rsr#r)) mw
+ | Pgoto s =>
+ Next (rsw#PC <- (Genv.symbol_address ge s Ptrofs.zero)) mw
+ | Pigoto r =>
+ Next (rsw#PC <- (rsr#r)) mw
+ | Pj_l l =>
+ par_goto_label f l rsr rsw mw
+ | Pcb bt r l =>
+ match cmp_for_btest bt with
+ | (Some c, Int) => par_eval_branch f l rsr rsw mw (Val.cmp_bool c rsr#r (Vint (Int.repr 0)))
+ | (Some c, Long) => par_eval_branch f l rsr rsw mw (Val.cmpl_bool c rsr#r (Vlong (Int64.repr 0)))
+ | (None, _) => Stuck
+ end
+ | Pcbu bt r l =>
+ match cmpu_for_btest bt with
+ | (Some c, Int) => par_eval_branch f l rsr rsw mw (Val_cmpu_bool c rsr#r (Vint (Int.repr 0)))
+ | (Some c, Long) => par_eval_branch f l rsr rsw mw (Val_cmplu_bool c rsr#r (Vlong (Int64.repr 0)))
+ | (None, _) => Stuck
+ end
+
+
+(** Pseudo-instructions *)
+ | Pbuiltin ef args res =>
+ Stuck (**r treated specially below *)
+ end
+ | None => Next (rsw#PC <- (rsr#PC)) mw
+end.
+
+
+Definition parexec_wio_bblock_aux (f: function) bdy ext size_b (rsr rsw: regset) (mr mw: mem): outcome :=
+ match parexec_wio_body bdy rsr rsw mr mw with
+ | Next rsw mw =>
+ let rsr := par_nextblock size_b rsr in
+ let rsw := par_nextblock size_b rsw in
+ parexec_control f ext rsr rsw mw
+ | Stuck => Stuck
+ end.
+
+Definition parexec_wio_bblock (f: function) (b: bblock) (rs: regset) (m: mem): outcome :=
+ parexec_wio_bblock_aux f (body b) (exit b) (Ptrofs.repr (size b)) rs rs m m.
+
+Definition parexec_bblock (f: function) (b: bblock) (rs: regset) (m: mem) (o: outcome): Prop :=
+ exists bdy1 bdy2, Permutation (bdy1++bdy2) (body b) /\
+ o=match parexec_wio_bblock_aux f bdy1 (exit b) (Ptrofs.repr (size b)) rs rs m m with
+ | Next rsw mw => parexec_wio_body bdy2 rs rsw m mw
+ | Stuck => Stuck
+ end.
+
+Lemma parexec_bblock_write_in_order f b rs m:
+ parexec_bblock f b rs m (parexec_wio_bblock f b rs m).
+Proof.
+ exists (body b). exists nil.
+ constructor 1.
+ - rewrite app_nil_r; auto.
+ - unfold parexec_wio_bblock.
+ destruct (parexec_wio_bblock_aux f _ _ _ _ _); simpl; auto.
+Qed.
+
+Inductive step: state -> trace -> state -> Prop :=
+ | exec_step_internal:
+ forall b ofs f bi rs m rs' m',
+ rs PC = Vptr b ofs ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bi ->
+ parexec_wio_bblock f bi rs m = Next rs' m' ->
+ (forall o, parexec_bblock f bi rs m o -> o=(Next rs' m')) ->
+ step (State rs m) E0 (State rs' m')
+ | exec_step_builtin:
+ forall b ofs f ef args res rs m vargs t vres rs' m' bi,
+ rs PC = Vptr b ofs ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ find_bblock (Ptrofs.unsigned ofs) f.(fn_blocks) = Some bi ->
+ exit bi = Some (PExpand (Pbuiltin ef args res)) ->
+ eval_builtin_args ge rs (rs SP) m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rs' = nextblock bi
+ (set_res res vres
+ (undef_regs (map preg_of (destroyed_by_builtin ef))
+ (rs#RTMP <- Vundef))) ->
+ step (State rs m) t (State rs' m')
+ | exec_step_external:
+ forall b ef args res rs m t rs' m',
+ rs PC = Vptr b Ptrofs.zero ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ external_call ef ge args m t res m' ->
+ extcall_arguments rs m (ef_sig ef) args ->
+ rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs RA) ->
+ step (State rs m) t (State rs' m')
+ .
+
+End RELSEM.
+
+(** Execution of whole programs. *)
+
+Definition semantics (p: program) :=
+ Semantics step (initial_state p) final_state (Genv.globalenv p).
+
+Lemma semantics_determinate: forall p, determinate (semantics p).
+Proof.
+Ltac Equalities :=
+ match goal with
+ | [ H1: ?a = ?b, H2: ?a = ?c |- _ ] =>
+ rewrite H1 in H2; inv H2; Equalities
+ | _ => idtac
+ end.
+ intros; constructor; simpl; intros.
+- (* determ *)
+ inv H; inv H0; Equalities.
+ + split. constructor. auto.
+ + unfold parexec_wio_bblock, parexec_wio_bblock_aux in H4. destruct (parexec_wio_body _ _ _ _ _ _); try discriminate.
+ rewrite H10 in H4. discriminate.
+ + unfold parexec_wio_bblock, parexec_wio_bblock_aux in H11. destruct (parexec_wio_body _ _ _ _ _ _); try discriminate.
+ rewrite H4 in H11. discriminate.
+ + assert (vargs0 = vargs) by (eapply eval_builtin_args_determ; eauto). subst vargs0.
+ exploit external_call_determ. eexact H6. eexact H13. intros [A B].
+ split. auto. intros. destruct B; auto. subst. auto.
+ + assert (args0 = args) by (eapply extcall_arguments_determ; eauto). subst args0.
+ exploit external_call_determ. eexact H3. eexact H8. intros [A B].
+ split. auto. intros. destruct B; auto. subst. auto.
+- (* trace length *)
+ red; intros. inv H; simpl.
+ omega.
+ eapply external_call_trace_length; eauto.
+ eapply external_call_trace_length; eauto.
+- (* initial states *)
+ inv H; inv H0. f_equal. congruence.
+- (* final no step *)
+ assert (NOTNULL: forall b ofs, Vnullptr <> Vptr b ofs).
+ { intros; unfold Vnullptr; destruct Archi.ptr64; congruence. }
+ inv H. unfold Vzero in H0. red; intros; red; intros.
+ inv H; rewrite H0 in *; eelim NOTNULL; eauto.
+- (* final states *)
+ inv H; inv H0. congruence.
+Qed. \ No newline at end of file
diff --git a/mppa_k1c/PostpassScheduling.v b/mppa_k1c/PostpassScheduling.v
index 6700e684..ab4bc9c9 100644
--- a/mppa_k1c/PostpassScheduling.v
+++ b/mppa_k1c/PostpassScheduling.v
@@ -352,8 +352,9 @@ Definition verified_schedule_nob (bb : bblock) : res (list bblock) :=
do tbb <- concat_all lbb;
do sizecheck <- verify_size bb lbb;
do schedcheck <- verify_schedule bb' tbb;
- do parcheck <- verify_par lbb;
- stick_header_code (header bb) lbb.
+ do res <- stick_header_code (header bb) lbb;
+ do parcheck <- verify_par res;
+ OK res.
Lemma verified_schedule_nob_size:
forall bb lbb, verified_schedule_nob bb = OK lbb -> size bb = size_blocks lbb.
@@ -379,7 +380,7 @@ Lemma verified_schedule_nob_header:
/\ Forall (fun b => header b = nil) lbb.
Proof.
intros. split.
- - monadInv H. unfold stick_header_code in EQ4. destruct (hd_error _); try discriminate. inv EQ4.
+ - monadInv H. unfold stick_header_code in EQ2. destruct (hd_error _); try discriminate. inv EQ2.
simpl. reflexivity.
- apply verified_schedule_nob_no_header_in_middle in H. assumption.
Qed.
diff --git a/mppa_k1c/PostpassSchedulingproof.v b/mppa_k1c/PostpassSchedulingproof.v
index b59c381c..4e33fc90 100644
--- a/mppa_k1c/PostpassSchedulingproof.v
+++ b/mppa_k1c/PostpassSchedulingproof.v
@@ -71,7 +71,7 @@ Lemma next_eq:
forall (rs rs': regset) m m',
rs = rs' -> m = m' -> Next rs m = Next rs' m'.
Proof.
- intros. congruence.
+ intros; apply f_equal2; auto.
Qed.
Lemma regset_double_set:
@@ -238,7 +238,7 @@ Proof.
rewrite <- Zplus_mod. auto.
Qed.
-Section PRESERVATION.
+Section PRESERVATION_ASMBLOCK.
Variables prog tprog: program.
Hypothesis TRANSL: match_prog prog tprog.
@@ -660,7 +660,7 @@ Proof.
eapply external_call_symbols_preserved; eauto. apply senv_preserved.
Qed.
-Theorem transf_program_correct:
+Theorem transf_program_correct_Asmblock:
forward_simulation (Asmblock.semantics prog) (Asmblock.semantics tprog).
Proof.
eapply forward_simulation_plus.
@@ -670,4 +670,174 @@ Proof.
- apply transf_step_correct.
Qed.
+End PRESERVATION_ASMBLOCK.
+
+
+
+
+Require Import Asmvliw.
+
+
+Lemma verified_par_checks_alls_bundles lb x: forall bundle,
+ verify_par lb = OK x ->
+ List.In bundle lb -> verify_par_bblock bundle = OK tt.
+Proof.
+ induction lb; simpl; try tauto.
+ intros bundle H; monadInv H.
+ destruct 1; subst; eauto.
+ destruct x0; auto.
+Qed.
+
+
+Lemma verified_schedule_nob_checks_alls_bundles bb lb bundle:
+ verified_schedule_nob bb = OK lb ->
+ List.In bundle lb -> verify_par_bblock bundle = OK tt.
+Proof.
+ unfold verified_schedule_nob. intros H;
+ monadInv H. destruct x3.
+ intros; eapply verified_par_checks_alls_bundles; eauto.
+Qed.
+
+Lemma verify_par_bblock_PExpand bb i:
+ exit bb = Some (PExpand i) -> verify_par_bblock bb = OK tt.
+Proof.
+ destruct bb as [h bdy ext H]; simpl.
+ intros; subst. destruct i.
+ generalize H.
+ rewrite <- AB.wf_bblock_refl in H.
+ destruct H as [H H0].
+ unfold AB.builtin_alone in H0. erewrite H0; eauto.
+Qed.
+
+Local Hint Resolve verified_schedule_nob_checks_alls_bundles.
+
+Lemma verified_schedule_checks_alls_bundles bb lb bundle:
+ verified_schedule bb = OK lb ->
+ List.In bundle lb -> verify_par_bblock bundle = OK tt.
+Proof.
+ unfold verified_schedule. remember (exit bb) as exb.
+ destruct exb as [c|]; eauto.
+ destruct c as [i|]; eauto.
+ destruct i; intros H. inversion_clear H; simpl.
+ intuition subst.
+ intros; eapply verify_par_bblock_PExpand; eauto.
+Qed.
+
+Lemma transf_blocks_checks_all_bundles lbb: forall lb bundle,
+ transf_blocks lbb = OK lb ->
+ List.In bundle lb -> verify_par_bblock bundle = OK tt.
+Proof.
+ induction lbb; simpl.
+ - intros lb bundle H; inversion_clear H. simpl; try tauto.
+ - intros lb bundle H0.
+ monadInv H0.
+ rewrite in_app. destruct 1; eauto.
+ eapply verified_schedule_checks_alls_bundles; eauto.
+Qed.
+
+Lemma find_bblock_Some_in lb:
+ forall ofs b, find_bblock ofs lb = Some b -> List.In b lb.
+Proof.
+ induction lb; simpl; try congruence.
+ intros ofs b.
+ destruct (zlt ofs 0); try congruence.
+ destruct (zeq ofs 0); eauto.
+ intros X; inversion X; eauto.
+Qed.
+
+Section PRESERVATION_ASMVLIW.
+
+Variables prog tprog: program.
+Hypothesis TRANSL: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Lemma all_bundles_are_checked b ofs f bundle:
+ Genv.find_funct_ptr (globalenv (Asmblock.semantics tprog)) b = Some (Internal f) ->
+ find_bblock ofs (fn_blocks f) = Some bundle ->
+ verify_par_bblock bundle = OK tt.
+Proof.
+ unfold match_prog, match_program in TRANSL.
+ unfold Genv.find_funct_ptr; simpl; intros X.
+ destruct (Genv.find_def_match_2 TRANSL b) as [|f0 y H]; try congruence.
+ destruct y as [tf0|]; try congruence.
+ inversion X as [H1]. subst. clear X.
+ remember (@Gfun fundef unit (Internal f)) as f2.
+ destruct H as [ctx' f1 f2 H0|]; try congruence.
+ inversion Heqf2 as [H2]. subst; clear Heqf2.
+ unfold transf_fundef, transf_partial_fundef in H.
+ destruct f1 as [f1|f1]; try congruence.
+ unfold transf_function, transl_function in H.
+ monadInv H. monadInv EQ.
+ destruct (zlt Ptrofs.max_unsigned (size_blocks (fn_blocks _))); simpl in *|-; try congruence.
+ injection EQ1; intros; subst.
+ monadInv EQ0. simpl in * |-.
+ intros; exploit transf_blocks_checks_all_bundles; eauto.
+ intros; eapply find_bblock_Some_in; eauto.
+Qed.
+
+Lemma checked_bundles_are_parexec_equiv f bundle rs rs' m m' o:
+ exec_bblock (globalenv (Asmblock.semantics tprog)) f bundle rs m = Next rs' m' ->
+ verify_par_bblock bundle = OK tt ->
+ parexec_bblock (globalenv (semantics tprog)) f bundle rs m o -> o = Next rs' m'.
+Proof.
+ intros. unfold verify_par_bblock in H0. destruct (Asmblockdeps.bblock_para_check _) eqn:BPC; try discriminate. clear H0.
+ simpl in H. simpl in H1.
+ eapply Asmblockdeps.bblock_para_check_correct; eauto.
+Qed.
+
+Lemma seqexec_parexec_equiv b ofs f bundle rs rs' m m' o:
+ Genv.find_funct_ptr (globalenv (Asmblock.semantics tprog)) b = Some (Internal f) ->
+ find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bundle ->
+ exec_bblock (globalenv (Asmblock.semantics tprog)) f bundle rs m = Next rs' m' ->
+ parexec_bblock (globalenv (semantics tprog)) f bundle rs m o -> o = Next rs' m'.
+Proof.
+ intros; eapply checked_bundles_are_parexec_equiv; eauto.
+ eapply all_bundles_are_checked; eauto.
+Qed.
+
+Lemma seqexec_parexec_wio b ofs f bundle rs rs' m m':
+ Genv.find_funct_ptr (globalenv (Asmblock.semantics tprog)) b = Some (Internal f) ->
+ find_bblock (Ptrofs.unsigned ofs) (fn_blocks f) = Some bundle ->
+ exec_bblock (globalenv (Asmblock.semantics tprog)) f bundle rs m = Next rs' m' ->
+ parexec_wio_bblock (globalenv (semantics tprog)) f bundle rs m = Next rs' m'.
+Proof.
+ intros; erewrite <- seqexec_parexec_equiv; eauto.
+ eapply parexec_bblock_write_in_order.
+Qed.
+
+
+Theorem transf_program_correct_Asmvliw:
+ forward_simulation (Asmblock.semantics tprog) (Asmvliw.semantics tprog).
+Proof.
+ eapply forward_simulation_step with (match_states:=fun (s1:Asmblock.state) s2 => s1=s2); eauto.
+ - intros; subst; auto.
+ - intros s1 t s1' H s2 H0; subst; inversion H; clear H; subst; eexists; split; eauto.
+ + eapply exec_step_internal; eauto.
+ eapply seqexec_parexec_wio; eauto.
+ intros; eapply seqexec_parexec_equiv; eauto.
+ + eapply exec_step_builtin; eauto.
+ + eapply exec_step_external; eauto.
+Qed.
+
+End PRESERVATION_ASMVLIW.
+
+
+
+
+Section PRESERVATION.
+
+Variables prog tprog: program.
+Hypothesis TRANSL: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Theorem transf_program_correct:
+ forward_simulation (Asmblock.semantics prog) (Asmvliw.semantics tprog).
+Proof.
+ eapply compose_forward_simulations.
+ eapply transf_program_correct_Asmblock; eauto.
+ eapply transf_program_correct_Asmvliw; eauto.
+Qed.
+
End PRESERVATION.
diff --git a/mppa_k1c/abstractbb/Parallelizability.v b/mppa_k1c/abstractbb/Parallelizability.v
index b6d1f142..065c0922 100644
--- a/mppa_k1c/abstractbb/Parallelizability.v
+++ b/mppa_k1c/abstractbb/Parallelizability.v
@@ -79,6 +79,24 @@ Proof.
+ intros H1; rewrite H1; simpl; auto.
Qed.
+Lemma prun_iw_app_None p1: forall m1 old p2,
+ prun_iw p1 m1 old = None ->
+ prun_iw (p1++p2) m1 old = None.
+Proof.
+ induction p1; simpl; try congruence.
+ intros; destruct (macro_prun _ _ _); simpl; auto.
+Qed.
+
+Lemma prun_iw_app_Some p1: forall m1 old m2 p2,
+ prun_iw p1 m1 old = Some m2 ->
+ prun_iw (p1++p2) m1 old = prun_iw p2 m2 old.
+Proof.
+ induction p1; simpl; try congruence.
+ intros; destruct (macro_prun _ _ _); simpl; auto.
+ congruence.
+Qed.
+
+
End PARALLEL.
End ParallelSemantics.