From 1d32156c5bc43f966abe2c12c317a27e092355c4 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Fri, 7 May 2021 17:30:43 +0200 Subject: idea of a new scheme to define the semantics of LOAD NOTRAP --- scheduling/BTL.v | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) (limited to 'scheduling/BTL.v') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index fb67d150..8ccab024 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') -- cgit From a1cde0686dfb40595423f40ccc40e18de3539e52 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 8 May 2021 07:31:05 +0200 Subject: idee pour simplifier la preuve: restreindre le "right_assoc" en "expand" ! --- scheduling/BTL.v | 66 +++++++++++++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 61 insertions(+), 5 deletions(-) (limited to 'scheduling/BTL.v') diff --git a/scheduling/BTL.v b/scheduling/BTL.v index 8ccab024..7f3d327b 100644 --- a/scheduling/BTL.v +++ b/scheduling/BTL.v @@ -785,11 +785,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 @@ -797,6 +792,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: -- cgit