aboutsummaryrefslogtreecommitdiff
path: root/src/Util/WordUtil.v
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2017-01-21 18:46:03 -0500
committerGravatar Jason Gross <jgross@mit.edu>2017-01-21 18:46:03 -0500
commit5597e03e4db81700bb0774a2986cd22e45810409 (patch)
treea456963d204417358eadd8e483961f3ac2ce8781 /src/Util/WordUtil.v
parent7754c805fe7e956560bd5747f0324e22dc5b7bb7 (diff)
Add weqb_hetero
Diffstat (limited to 'src/Util/WordUtil.v')
-rw-r--r--src/Util/WordUtil.v34
1 files changed, 33 insertions, 1 deletions
diff --git a/src/Util/WordUtil.v b/src/Util/WordUtil.v
index acf27e225..7f8d7b824 100644
--- a/src/Util/WordUtil.v
+++ b/src/Util/WordUtil.v
@@ -7,6 +7,7 @@ Require Import Coq.Numbers.Natural.Peano.NPeano Util.NatUtil.
Require Import Coq.micromega.Psatz.
Require Import Coq.Bool.Bool.
+Require Import Crypto.Util.Bool.
Require Import Crypto.Util.NatUtil.
Require Import Crypto.Util.Tactics.
@@ -21,6 +22,38 @@ Create HintDb push_wordToN discriminated.
Hint Extern 1 => autorewrite with pull_wordToN in * : pull_wordToN.
Hint Extern 1 => autorewrite with push_wordToN in * : push_wordToN.
+Module Word.
+ Fixpoint weqb_hetero sz1 sz2 (x : word sz1) (y : word sz2) : bool :=
+ match x, y with
+ | WO, WO => true
+ | WO, _ => false
+ | WS b _ x', WS b' _ y'
+ => eqb b b' && @weqb_hetero _ _ x' y'
+ | WS _ _ _, _
+ => false
+ end%bool.
+ Global Arguments weqb_hetero {_ _} _ _.
+ Theorem weqb_hetero_true_iff : forall sz1 x sz2 y,
+ @weqb_hetero sz1 sz2 x y = true <-> exists pf : sz1 = sz2, eq_rect _ word x _ pf = y.
+ Proof.
+ induction x; intros sz2 y; (split; [ destruct y | ]);
+ repeat first [ progress simpl
+ | intro
+ | congruence
+ | exists eq_refl
+ | progress destruct_head ex
+ | progress destruct_head bool
+ | progress subst
+ | progress split_andb
+ | match goal with
+ | [ IHx : forall sz2 y, weqb_hetero _ y = true <-> _, H : weqb_hetero _ _ = true |- _ ]
+ => apply IHx in H
+ | [ IHx : forall sz2 y, weqb_hetero _ y = true <-> _ |- weqb_hetero ?x ?x = true ]
+ => apply IHx
+ end ].
+ Qed.
+End Word.
+
(* Utility Tactics *)
Ltac word_util_arith := omega.
@@ -1211,4 +1244,3 @@ Lemma wlast_combine : forall sz a b, @wlast sz (combine a (WS b WO)) = b.
Proof.
intros; unfold wlast; rewrite split2_combine; cbv; reflexivity.
Qed.
-