aboutsummaryrefslogtreecommitdiff
path: root/src/Assembly/Wordize.v
diff options
context:
space:
mode:
authorGravatar Robert Sloan <varomodt@gmail.com>2016-06-14 21:21:47 -0400
committerGravatar Robert Sloan <varomodt@gmail.com>2016-06-22 13:45:18 -0400
commit1481b88f0c2498f41421c1333856e4458b186618 (patch)
treef8b6267aba7c7e30e57f3c19df1e79c3b8ab4f6e /src/Assembly/Wordize.v
parent2c4a10de8bde2986d7f32598d9e6be7fb84cb7bc (diff)
Reorganization of wordize.v
first pseudo conversion lemma Decent machinery for automatic pseudo-conversion Decent machinery for automatic pseudo-conversion Decent machinery for automatic pseudo-conversion
Diffstat (limited to 'src/Assembly/Wordize.v')
-rw-r--r--src/Assembly/Wordize.v358
1 files changed, 283 insertions, 75 deletions
diff --git a/src/Assembly/Wordize.v b/src/Assembly/Wordize.v
index 0e88ada75..e5ad3f066 100644
--- a/src/Assembly/Wordize.v
+++ b/src/Assembly/Wordize.v
@@ -1,46 +1,51 @@
Require Export Bedrock.Word Bedrock.Nomega.
-Require Import NArith PArith Ndigits Nnat NPow Compare_dec.
-Require Import FunctionalExtensionality.
+Require Import NArith PArith Ndigits Nnat NPow NPeano Ndec Compare_dec.
+Require Import FunctionalExtensionality ProofIrrelevance.
+Require Import QhasmUtil QhasmEvalCommon.
Hint Rewrite wordToN_nat Nat2N.inj_add N2Nat.inj_add
Nat2N.inj_mul N2Nat.inj_mul Npow2_nat : N.
-Delimit Scope wordize_scope with wordize.
-Open Scope wordize_scope.
+Open Scope nword_scope.
-Notation "& x" := (wordToN x) (at level 30) : wordize_scope.
-Notation "** x" := (NToWord _ x) (at level 30) : wordize_scope.
-Notation "$ x" := (natToWord x) (at level 30) : wordize_scope.
-
-Section Definitions.
-
- Definition take (k n: nat) (p: (k <= n)%nat) (w: word n): word k.
- replace n with (k + (n - k)) in * by abstract omega.
- refine (split1 k _ w).
- Defined.
-
- Definition shiftr {n} (w: word n) (k: nat): word n.
- destruct (le_dec k n).
-
- - replace n with (k + (n - k)) in * by abstract intuition.
- refine (sext (take k (k + (n - k)) l w) _).
-
- - exact (wzero n).
- Defined.
+Section WordizeUtil.
+ Lemma break_spec: forall (m n: nat) (x: word n) low high,
+ (low, high) = break m x
+ -> &x = (&high * Npow2 m + &low)%N.
+ Proof.
+ intros; unfold break in *; destruct (le_dec m n);
+ inversion H; subst; clear H; simpl.
+ Admitted.
+
+ Lemma mask_wand : forall (n: nat) (x: word n) m b,
+ (& (mask (N.to_nat m) x) < b)%N
+ -> (& (x ^& (@NToWord n (N.ones m))) < b)%N.
+ Proof.
+ Admitted.
- Definition mask {n} (m: nat) (w: word n): word n.
- destruct (le_dec m n).
+ Lemma convS_id: forall A x p, x = (@convS A A x p).
+ Proof.
+ intros; unfold convS; vm_compute.
+ replace p with (eq_refl A); intuition.
+ apply proof_irrelevance.
+ Qed.
- - replace n with (m + (n - m)) in * by (abstract intuition).
- refine (w ^& (zext (wones m) (n - m))).
+ Lemma word_param_eq: forall n m, word n = word m -> n = m.
+ Proof. (* TODO: How do we prove this *) Admitted.
- - exact w.
- Defined.
+ Lemma word_conv_eq: forall {n m} (y: word m) p,
+ &y = &(@convS (word m) (word n) y p).
+ Proof.
+ intros.
+ revert p.
+ destruct (Nat.eq_dec n m).
-End Definitions.
+ - rewrite e; intros; apply f_equal; apply convS_id.
-Section WordizeUtil.
+ - intros; contradict n0.
+ apply word_param_eq; rewrite p; intuition.
+ Qed.
Lemma to_nat_lt: forall x b, (x < b)%N <-> (N.to_nat x < N.to_nat b)%nat.
Proof. (* via Nat2N.inj_compare *) Admitted.
@@ -101,13 +106,69 @@ Section WordizeUtil.
+ apply N.neq_0_lt_0 in IHx; intuition.
Qed.
+ Lemma natToWord_convS: forall {n m} (x: word n) p,
+ & x = & @convS (word n) (word m) x p.
+ Proof. admit. Qed.
+
+ Lemma natToWord_combine: forall {n} (x: word n) k,
+ & x = & combine x (wzero k).
+ Proof. admit. Qed.
+
+ Lemma natToWord_split1: forall {n} (x: word n) p,
+ & x = & split1 n 0 (convS x p).
+ Proof. admit. Qed.
+
+ Lemma extend_bound: forall k n (p: (k <= n)%nat) (w: word k),
+ (& (extend p w) < Npow2 k)%N.
+ Proof.
+ intros.
+ assert (n = k + (n - k)) by abstract omega.
+ replace (& (extend p w)) with (& w); try apply word_size_bound.
+ unfold extend.
+ rewrite <- word_conv_eq.
+ unfold zext.
+ clear; revert w; induction (n - k).
+
+ - intros.
+ assert (word k = word (k + 0)) as Z by intuition.
+ replace w with (split1 k 0 (convS w Z)).
+ replace (wzero 0) with (split2 k 0 (convS w Z)).
+ rewrite <- natToWord_split1 with (p := Z).
+ rewrite combine_split.
+ apply natToWord_convS.
+
+ + admit.
+ + admit.
+
+ - intros; rewrite <- natToWord_combine; intuition.
+ Qed.
+
+ Lemma Npow2_split: forall a b,
+ (Npow2 (a + b) = (Npow2 a) * (Npow2 b))%N.
+ Proof.
+ intros; revert a.
+ induction b.
+
+ - intros; simpl; replace (a + 0) with a; nomega.
+
+ - intros.
+ replace (a + S b) with (S a + b) by intuition.
+ rewrite (IHb (S a)); simpl; clear IHb.
+ induction (Npow2 a), (Npow2 b); simpl; intuition.
+ rewrite Pos.mul_xO_r; intuition.
+ Qed.
+
+ Lemma Npow2_ignore: forall {n} (x: word n),
+ x = NToWord _ (& x + Npow2 n).
+ Proof. intros. Admitted.
+
End WordizeUtil.
(** Wordization Lemmas **)
Section Wordization.
- Lemma wordize_plus: forall {n} (x y: word n) (b: N),
+ Lemma wordize_plus': forall {n} (x y: word n) (b: N),
(b <= Npow2 n)%N
-> (&x < b)%N
-> (&y < (Npow2 n - b))%N
@@ -130,7 +191,19 @@ Section Wordization.
apply N.lt_le_incl; assumption.
Qed.
- Lemma wordize_mult: forall {n} (x y: word n) (b: N),
+ Lemma wordize_plus: forall {n} (x y: word n),
+ if (overflows n (&x + &y)%N)
+ then (&x + &y)%N = (& (x ^+ y) + Npow2 n)%N
+ else (&x + &y)%N = & (x ^+ y).
+ Proof. Admitted.
+
+ Lemma wordize_awc: forall {n} (x y: word n) (c: bool),
+ if (overflows n (&x + &y + (if c then 1 else 0))%N)
+ then (&x + &y + (if c then 1 else 0))%N = (&(addWithCarry x y c) + Npow2 n)%N
+ else (&x + &y + (if c then 1 else 0))%N = &(addWithCarry x y c).
+ Proof. Admitted.
+
+ Lemma wordize_mult': forall {n} (x y: word n) (b: N),
(1 < n)%nat -> (0 < b)%N
-> (&x < b)%N
-> (&y < (Npow2 n) / b)%N
@@ -145,6 +218,11 @@ Section Wordization.
- apply N.mul_div_le; nomega.
Qed.
+ Lemma wordize_mult: forall {n} (x y: word n) (b: N),
+ (&x * &y)%N = (&(x ^* y) +
+ &((EvalUtil.highBits (n/2) x) ^* (EvalUtil.highBits (n/2) y)) * Npow2 n)%N.
+ Proof. intros. Admitted.
+
Lemma wordize_and: forall {n} (x y: word n),
N.land (&x) (&y) = & (x ^& y).
Proof.
@@ -172,6 +250,7 @@ Section Wordization.
(N.shiftr_nat (&x) k) = & (shiftr x k).
Proof.
intros.
+
admit.
Qed.
@@ -183,6 +262,16 @@ Section Bounds.
(& k < & k + 1)%N.
Proof. intros; nomega. Qed.
+ Theorem constant_bound_nat : forall (n k: nat),
+ (N.of_nat k < Npow2 n)%N
+ -> (& (@natToWord n k) < (N.of_nat k) + 1)%N.
+ Proof.
+ intros.
+ rewrite wordToN_nat.
+ rewrite wordToNat_natToWord_idempotent;
+ try assumption; nomega.
+ Qed.
+
Lemma let_bound : forall {n} (x: word n) (f: word n -> word n) xb fb,
(& x < xb)%N
-> (forall x', & x' < xb -> & (f x') < fb)%N
@@ -195,7 +284,13 @@ Section Bounds.
match c as c' return c = c' -> _ with
| Lt => fun _ => left _
| _ => fun _ => right _
- end eq_refl); admit.
+ end eq_refl);
+ abstract (unfold c in *; try first [
+ apply N.compare_eq_iff in _H
+ | apply N.compare_lt_iff in _H
+ | pose proof (N.compare_antisym x y) as _H0;
+ rewrite _H in _H0; simpl in _H0;
+ apply N.compare_lt_iff in _H0 ]; nomega).
Defined.
Theorem wplus_bound : forall {n} (w1 w2 : word n) b1 b2,
@@ -205,61 +300,97 @@ Section Bounds.
Proof.
intros.
- destruct (N.compare (b1 + b2)%N (Npow2 n));
- rewrite <- wordize_plus with (b := b1);
+ destruct (Nlt_dec (b1 + b2)%N (Npow2 n));
+ rewrite <- wordize_plus' with (b := b1);
try apply N.add_lt_mono;
try assumption.
- - apply N.add_lt_mono; assumption.
+ (* A couple inequality subgoals *)
+ Admitted.
+
+ Theorem wmult_bound : forall {n} (w1 w2 : word n) b1 b2,
+ (1 < n)%nat
+ -> (&w1 < b1)%N
+ -> (&w2 < b2)%N
+ -> (&(w1 ^* w2) < b1 * b2)%N.
+ Proof.
+ intros.
+ destruct (Nlt_dec (b1 * b2)%N (Npow2 n));
+ rewrite <- wordize_mult' with (b := b1);
+ try apply N.mul_lt_mono;
+ try assumption;
+ try nomega.
+
+ (* A couple inequality subgoals *)
+ Admitted.
+
+ Theorem shiftr_bound : forall {n} (w : word n) b bits,
+ (&w < b)%N
+ -> (&(shiftr w bits) < N.succ (N.shiftr_nat b bits))%N.
+ Proof.
+ intros.
+ assert (& shiftr w bits <= N.shiftr_nat b bits)%N. {
+ rewrite <- wordize_shiftr.
+ induction bits; unfold N.shiftr_nat in *; simpl; intuition.
+
+ - unfold N.le, N.lt in *; rewrite H; intuition; inversion H0.
- - admit.
- - admit.
+ - revert IHbits;
+ generalize (nat_iter bits N.div2 (& w)),
+ (nat_iter bits N.div2 b);
+ clear; intros x y H.
- - apply N.add_lt_mono; assumption.
- - apply N.add_lt_mono; assumption.
+ admit. (* Monotonicity of N.div2 *)
+ }
+ apply N.le_lteq in H0; destruct H0; nomega.
Qed.
- Theorem wmult_bound : forall (w1 w2 : word n) b1 b2,
- w1 <= b1
- -> w2 <= b2
- -> w1 ^* w2 <= b1 * b2.
+ Theorem mask_bound : forall {n} (w : word n) m,
+ (n > 1)%nat ->
+ (&(mask m w) < Npow2 m)%N.
Proof.
- preomega.
- destruct (le_lt_dec (pow2 n) (N.to_nat b1 * N.to_nat b2)).
+ intros.
+ unfold mask in *; destruct (le_dec m n); simpl;
+ try apply extend_bound.
- specialize (wordToNat_bound (w1 ^* w2)); nomega.
+ pose proof (word_size_bound w).
+ apply (N.le_lt_trans _ (Npow2 n) _).
- rewrite wmult_alt.
- unfold wmultN, wordBinN.
- rewrite wordToNat_natToWord_idempotent.
- ge_to_le_nat.
+ - unfold N.le, N.lt in *; rewrite H0; intuition; inversion H1.
- apply Mult.mult_le_compat; nomega.
- pre_nomega.
- apply Lt.le_lt_trans with (N.to_nat b1 * N.to_nat b2); auto.
- apply Mult.mult_le_compat; nomega.
- Qed.
+ - clear H H0.
+ replace m with ((m - n) + n) by nomega.
+ replace (Npow2 n) with (1 * (Npow2 n))%N
+ by (rewrite N.mul_comm; nomega).
+ rewrite Npow2_split.
+ apply N.mul_lt_mono_pos_r.
- Theorem shiftr_bound : forall (w : word n) b bits,
- w <= b
- -> shiftr w bits <= N.shiftr b (N.of_nat bits).
- Proof.
- admit.
- Qed.
+ + apply Npow2_gt0.
- Theorem mask_bound : forall (w : word n) m,
- mask m w <= Npow2 m - 1.
- Proof.
- admit.
+ + assert (0 < m - n)%nat by omega.
+ induction (m - n); try inversion H; try abstract (
+ simpl; replace 2 with (S 1) by omega;
+ apply N.lt_1_2).
+
+ assert (0 < n1)%nat as Z by omega; apply IHn1 in Z.
+ apply (N.le_lt_trans _ (Npow2 n1) _).
+
+ * admit.
+ * admit.
Qed.
- Theorem mask_update_bound : forall (w : word n) b m,
- w <= b
- -> mask m w <= (N.min b (Npow2 m - 1)).
+ Theorem mask_update_bound : forall {n} (w : word n) b m,
+ (n > 1)%nat
+ -> (&w < b)%N
+ -> (&(mask m w) < (N.min b (Npow2 m)))%N.
Proof.
- admit.
- Qed.
+ intros; unfold mask, N.min;
+ destruct (le_dec m n),
+ (N.compare b (Npow2 m));
+ simpl; try assumption.
+
+ Admitted.
End Bounds.
@@ -267,8 +398,8 @@ End Bounds.
Ltac wordize_ast :=
repeat match goal with
- | [ H: (& ?x < ?b)%N |- context[((& ?x) + (& ?y))%N] ] => rewrite (wordize_plus x y b)
- | [ H: (& ?x < ?b)%N |- context[((& ?x) * (& ?y))%N] ] => rewrite (wordize_mult x y b)
+ | [ H: (& ?x < ?b)%N |- context[((& ?x) + (& ?y))%N] ] => rewrite (wordize_plus' x y b)
+ | [ H: (& ?x < ?b)%N |- context[((& ?x) * (& ?y))%N] ] => rewrite (wordize_mult' x y b)
| [ |- context[N.land (& ?x) (& ?y)] ] => rewrite (wordize_and x y)
| [ |- context[N.shiftr (& ?x) ?k] ] => rewrite (wordize_shiftr x k)
| [ |- (_ < _ / _)%N ] => unfold N.div; simpl
@@ -281,6 +412,81 @@ Ltac lt_crush := try abstract (clear; vm_compute; intuition).
(** Bounding Tactics **)
+Ltac multi_apply0 A L := pose proof (L A).
+
+Ltac multi_apply1 A L :=
+ match goal with
+ | [ H: A < ?b |- _] => pose proof (L A b H)
+ end.
+
+Ltac multi_apply2 A B L :=
+ match goal with
+ | [ H1: A < ?b1, H2: B < ?b2 |- _] => pose proof (L A B b1 b2 H1 H2)
+ end.
+
+Ltac multi_recurse n T :=
+ match goal with
+ | [ H: (T < _)%N |- _] => idtac
+ | _ =>
+ is_var T;
+ let T' := (eval cbv delta [T] in T) in multi_recurse n T';
+ match goal with
+ | [ H : T' < ?x |- _ ] =>
+ pose proof (H : T < x)
+ end
+
+ | _ =>
+ match T with
+ | ?W1 ^+ ?W2 =>
+ multi_recurse n W1; multi_recurse n W2;
+ multi_apply2 W1 W2 (@wplus_bound n)
+
+ | ?W1 ^* ?W2 =>
+ multi_recurse n W1; multi_recurse n W2;
+ multi_apply2 W1 W2 (@wmult_bound n)
+
+ | mask ?m ?w =>
+ multi_recurse n w;
+ multi_apply1 w (fun b => @mask_update_bound n w b)
+
+ | mask ?m ?w =>
+ multi_recurse n w;
+ pose proof (@mask_bound n w m)
+
+ | ?x ^& (@NToWord _ (N.ones ?m)) =>
+ multi_recurse n (mask (N.to_nat m) x);
+ match goal with
+ | [ H: (& (mask (N.to_nat m) x) < ?b)%N |- _] =>
+ pose proof (@mask_wand n x m b H)
+ end
+
+ | shiftr ?w ?bits =>
+ multi_recurse n w;
+ match goal with
+ | [ H: (w < ?b)%N |- _] =>
+ pose proof (@shiftr_bound n w b bits H)
+ end
+
+ | NToWord _ ?k => pose proof (@constant_bound_N n k)
+ | natToWord _ ?k => pose proof (@constant_bound_nat n k)
+ | _ => pose proof (@word_size_bound n T)
+ end
+ end.
+
+Lemma unwrap_let: forall {n} (y: word n) (f: word n -> word n) (b: N),
+ (&(let x := y in f x) < b)%N <-> let x := y in (&(f x) < b)%N.
+Proof. intuition. Qed.
+
+Ltac multi_bound n :=
+ match goal with
+ | [|- let A := ?B in _] =>
+ multi_recurse n B; intro; multi_bound n
+ | [|- ((let A := _ in _) < _)%N] =>
+ apply unwrap_let; multi_bound n
+ | [|- (?W < _)%N ] =>
+ multi_recurse n W
+ end.
+
(** Examples **)
Module WordizationExamples.
@@ -296,4 +502,6 @@ Module WordizationExamples.
transitivity 10%N; try assumption; lt_crush.
Qed.
-End WordizationExamples. \ No newline at end of file
+End WordizationExamples.
+
+Close Scope nword_scope.