aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--doc/index-kvx.html2
-rw-r--r--kvx/abstractbb/AbstractBasicBlocksDef.v44
-rw-r--r--kvx/abstractbb/ImpSimuTest.v27
-rw-r--r--kvx/abstractbb/Parallelizability.v25
-rw-r--r--kvx/abstractbb/SeqSimuTheory.v21
5 files changed, 72 insertions, 47 deletions
diff --git a/doc/index-kvx.html b/doc/index-kvx.html
index ff3fbc17..4f666cc3 100644
--- a/doc/index-kvx.html
+++ b/doc/index-kvx.html
@@ -325,7 +325,7 @@ This IR is generic over the processor, even if currently, only used for KVX.
<TD>Flattening bundles (only a bureaucratic operation)</TD>
<TD>Asmvliw to Asm</TD>
<TD><A HREF="html/compcert.kvx.Asmgen.html"><I>Asmgen</I></A></TD>
- <TD><A HREF="html/compcert.kvx.Asmgenproof.html"><I>Asmgenproof</I></A></TD>
+ <TD><A HREF="html/compcert.kvx.Asmgenproof.html"><I>Asmgenproof</I></A> (whole simulation proof from <tt>Mach</tt> to <tt>Asm</tt>)</TD>
</TR>
</TABLE>
diff --git a/kvx/abstractbb/AbstractBasicBlocksDef.v b/kvx/abstractbb/AbstractBasicBlocksDef.v
index 0b1c502d..948ed660 100644
--- a/kvx/abstractbb/AbstractBasicBlocksDef.v
+++ b/kvx/abstractbb/AbstractBasicBlocksDef.v
@@ -45,7 +45,7 @@ End LangParam.
-(** * Syntax and (sequential) semantics of "basic blocks" *)
+(** * Syntax and (sequential) semantics of "abstract basic blocks" *)
Module MkSeqLanguage(P: LangParam).
Export P.
@@ -62,12 +62,12 @@ Definition assign (m: mem) (x:R.t) (v: value): mem
:= fun y => if R.eq_dec x y then v else m y.
-(** expressions *)
+(** Expressions *)
Inductive exp :=
- | PReg (x:R.t)
- | Op (o:op) (le: list_exp)
- | Old (e: exp)
+ | PReg (x:R.t) (**r pseudo-register *)
+ | Op (o:op) (le: list_exp) (**r operation *)
+ | Old (e: exp) (**r evaluation of [e] in the initial state of the instruction (see [inst] below) *)
with list_exp :=
| Enil
| Econs (e:exp) (le:list_exp)
@@ -95,7 +95,8 @@ with list_exp_eval (le: list_exp) (m old: mem): option (list value) :=
| LOld le => list_exp_eval le old old
end.
-Definition inst := list (R.t * exp). (* = a sequence of assignments *)
+(** An instruction represents a sequence of assignments where [Old] refers to the initial state of the sequence. *)
+Definition inst := list (R.t * exp).
Fixpoint inst_run (i: inst) (m old: mem): option mem :=
match i with
@@ -107,6 +108,7 @@ Fixpoint inst_run (i: inst) (m old: mem): option mem :=
end
end.
+(** A basic block is a sequence of instructions. *)
Definition bblock := list inst.
Fixpoint run (p: bblock) (m: mem): option mem :=
@@ -250,12 +252,16 @@ Qed.
End SEQLANG.
-Module Terms.
-(** terms in the symbolic evaluation
-NB: such a term represents the successive computations in one given pseudo-register
+(** * Terms in the symbolic execution *)
+
+(** Such a term represents the successive computations in one given pseudo-register.
+The [hid] has no formal semantics: it is only used by the hash-consing oracle (itself dynamically checked to behave like an identity function).
+
*)
+Module Terms.
+
Inductive term :=
| Input (x:R.t) (hid:hashcode)
| App (o: op) (l: list_term) (hid:hashcode)
@@ -334,11 +340,21 @@ Proof.
- rewrite IHl; clear IHl. intuition (congruence || eauto).
Qed.
+(** * Rewriting rules in the symbolic execution *)
+
+(** The symbolic execution is parametrized by rewriting rules on pseudo-terms. *)
+
Record pseudo_term: Type := intro_fail {
mayfail: list term;
effect: term
}.
+(** Simulation relation between a term and a pseudo-term *)
+
+Definition match_pt (t: term) (pt: pseudo_term) :=
+ (forall ge m, term_eval ge t m <> None <-> allvalid ge pt.(mayfail) m)
+ /\ (forall ge m0 m1, term_eval ge t m0 = Some m1 -> term_eval ge pt.(effect) m0 = Some m1).
+
Lemma inf_option_equivalence (A:Type) (o1 o2: option A):
(o1 <> None -> o1 = o2) <-> (forall m1, o1 = Some m1 -> o2 = Some m1).
Proof.
@@ -346,10 +362,6 @@ Proof.
symmetry; eauto.
Qed.
-Definition match_pt (t: term) (pt: pseudo_term) :=
- (forall ge m, term_eval ge t m <> None <-> allvalid ge pt.(mayfail) m)
- /\ (forall ge m0 m1, term_eval ge t m0 = Some m1 -> term_eval ge pt.(effect) m0 = Some m1).
-
Lemma intro_fail_correct (l: list term) (t: term) :
(forall ge m, term_eval ge t m <> None <-> allvalid ge l m) -> match_pt t (intro_fail l t).
Proof.
@@ -357,6 +369,7 @@ Proof.
Qed.
Hint Resolve intro_fail_correct: wlp.
+(** The default reduction of a term to a pseudo-term *)
Definition identity_fail (t: term):= intro_fail [t] t.
Lemma identity_fail_correct (t: term): match_pt t (identity_fail t).
@@ -366,6 +379,7 @@ Qed.
Global Opaque identity_fail.
Hint Resolve identity_fail_correct: wlp.
+(** The reduction for constant term *)
Definition nofail (is_constant: op -> bool) (t: term):=
match t with
| Input x _ => intro_fail ([])%list t
@@ -385,6 +399,7 @@ Qed.
Global Opaque nofail.
Hint Resolve nofail_correct: wlp.
+(** Term equivalence preserve the simulation by pseudo-terms *)
Definition term_equiv t1 t2:= forall ge m, term_eval ge t1 m = term_eval ge t2 m.
Global Instance term_equiv_Equivalence : Equivalence term_equiv.
@@ -401,6 +416,7 @@ Proof.
Qed.
Hint Resolve match_pt_term_equiv: wlp.
+(** Other generic reductions *)
Definition app_fail (l: list term) (pt: pseudo_term): pseudo_term :=
{| mayfail := List.rev_append l pt.(mayfail); effect := pt.(effect) |}.
@@ -431,6 +447,8 @@ Extraction Inline app_fail.
Import ImpCore.Notations.
Local Open Scope impure_scope.
+(** Specification of rewriting functions in parameter of the symbolic execution: in the impure monad, because the rewriting functions produce hash-consed terms (wrapped in pseudo-terms).
+*)
Record reduction:= {
result:> term -> ?? pseudo_term;
result_correct: forall t, WHEN result t ~> pt THEN match_pt t pt;
diff --git a/kvx/abstractbb/ImpSimuTest.v b/kvx/abstractbb/ImpSimuTest.v
index c914eee1..97d1a234 100644
--- a/kvx/abstractbb/ImpSimuTest.v
+++ b/kvx/abstractbb/ImpSimuTest.v
@@ -10,9 +10,11 @@
(* *)
(* *************************************************************)
-(** Implementation of a symbolic execution of sequential semantics of Abstract Basic Blocks
+(** Implementation of a simulation test (ie a "scheduling verifier") for the sequential semantics of Abstract Basic Blocks.
-with imperative hash-consing, and rewriting.
+It is based on a symbolic execution procedure of Abstract Basic Blocks with imperative hash-consing and rewriting.
+
+It also provides debugging information when the test fails.
*)
@@ -32,6 +34,7 @@ Import ListNotations.
Local Open Scope list_scope.
+(** * Interface of (impure) equality tests for operators *)
Module Type ImpParam.
Include LangParam.
@@ -54,6 +57,8 @@ Include MkSeqLanguage LP.
End ISeqLanguage.
+(** * A generic dictinary on PseudoRegisters with an impure equality test *)
+
Module Type ImpDict.
Declare Module R: PseudoRegisters.
@@ -91,26 +96,27 @@ Parameter eq_test_correct: forall A (d1 d2: t A),
(* NB: we could also take an eq_test on R.t (but not really useful with "pure" dictionaries *)
-
-(* only for debugging *)
+(** only for debugging *)
Parameter not_eq_witness: forall {A}, t A -> t A -> ?? option R.t.
End ImpDict.
+(** * Specification of the provided tests *)
Module Type ImpSimuInterface.
Declare Module CoreL: ISeqLanguage.
Import CoreL.
Import Terms.
+(** the silent test (without debugging informations) *)
Parameter bblock_simu_test: reduction -> bblock -> bblock -> ?? bool.
Parameter bblock_simu_test_correct: forall reduce (p1 p2 : bblock),
WHEN bblock_simu_test reduce p1 p2 ~> b
THEN b = true -> forall ge : genv, bblock_simu ge p1 p2.
-
+(** the verbose test extended with debugging informations *)
Parameter verb_bblock_simu_test
: reduction ->
(R.t -> ?? pstring) ->
@@ -127,6 +133,7 @@ Parameter verb_bblock_simu_test_correct:
End ImpSimuInterface.
+(** * Implementation of the provided tests *)
Module ImpSimu (L: ISeqLanguage) (Dict: ImpDict with Module R:=L.LP.R): ImpSimuInterface with Module CoreL := L.
@@ -168,7 +175,7 @@ Section SimuWithReduce.
Variable reduce: reduction.
-Section CanonBuilding.
+Section CanonBuilding. (** Implementation of the symbolic execution (ie a "canonical form" representing the semantics of an abstract basic block) *)
Variable hC_term: hashinfo term -> ?? term.
Hypothesis hC_term_correct: forall t, WHEN hC_term t ~> t' THEN forall ge m, term_eval ge (hdata t) m = term_eval ge t' m.
@@ -1117,9 +1124,9 @@ Extraction Inline lift.
End ImpSimu.
-Require Import FMapPositive.
-
+(** * Implementation of the Dictionary (based on PositiveMap) *)
+Require Import FMapPositive.
Require Import PArith.
Require Import FMapPositive.
@@ -1206,7 +1213,7 @@ Proof.
Qed.
Global Opaque eq_test.
-(* ONLY FOR DEBUGGING INFO: get some key of a non-empty d *)
+(** ONLY FOR DEBUGGING INFO: get some key of a non-empty d *)
Fixpoint pick {A} (d: t A): ?? R.t :=
match d with
| Leaf _ => FAILWITH "unexpected empty dictionary"
@@ -1219,7 +1226,7 @@ Fixpoint pick {A} (d: t A): ?? R.t :=
RET (xO p)
end.
-(* ONLY FOR DEBUGGING INFO: find one variable on which d1 and d2 differs *)
+(** ONLY FOR DEBUGGING INFO: find one variable on which d1 and d2 differs *)
Fixpoint not_eq_witness {A} (d1 d2: t A): ?? option R.t :=
match d1, d2 with
| Leaf _, Leaf _ => RET None
diff --git a/kvx/abstractbb/Parallelizability.v b/kvx/abstractbb/Parallelizability.v
index feebeee5..79ec9038 100644
--- a/kvx/abstractbb/Parallelizability.v
+++ b/kvx/abstractbb/Parallelizability.v
@@ -26,7 +26,7 @@ Require Import Sorting.Permutation.
Require Import Bool.
Local Open Scope lazy_bool_scope.
-
+(** * Definition of the parallel semantics *)
Module ParallelSemantics (L: SeqLanguage).
Export L.
@@ -590,17 +590,17 @@ End PARALLELI.
End ParallelizablityChecking.
-Module Type PseudoRegSet.
-
-Declare Module R: PseudoRegisters.
-
-(** We assume a datatype [t] refining (list R.t)
+(** * We assume a datatype [PseudoRegSet.t] refining [list R.t] *)
+(**
This data-refinement is given by an abstract "invariant" match_frame below,
preserved by the following operations.
-
*)
+Module Type PseudoRegSet.
+
+Declare Module R: PseudoRegisters.
+
Parameter t: Type.
Parameter match_frame: t -> (list R.t) -> Prop.
@@ -716,6 +716,11 @@ End ParallelChecks.
+(** * Implementing the datatype [PosPseudoRegSet.t] refining [list R.t] *)
+
+(* This data-refinement is given by an abstract "invariant" match_frame below,
+preserved by the following operations.
+*)
Require Import PArith.
Require Import MSets.MSetPositive.
@@ -724,12 +729,6 @@ Module PosPseudoRegSet <: PseudoRegSet with Module R:=Pos.
Module R:=Pos.
-(** We assume a datatype [t] refining (list R.t)
-
-This data-refinement is given by an abstract "invariant" match_frame below,
-preserved by the following operations.
-
-*)
Definition t:=PositiveSet.t.
diff --git a/kvx/abstractbb/SeqSimuTheory.v b/kvx/abstractbb/SeqSimuTheory.v
index 61f8f2ec..a957c50a 100644
--- a/kvx/abstractbb/SeqSimuTheory.v
+++ b/kvx/abstractbb/SeqSimuTheory.v
@@ -55,13 +55,14 @@ with list_term_eval ge (l: list_term) (m: mem) {struct l}: option (list value) :
end
end.
-(* the symbolic memory:
- - pre: pre-condition expressing that the computation has not yet abort on a None.
- - post: the post-condition for each pseudo-register
-*)
-Record smem:= {pre: genv -> mem -> Prop; post:> R.t -> term}.
-
-(** initial symbolic memory *)
+(** The (abstract) symbolic memory state *)
+Record smem :=
+{
+ pre: genv -> mem -> Prop; (**r pre-condition expressing that the computation has not yet abort on a None. *)
+ post:> R.t -> term (**r the output term computed on each pseudo-register *)
+}.
+
+(** Initial symbolic memory state *)
Definition smem_empty := {| pre:=fun _ _ => True; post:=(fun x => Input x) |}.
Fixpoint exp_term (e: exp) (d old: smem) : term :=
@@ -78,11 +79,12 @@ with list_exp_term (le: list_exp) (d old: smem) : list_term :=
end.
-(** assignment of the symbolic memory *)
+(** assignment of the symbolic memory state *)
Definition smem_set (d:smem) x (t:term) :=
{| pre:=(fun ge m => (term_eval ge (d x) m) <> None /\ (d.(pre) ge m));
post:=fun y => if R.eq_dec x y then t else d y |}.
+(** Simulation theory: the theorem [bblock_smem_simu] ensures that the simulation between two abstract basic blocks is implied by the simulation between their symbolic execution. *)
Section SIMU_THEORY.
Variable ge: genv.
@@ -375,8 +377,7 @@ Qed.
End SIMU_THEORY.
-(** REMARKS: more abstract formulation of the proof...
- but relying on functional_extensionality.
+(** REMARK: this theorem reformulates the lemma above in a more abstract way (but relies on functional_extensionality axiom).
*)
Definition smem_correct ge (d: smem) (m: mem) (om: option mem): Prop:=
forall m', om=Some m' <-> (d.(pre) ge m /\ forall x, term_eval ge (d x) m = Some (m' x)).