From 766e4bd7ae4f299dcd159216813062a93c58c2c3 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Mon, 13 Aug 2018 19:20:04 -0400 Subject: Move a lemma --- src/Experiments/NewPipeline/Toplevel1.v | 13 +------------ src/Util/ListUtil/FoldBool.v | 12 ++++++++++++ 2 files changed, 13 insertions(+), 12 deletions(-) (limited to 'src') diff --git a/src/Experiments/NewPipeline/Toplevel1.v b/src/Experiments/NewPipeline/Toplevel1.v index 727426082..869abea8e 100644 --- a/src/Experiments/NewPipeline/Toplevel1.v +++ b/src/Experiments/NewPipeline/Toplevel1.v @@ -1750,18 +1750,7 @@ Module Import UnsaturatedSolinas. = type.app_curried (t:=to_bytesT) (freeze_to_bytesmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n machine_wordsize m_enc) f)) /\ (Positional.eval (weight 8 1) n_bytes (type.app_curried (t:=to_bytesT) (freeze_to_bytesmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n machine_wordsize m_enc) f)) = (Positional.eval (weight (Qnum limbwidth) (Z.pos (Qden limbwidth))) n (fst f) mod m))). - (** XXX TODO MOVE ME *) - Lemma fold_andb_map_snoc A B f x xs y ys - : @fold_andb_map A B f (xs ++ [x]) (ys ++ [y]) = @fold_andb_map A B f xs ys && f x y. - Proof. - clear. - revert ys; induction xs as [|x' xs'], ys as [|y' ys']; cbn; - rewrite ?Bool.andb_true_r, ?Bool.andb_false_r; - try (destruct ys'; cbn; rewrite Bool.andb_false_r); - try (destruct xs'; cbn; rewrite Bool.andb_false_r); - try reflexivity. - rewrite IHxs', Bool.andb_assoc; reflexivity. - Qed. + (** XXX TODO MOVE ME? *) Lemma eval_is_bounded_by wt n' bounds bounds' f (H : ZRange.type.base.option.is_bounded_by (t:=base.type.list base.type.Z) bounds f = true) (Hb : bounds = Some (List.map (@Some _) bounds')) diff --git a/src/Util/ListUtil/FoldBool.v b/src/Util/ListUtil/FoldBool.v index b69081cab..f6cac6118 100644 --- a/src/Util/ListUtil/FoldBool.v +++ b/src/Util/ListUtil/FoldBool.v @@ -55,3 +55,15 @@ Proof. intuition (congruence || (subst; auto)). apply (H1 (_, _)); auto. Qed. + +Lemma fold_andb_map_snoc A B f x xs y ys + : @fold_andb_map A B f (xs ++ [x]) (ys ++ [y]) = @fold_andb_map A B f xs ys && f x y. +Proof. + clear. + revert ys; induction xs as [|x' xs'], ys as [|y' ys']; cbn; + rewrite ?Bool.andb_true_r, ?Bool.andb_false_r; + try (destruct ys'; cbn; rewrite Bool.andb_false_r); + try (destruct xs'; cbn; rewrite Bool.andb_false_r); + try reflexivity. + rewrite IHxs', Bool.andb_assoc; reflexivity. +Qed. -- cgit v1.2.3