From 3830a91a4711c4570394e02e93e4e08db88eac6f Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Fri, 12 Apr 2019 17:35:18 +0200 Subject: Support a "select" operation between two values `Val.select ob v1 v2 ty` is a conditional operation that chooses between the values `v1` and `v2` depending on the comparison `ob : option bool`. If `ob` is `None`, `Vundef` is returned. If the selected value does not match type `ty`, `Vundef` is returned. This operation will be used to model a "select" (or "conditional move") operation at the CminorSel/RTL/LTL/Mach level. --- backend/ValueDomain.v | 58 +++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 58 insertions(+) (limited to 'backend/ValueDomain.v') diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v index 3ba2a35b..f6afa836 100644 --- a/backend/ValueDomain.v +++ b/backend/ValueDomain.v @@ -2824,6 +2824,64 @@ Proof. intros. inv H; simpl in H0; congruence. Qed. +(** Select either returns one of its arguments, or Vundef. *) + +Definition add_undef (x: aval) := + match x with + | Vbot => ntop + | I i => + if Int.lt i Int.zero + then sgn Pbot (ssize i) + else uns Pbot (usize i) + | L _ | F _ | FS _ => ntop + | _ => x + end. + +Lemma add_undef_sound: + forall v x, vmatch v x -> vmatch v (add_undef x). +Proof. + destruct 1; simpl; auto with va. + destruct (Int.lt i Int.zero). + apply vmatch_sgn; apply is_sgn_ssize. + apply vmatch_uns; apply is_uns_usize. +Qed. + +Lemma add_undef_undef: + forall x, vmatch Vundef (add_undef x). +Proof. + destruct x; simpl; auto with va. + destruct (Int.lt n Int.zero); auto with va. +Qed. + +Lemma add_undef_normalize: + forall v x ty, vmatch v x -> vmatch (Val.normalize v ty) (add_undef x). +Proof. + intros. destruct (Val.lessdef_normalize v ty); + auto using add_undef_sound, add_undef_undef. +Qed. + +Definition select (ab: abool) (x y: aval) := + match ab with + | Bnone => ntop + | Just b | Maybe b => add_undef (if b then x else y) + | Btop => add_undef (vlub x y) + end. + +Lemma select_sound: + forall ob v w ab x y ty, + cmatch ob ab -> vmatch v x -> vmatch w y -> + vmatch (Val.select ob v w ty) (select ab x y). +Proof. + unfold Val.select, select; intros. inv H. +- auto with va. +- apply add_undef_normalize; destruct b; auto. +- apply add_undef_undef. +- apply add_undef_normalize; destruct b; auto. +- destruct ob as [b|]. ++ apply add_undef_normalize. destruct b; [apply vmatch_lub_l|apply vmatch_lub_r]; auto. ++ apply add_undef_undef. +Qed. + (** Normalization at load time *) Definition vnormalize (chunk: memory_chunk) (v: aval) := -- cgit