Module NeedOp


Require Import Coqlib.
Require Import AST Integers Floats.
Require Import Values Memory Globalenvs.
Require Import Op RTL.
Require Import NeedDomain.

Neededness analysis for RISC-V operators

Definition op1 (nv: nval) := nv :: nil.
Definition op2 (nv: nval) := nv :: nv :: nil.

Definition needs_of_condition (cond: condition): list nval := nil.

Definition needs_of_operation (op: operation) (nv: nval): list nval :=
  match op with
  | Omove => op1 nv
  | Ocopy => op2 (default nv)
  | Ocopyimm _ => op1 (default nv)
  | Owelldef _ => op1 (default nv)
  | Ointconst n => nil
  | Olongconst n => nil
  | Ofloatconst n => nil
  | Osingleconst n => nil
  | Oaddrsymbol id ofs => nil
  | Oaddrstack ofs => nil
  | Ocast8signed => op1 (sign_ext 8 nv)
  | Ocast16signed => op1 (sign_ext 16 nv)
  | Oadd => op2 (modarith nv)
  | Oaddimm n => op1 (modarith nv)
  | Oneg => op1 (modarith nv)
  | Osub => op2 (default nv)
  | Omul => op2 (modarith nv)
  | Omulhs | Omulhu | Odiv | Odivu | Omod | Omodu => op2 (default nv)
  | Oand => op2 (bitwise nv)
  | Oandimm n => op1 (andimm nv n)
  | Oor => op2 (bitwise nv)
  | Oorimm n => op1 (orimm nv n)
  | Oxor => op2 (bitwise nv)
  | Oxorimm n => op1 (bitwise nv)
  | Oshl | Oshr | Oshru => op2 (default nv)
  | Oshlimm n => op1 (shlimm nv n)
  | Oshrimm n => op1 (shrimm nv n)
  | Oshruimm n => op1 (shruimm nv n)
  | Oshrximm n => op1 (default nv)
  | Omakelong => op2 (default nv)
  | Olowlong | Ohighlong => op1 (default nv)
  | Ocast32signed => op1 (default nv)
  | Ocast32unsigned => op1 (default nv)
  | Oaddl => op2 (default nv)
  | Oaddlimm n => op1 (default nv)
  | Onegl => op1 (default nv)
  | Osubl => op2 (default nv)
  | Omull => op2 (default nv)
  | Omullhs | Omullhu | Odivl | Odivlu | Omodl | Omodlu => op2 (default nv)
  | Oandl => op2 (default nv)
  | Oandlimm n => op1 (default nv)
  | Oorl => op2 (default nv)
  | Oorlimm n => op1 (default nv)
  | Oxorl => op2 (default nv)
  | Oxorlimm n => op1 (default nv)
  | Oshll | Oshrl | Oshrlu => op2 (default nv)
  | Oshllimm n => op1 (default nv)
  | Oshrlimm n => op1 (default nv)
  | Oshrluimm n => op1 (default nv)
  | Oshrxlimm n => op1 (default nv)
  | Onegf | Oabsf | Osqrtf => op1 (default nv)
  | Oaddf | Osubf | Omulf | Odivf => op2 (default nv)
  | Onegfs | Oabsfs | Osqrtfs => op1 (default nv)
  | Oaddfs | Osubfs | Omulfs | Odivfs => op2 (default nv)
  | Ofloatofsingle | Osingleoffloat => op1 (default nv)
  | Ointoffloat | Ointuoffloat | Ofloatofint | Ofloatofintu => op1 (default nv)
  | Olongoffloat | Olonguoffloat | Ofloatoflong | Ofloatoflongu => op1 (default nv)
  | Ointofsingle | Ointuofsingle | Osingleofint | Osingleofintu => op1 (default nv)
  | Olongofsingle | Olonguofsingle | Osingleoflong | Osingleoflongu => op1 (default nv)
  | Ocmp c => needs_of_condition c
  | OEseqw _ => op2 (default nv)
  | OEsnew _ => op2 (default nv)
  | OEsequw _ => op2 (default nv)
  | OEsneuw _ => op2 (default nv)
  | OEsltw _ => op2 (default nv)
  | OEsltuw _ => op2 (default nv)
  | OEsltiw _ => op1 (default nv)
  | OEsltiuw _ => op1 (default nv)
  | OExoriw _ => op1 (bitwise nv)
  | OEluiw _ => op1 (default nv)
  | OEaddiw _ _ => op1 (default nv)
  | OEandiw n => op1 (andimm nv n)
  | OEoriw n => op1 (orimm nv n)
  | OEseql _ => op2 (default nv)
  | OEsnel _ => op2 (default nv)
  | OEsequl _ => op2 (default nv)
  | OEsneul _ => op2 (default nv)
  | OEsltl _ => op2 (default nv)
  | OEsltul _ => op2 (default nv)
  | OEsltil _ => op1 (default nv)
  | OEsltiul _ => op1 (default nv)
  | OExoril _ => op1 (default nv)
  | OEluil _ => op1 (default nv)
  | OEaddil _ _ => op1 (default nv)
  | OEandil _ => op1 (default nv)
  | OEoril _ => op1 (default nv)
  | OEloadli _ => op1 (default nv)
  | OEmayundef _ => op2 (default nv)
  | OEfeqd => op2 (default nv)
  | OEfltd => op2 (default nv)
  | OEfled => op2 (default nv)
  | OEfeqs => op2 (default nv)
  | OEflts => op2 (default nv)
  | OEfles => op2 (default nv)
  | Obits_of_single => op1 (default nv)
  | Obits_of_float => op1 (default nv)
  | Osingle_of_bits => op1 (default nv)
  | Ofloat_of_bits => op1 (default nv)
  | Oselectl => All :: nv :: nv :: nil
  | Ogetcanary => nil
  end.

Definition operation_is_redundant (op: operation) (nv: nval): bool :=
  match op with
  | Ocast8signed => sign_ext_redundant 8 nv
  | Ocast16signed => sign_ext_redundant 16 nv
  | Oandimm n => andimm_redundant nv n
  | Oorimm n => orimm_redundant nv n
  | _ => false
  end.

Ltac InvAgree :=
  match goal with
  | [H: vagree_list nil _ _ |- _ ] => inv H; InvAgree
  | [H: vagree_list (_::_) _ _ |- _ ] => inv H; InvAgree
  | _ => idtac
  end.

Ltac TrivialExists :=
  match goal with
  | [ |- exists v, Some ?x = Some v /\ _ ] => exists x; split; auto
  | _ => idtac
  end.

Section SOUNDNESS.

Variable ge: genv.
Variable sp: block.
Variables m m': mem.
Hypothesis PERM: forall b ofs k p, Mem.perm m b ofs k p -> Mem.perm m' b ofs k p.

Lemma needs_of_condition_sound:
  forall cond args b args',
  eval_condition cond args m = Some b ->
  vagree_list args args' (needs_of_condition cond) ->
  eval_condition cond args' m' = Some b.
Proof.
  intros. unfold needs_of_condition in H0.
  eapply default_needs_of_condition_sound; eauto.
Qed.

Lemma needs_of_operation_sound:
  forall op args v nv args',
  eval_operation ge (Vptr sp Ptrofs.zero) op args m = Some v ->
  vagree_list args args' (needs_of_operation op nv) ->
  nv <> Nothing ->
  exists v',
     eval_operation ge (Vptr sp Ptrofs.zero) op args' m' = Some v'
  /\ vagree v v' nv.
Proof.
  unfold needs_of_operation; intros; destruct op; try (eapply default_needs_of_operation_sound; eauto; fail);
  simpl in *; FuncInv; InvAgree; TrivialExists.
- apply sign_ext_sound; auto. compute; auto.
- apply sign_ext_sound; auto. compute; auto.
- apply add_sound; auto.
- apply add_sound; auto with na.
- apply neg_sound; auto.
- apply mul_sound; auto.
- apply and_sound; auto.
- apply andimm_sound; auto.
- apply or_sound; auto.
- apply orimm_sound; auto.
- apply xor_sound; auto.
- apply xor_sound; auto with na.
- apply shlimm_sound; auto.
- apply shrimm_sound; auto.
- apply shruimm_sound; auto.
- fold (Val.and (Vint n) v0);
  fold (Val.and (Vint n) v2);
  rewrite (Val.and_commut (Vint n) v0);
  rewrite (Val.and_commut (Vint n) v2);
  apply andimm_sound; auto.
- fold (Val.or (Vint n) v0);
  fold (Val.or (Vint n) v2);
  rewrite (Val.or_commut (Vint n) v0);
  rewrite (Val.or_commut (Vint n) v2);
  apply orimm_sound; auto.
- apply xor_sound; auto with na.
- (* selectl *)
  unfold ExtValues.select01_long.
  destruct v0; auto with na.
  assert (Val.lessdef (Vint i) v4) as LESSDEF by auto with na.
  inv LESSDEF.
  destruct (Int.eq i Int.one).
  { apply normalize_sound; auto. }
  destruct (Int.eq i Int.zero).
  { apply normalize_sound; auto. }
  cbn. auto with na.
Qed.

Lemma operation_is_redundant_sound:
  forall op nv arg1 args v arg1' args',
  operation_is_redundant op nv = true ->
  eval_operation ge (Vptr sp Ptrofs.zero) op (arg1 :: args) m = Some v ->
  vagree_list (arg1 :: args) (arg1' :: args') (needs_of_operation op nv) ->
  vagree v arg1' nv.
Proof.
  intros. destruct op; simpl in *; try discriminate; inv H1; FuncInv; subst.
- apply sign_ext_redundant_sound; auto. lia.
- apply sign_ext_redundant_sound; auto. lia.
- apply andimm_redundant_sound; auto.
- apply orimm_redundant_sound; auto.
Qed.

End SOUNDNESS.