From 0eb1e1e6dd80154fb9841697b2c19482ece2f7da Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 30 Jan 2021 15:52:21 +0100 Subject: select through bitwise operations --- riscV/ExtValues.v | 40 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 40 insertions(+) create mode 100644 riscV/ExtValues.v (limited to 'riscV/ExtValues.v') diff --git a/riscV/ExtValues.v b/riscV/ExtValues.v new file mode 100644 index 00000000..0f29e837 --- /dev/null +++ b/riscV/ExtValues.v @@ -0,0 +1,40 @@ +Require Import Coqlib. +Require Import Integers. +Require Import Values. +Require Import Floats. + +Definition bitwise_select_int b vtrue vfalse := + Val.or (Val.and (Val.neg b) vtrue) + (Val.and (Val.sub b Vone) vfalse). + +Lemma bitwise_select_int_true : + forall vtrue vfalse, + bitwise_select_int (Val.of_optbool (Some true)) (Vint vtrue) (Vint vfalse) + = Vint vtrue. +Proof. + intros. cbn. f_equal. + change (Int.neg Int.one) with Int.mone. + rewrite Int.and_commut. + rewrite Int.and_mone. + rewrite Int.sub_idem. + rewrite Int.and_commut. + rewrite Int.and_zero. + apply Int.or_zero. +Qed. + +Lemma bitwise_select_int_false : + forall vtrue vfalse, + bitwise_select_int (Val.of_optbool (Some false)) (Vint vtrue) (Vint vfalse) + = Vint vfalse. +Proof. + intros. cbn. f_equal. + rewrite Int.neg_zero. + rewrite Int.and_commut. + rewrite Int.and_zero. + rewrite Int.sub_zero_r. + change (Int.neg Int.one) with Int.mone. + rewrite Int.and_commut. + rewrite Int.and_mone. + rewrite Int.or_commut. + apply Int.or_zero. +Qed. -- cgit