aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling/BTL.v
diff options
context:
space:
mode:
authorLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-05-09 22:41:48 +0200
committerLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-05-09 22:41:48 +0200
commit3043510dc2bdaa5d151656a667f1f7988689a75a (patch)
treedc7bf030c9811bccbf9d7dc142b34bfe7ff00161 /scheduling/BTL.v
parentc78393a5d4d13d0f2cd7f0a73756d6bb598b5ae4 (diff)
parentd422c63cbcad7ba156d5d324e0221db9d13f9559 (diff)
downloadcompcert-kvx-3043510dc2bdaa5d151656a667f1f7988689a75a.tar.gz
compcert-kvx-3043510dc2bdaa5d151656a667f1f7988689a75a.zip
Merge branch 'BTL' of gricad-gitlab.univ-grenoble-alpes.fr:sixcy/CompCert into BTL
Diffstat (limited to 'scheduling/BTL.v')
-rw-r--r--scheduling/BTL.v88
1 files changed, 83 insertions, 5 deletions
diff --git a/scheduling/BTL.v b/scheduling/BTL.v
index c8dfa0af..d7ffa118 100644
--- a/scheduling/BTL.v
+++ b/scheduling/BTL.v
@@ -175,6 +175,23 @@ Definition find_function (ros: reg + ident) (rs: regset) : option fundef :=
Local Open Scope option_monad_scope.
+(* (* TODO: a new (hopefully simpler) scheme to support "NOTRAP" wrt current scheme of RTL *)
+
+Inductive has_loaded sp rs m chunk addr args v: trapping_mode -> Prop :=
+ | has_loaded_normal a trap
+ (EVAL: eval_addressing ge sp addr rs##args = Some a)
+ (LOAD: Mem.loadv chunk m a = Some v)
+ : has_loaded sp rs m chunk addr args v trap
+ | has_loaded_default
+ (LOAD: forall a, eval_addressing ge sp addr rs##args = Some a -> Mem.loadv chunk m a = None)
+ (DEFAULT: v = default_notrap_load_value chunk)
+ : has_loaded sp rs m chunk addr args v NOTRAP
+ .
+
+(* TODO: move this scheme in "Memory" module if this scheme is useful ! *)
+
+*)
+
(** internal big-step execution of one iblock *)
Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option final -> Prop :=
| exec_final rs m fin: iblock_istep sp rs m (BF fin) rs m (Some fin)
@@ -186,6 +203,11 @@ Inductive iblock_istep sp: regset -> mem -> iblock -> regset -> mem -> option fi
(EVAL: eval_addressing ge sp addr rs##args = Some a)
(LOAD: Mem.loadv chunk m a = Some v)
: iblock_istep sp rs m (Bload TRAP chunk addr args dst) (rs#dst <- v) m None
+(* TODO: replace [exec_load_TRAP] above by this one. See "new scheme" for "NOTRAP" above
+ | exec_load rs m trap chunk addr args dst v
+ (LOAD: has_loaded sp rs m chunk addr args v trap)
+ : iblock_istep sp rs m (Bload trap chunk addr args dst) (rs#dst <- v) m None
+*)
| exec_store rs m chunk addr args src a m'
(EVAL: eval_addressing ge sp addr rs##args = Some a)
(STORE: Mem.storev chunk m a rs#src = Some m')
@@ -764,11 +786,6 @@ Definition verify_function dupmap f f' : res unit :=
do _ <- verify_is_copy dupmap (fn_entrypoint f) (RTL.fn_entrypoint f');
verify_cfg dupmap (fn_code f) (RTL.fn_code f').
-Definition is_seq (ib: iblock): bool :=
- match ib with
- | Bseq _ _ => true
- | _ => false
- end.
Definition is_atom (ib: iblock): bool :=
match ib with
@@ -776,6 +793,67 @@ Definition is_atom (ib: iblock): bool :=
| _ => true
end.
+Inductive is_expand: iblock -> Prop :=
+ | exp_Bseq ib1 ib2:
+ is_atom ib1 = true ->
+ is_expand ib2 ->
+ is_expand (Bseq ib1 ib2)
+ | exp_Bcond cond args ib1 ib2 i:
+ is_expand ib1 ->
+ is_expand ib2 ->
+ is_expand (Bcond cond args ib1 ib2 i)
+ | exp_others ib:
+ is_atom ib = true ->
+ is_expand ib
+ .
+Local Hint Constructors is_expand: core.
+
+Fixpoint expand (ib: iblock) (k: option iblock): iblock :=
+ match ib with
+ | Bseq ib1 ib2 => expand ib1 (Some (expand ib2 k))
+ | Bcond cond args ib1 ib2 i =>
+ Bcond cond args (expand ib1 k) (expand ib2 k) i
+ | ib =>
+ match k with
+ | None => ib
+ | Some rem => Bseq ib rem
+ end
+ end.
+
+Lemma expand_correct ib: forall k,
+ (match k with Some rem => is_expand rem | None => True end)
+ -> is_expand (expand ib k).
+Proof.
+ induction ib; simpl; intros; try autodestruct; auto.
+Qed.
+
+Lemma expand_None_correct ib:
+ is_expand (expand ib None).
+Proof.
+ apply expand_correct; simpl; auto.
+Qed.
+
+Definition expand_code (cfg: code): code :=
+ (PTree.map (fun _ ib => {| entry:=expand ib.(entry) None; input_regs := ib.(input_regs) |}) cfg).
+
+Lemma expand_code_correct cfg pc ib :
+ (expand_code cfg)!pc=Some ib -> is_expand (ib.(entry)).
+Proof.
+ unfold expand_code.
+ rewrite PTree.gmap.
+ destruct (cfg!pc); simpl; intros; try_simplify_someHyps.
+ apply expand_None_correct; auto.
+Qed.
+
+
+(* TODO: eliminer toute la suite quand "right_assoc" sera remplace par "expand" *)
+
+Definition is_seq (ib: iblock): bool :=
+ match ib with
+ | Bseq _ _ => true
+ | _ => false
+ end.
+
Inductive is_right_assoc: iblock -> Prop :=
| ira_Bseq ib1 ib2: