aboutsummaryrefslogtreecommitdiff
path: root/src
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2016-11-14 21:46:40 -0500
committerGravatar Jason Gross <jgross@mit.edu>2016-11-14 21:46:40 -0500
commit43c5265c24bd1df125f8de00d1f89379a920659a (patch)
treefdb435c27c7fa1fdc22c9ca8cf107e52dc094a5c /src
parent4faedd36b571f6c07eab1e348d1df655e1123eda (diff)
Support for 128-bit words
I haven't found a good way to genericize the proofs of relatedness things, mostly because Modules and functors are annoying.
Diffstat (limited to 'src')
-rw-r--r--src/Experiments/Ed25519Extraction.v30
-rw-r--r--src/Reflection/Z/Interpretations.v812
-rw-r--r--src/Reflection/Z/Interpretations128.v13
-rw-r--r--src/Reflection/Z/Interpretations128/Relations.v (renamed from src/Reflection/Z/Interpretations/Relations.v)193
-rw-r--r--src/Reflection/Z/Interpretations128/RelationsCombinations.v359
-rw-r--r--src/Reflection/Z/Interpretations64.v13
-rw-r--r--src/Reflection/Z/Interpretations64/Relations.v531
-rw-r--r--src/Reflection/Z/Interpretations64/RelationsCombinations.v (renamed from src/Reflection/Z/Interpretations/RelationsCombinations.v)5
-rw-r--r--src/Reflection/Z/InterpretationsGen.v40
-rw-r--r--src/Specific/GF25519Bounded.v26
-rw-r--r--src/Specific/GF25519Reflective.v14
-rw-r--r--src/Specific/GF25519Reflective/Common.v75
-rw-r--r--src/Specific/GF25519Reflective/CommonBinOp.v8
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOp.v8
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOpWireToFE.v8
-rw-r--r--src/SpecificGen/GF2213_32Bounded.v26
-rw-r--r--src/SpecificGen/GF2213_32Reflective.v14
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Common.v75
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonBinOp.v8
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOp.v8
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v8
-rw-r--r--src/SpecificGen/GF2519_32Bounded.v26
-rw-r--r--src/SpecificGen/GF2519_32Reflective.v14
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Common.v75
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonBinOp.v8
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOp.v8
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v8
-rw-r--r--src/SpecificGen/GF25519_32Bounded.v26
-rw-r--r--src/SpecificGen/GF25519_32Reflective.v14
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Common.v75
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonBinOp.v8
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOp.v8
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v8
-rw-r--r--src/SpecificGen/GF25519_64Bounded.v40
-rw-r--r--src/SpecificGen/GF25519_64BoundedCommon.v166
-rw-r--r--src/SpecificGen/GF25519_64Reflective.v36
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Common.v77
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonBinOp.v8
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOp.v8
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v8
-rw-r--r--src/SpecificGen/GF41417_32Bounded.v26
-rw-r--r--src/SpecificGen/GF41417_32Reflective.v14
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Common.v75
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonBinOp.v8
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOp.v8
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v8
-rw-r--r--src/SpecificGen/GF5211_32Bounded.v26
-rw-r--r--src/SpecificGen/GF5211_32Reflective.v14
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Common.v75
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonBinOp.v8
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOp.v8
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v8
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v8
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v8
-rwxr-xr-xsrc/SpecificGen/copy_bounds.sh6
67 files changed, 1704 insertions, 1587 deletions
diff --git a/src/Experiments/Ed25519Extraction.v b/src/Experiments/Ed25519Extraction.v
index b498a8c4e..8ec018c5f 100644
--- a/src/Experiments/Ed25519Extraction.v
+++ b/src/Experiments/Ed25519Extraction.v
@@ -242,25 +242,25 @@ Extract Inductive Word.word => "[Prelude.Bool]" [ "[]" "(:)" ]
Extraction Inline LetIn.Let_In.
(* Word64 *)
-Import Crypto.Reflection.Z.Interpretations.
-Extract Inlined Constant Word64.word64 => "Data.Word.Word64".
+Import Crypto.Reflection.Z.Interpretations64.
+Extract Inlined Constant WordW.wordW => "Data.Word.Word64".
Extract Inlined Constant GF25519BoundedCommon.word64 => "Data.Word.Word64".
Extract Inlined Constant GF25519BoundedCommon.w64eqb => "(Prelude.==)".
-Extract Inlined Constant Word64.word64ToZ => "Prelude.fromIntegral".
-Extract Inlined Constant Word64.ZToWord64 => "Prelude.fromIntegral".
+Extract Inlined Constant WordW.wordWToZ => "Prelude.fromIntegral".
+Extract Inlined Constant WordW.ZToWordW => "Prelude.fromIntegral".
Extract Inlined Constant GF25519BoundedCommon.word64ToZ => "Prelude.fromIntegral".
Extract Inlined Constant GF25519BoundedCommon.NToWord64 => "Prelude.fromIntegral".
Extract Inlined Constant GF25519BoundedCommon.ZToWord64 => "Prelude.fromIntegral".
-Extract Inlined Constant Word64.add => "(Prelude.+)".
-Extract Inlined Constant Word64.mul => "(Prelude.*)".
-Extract Inlined Constant Word64.sub => "(Prelude.-)".
-Extract Inlined Constant Word64.land => "(Data.Bits..&.)".
-Extract Inlined Constant Word64.lor => "(Data.Bits..|.)".
-Extract Inlined Constant Word64.neg => "(\_ w -> Prelude.negate w)". (* FIXME: reification: drop arg1 *)
-Extract Inlined Constant Word64.shr => "(\w n -> Data.Bits.shiftR w (Prelude.fromIntegral n))".
-Extract Inlined Constant Word64.shl => "(\w n -> Data.Bits.shiftL w (Prelude.fromIntegral n))".
-Extract Inlined Constant Word64.cmovle => "(\x y r1 r2 -> if x Prelude.<= y then r1 else r2)".
-Extract Inlined Constant Word64.cmovne => "(\x y r1 r2 -> if x Prelude.== y then r1 else r2)".
+Extract Inlined Constant WordW.add => "(Prelude.+)".
+Extract Inlined Constant WordW.mul => "(Prelude.*)".
+Extract Inlined Constant WordW.sub => "(Prelude.-)".
+Extract Inlined Constant WordW.land => "(Data.Bits..&.)".
+Extract Inlined Constant WordW.lor => "(Data.Bits..|.)".
+Extract Inlined Constant WordW.neg => "(\_ w -> Prelude.negate w)". (* FIXME: reification: drop arg1 *)
+Extract Inlined Constant WordW.shr => "(\w n -> Data.Bits.shiftR w (Prelude.fromIntegral n))".
+Extract Inlined Constant WordW.shl => "(\w n -> Data.Bits.shiftL w (Prelude.fromIntegral n))".
+Extract Inlined Constant WordW.cmovle => "(\x y r1 r2 -> if x Prelude.<= y then r1 else r2)".
+Extract Inlined Constant WordW.cmovne => "(\x y r1 r2 -> if x Prelude.== y then r1 else r2)".
(* inlining, primarily to reduce polymorphism *)
Extraction Inline dec_eq_Z dec_eq_N dec_eq_sig_hprop.
@@ -277,7 +277,7 @@ Extraction Inline Crypto.Util.HList.mapt' Crypto.Util.HList.mapt Crypto.Util.Tup
Extraction Inline GF25519BoundedCommon.exist_fe25519 GF25519BoundedCommon.exist_fe25519W GF25519BoundedCommon.proj1_fe25519W.
Extraction Inline Crypto.Specific.GF25519Bounded.mulW Crypto.Specific.GF25519Bounded.addW Crypto.Specific.GF25519Bounded.subW.
-(*done-at-haskell-level: Extraction Inline Crypto.Specific.GF25519Bounded.mul Crypto.Specific.GF25519Bounded.add Crypto.Specific.GF25519Bounded.sub Crypto.Specific.GF25519Bounded.inv Crypto.Specific.GF25519Bounded.sqrt. *)
+(*done-at-haskell-level: Extraction Inline Crypto.Specific.GF25519Bounded.mul Crypto.Specific.GF25519Bounded.add Crypto.Specific.GF25519Bounded.sub Crypto.Specific.GF25519Bounded.inv Crypto.Specific.GF25519Bounded.sqrt. *)
Extraction Implicit Ed25519.SHA512 [ 1 ].
Extract Constant Ed25519.SHA512 =>
diff --git a/src/Reflection/Z/Interpretations.v b/src/Reflection/Z/Interpretations.v
deleted file mode 100644
index 5a015651a..000000000
--- a/src/Reflection/Z/Interpretations.v
+++ /dev/null
@@ -1,812 +0,0 @@
-(** * Interpretation of PHOAS syntax for expression trees on ℤ *)
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Application.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemListZOperations.
-Require Import Crypto.Util.Equality.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Option.
-Require Crypto.Util.Tuple.
-Require Crypto.Util.HList.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.
-Require Import Bedrock.Word.
-Require Import Crypto.Util.WordUtil.
-Export Reflection.Syntax.Notations.
-
-Local Notation eta x := (fst x, snd x).
-Local Notation eta3 x := (eta (fst x), snd x).
-Local Notation eta4 x := (eta3 (fst x), snd x).
-
-Module Z.
- Definition interp_base_type (t : base_type) : Type := interp_base_type t.
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := interp_op src dst f.
-End Z.
-
-Module LiftOption.
- Section lift_option.
- Context (T : Type).
-
- Definition interp_flat_type (t : flat_type base_type)
- := option (interp_flat_type (fun _ => T) t).
-
- Definition interp_base_type' (t : base_type)
- := match t with
- | TZ => option T
- end.
-
- Definition of' {t} : Syntax.interp_flat_type interp_base_type' t -> interp_flat_type t
- := @smart_interp_flat_map
- base_type
- interp_base_type' interp_flat_type
- (fun t => match t with TZ => fun x => x end)
- (fun _ _ x y => match x, y with
- | Some x', Some y' => Some (x', y')
- | _, _ => None
- end)
- t.
-
- Fixpoint to' {t} : interp_flat_type t -> Syntax.interp_flat_type interp_base_type' t
- := match t return interp_flat_type t -> Syntax.interp_flat_type interp_base_type' t with
- | Tbase TZ => fun x => x
- | Prod A B => fun x => (@to' A (option_map (@fst _ _) x),
- @to' B (option_map (@snd _ _) x))
- end.
-
- Definition lift_relation {interp_base_type2}
- (R : forall t, T -> interp_base_type2 t -> Prop)
- : forall t, interp_base_type' t -> interp_base_type2 t -> Prop
- := fun t x y => match of' (t:=Tbase t) x with
- | Some x' => R t x' y
- | None => True
- end.
-
- Definition Some {t} (x : T) : interp_base_type' t
- := match t with
- | TZ => Some x
- end.
- End lift_option.
- Global Arguments of' {T t} _.
- Global Arguments to' {T t} _.
- Global Arguments Some {T t} _.
- Global Arguments lift_relation {T _} R _ _ _.
-
- Section lift_option2.
- Context (T U : Type).
- Definition lift_relation2 (R : T -> U -> Prop)
- : forall t, interp_base_type' T t -> interp_base_type' U t -> Prop
- := fun t x y => match of' (t:=Tbase t) x, of' (t:=Tbase t) y with
- | Datatypes.Some x', Datatypes.Some y' => R x' y'
- | None, None => True
- | _, _ => False
- end.
- End lift_option2.
- Global Arguments lift_relation2 {T U} R _ _ _.
-End LiftOption.
-
-Module Word64.
- Definition bit_width : nat := 64.
- Definition word64 := word bit_width.
- Delimit Scope word64_scope with word64.
- Bind Scope word64_scope with word64.
-
- Definition word64ToZ (x : word64) : Z
- := Z.of_N (wordToN x).
- Definition ZToWord64 (x : Z) : word64
- := NToWord _ (Z.to_N x).
-
- Ltac fold_Word64_Z :=
- repeat match goal with
- | [ |- context G[NToWord bit_width (Z.to_N ?x)] ]
- => let G' := context G [ZToWord64 x] in change G'
- | [ |- context G[Z.of_N (wordToN ?x)] ]
- => let G' := context G [word64ToZ x] in change G'
- | [ H : context G[NToWord bit_width (Z.to_N ?x)] |- _ ]
- => let G' := context G [ZToWord64 x] in change G' in H
- | [ H : context G[Z.of_N (wordToN ?x)] |- _ ]
- => let G' := context G [word64ToZ x] in change G' in H
- end.
-
- Create HintDb push_word64ToZ discriminated.
- Hint Extern 1 => progress autorewrite with push_word64ToZ in * : push_word64ToZ.
-
- Lemma bit_width_pos : (0 < Z.of_nat bit_width)%Z.
- Proof. simpl; omega. Qed.
- Local Hint Resolve bit_width_pos : zarith.
-
- Ltac arith := solve [ omega | auto using bit_width_pos with zarith ].
-
- Lemma word64ToZ_bound w : (0 <= word64ToZ w < 2^Z.of_nat bit_width)%Z.
- Proof.
- pose proof (wordToNat_bound w) as H.
- apply Nat2Z.inj_lt in H.
- rewrite Zpow_pow2, Z2Nat.id in H by (apply Z.pow_nonneg; omega).
- unfold word64ToZ.
- rewrite wordToN_nat, nat_N_Z; omega.
- Qed.
-
- Lemma word64ToZ_log_bound w : (0 <= word64ToZ w /\ Z.log2 (word64ToZ w) < Z.of_nat bit_width)%Z.
- Proof.
- pose proof (word64ToZ_bound w) as H.
- destruct (Z_zerop (word64ToZ w)) as [H'|H'].
- { rewrite H'; simpl; omega. }
- { split; [ | apply Z.log2_lt_pow2 ]; try omega. }
- Qed.
-
- Lemma ZToWord64_word64ToZ (x : word64) : ZToWord64 (word64ToZ x) = x.
- Proof.
- unfold ZToWord64, word64ToZ.
- rewrite N2Z.id, NToWord_wordToN.
- reflexivity.
- Qed.
- Hint Rewrite ZToWord64_word64ToZ : push_word64ToZ.
-
- Lemma word64ToZ_ZToWord64 (x : Z) : (0 <= x < 2^Z.of_nat bit_width)%Z -> word64ToZ (ZToWord64 x) = x.
- Proof.
- unfold ZToWord64, word64ToZ; intros [H0 H1].
- pose proof H1 as H1'; apply Z2Nat.inj_lt in H1'; [ | omega.. ].
- rewrite <- Z.pow_Z2N_Zpow in H1' by omega.
- replace (Z.to_nat 2) with 2%nat in H1' by reflexivity.
- rewrite wordToN_NToWord_idempotent, Z2N.id by (omega || auto using bound_check_nat_N).
- reflexivity.
- Qed.
- Hint Rewrite word64ToZ_ZToWord64 using arith : push_word64ToZ.
-
- Definition add : word64 -> word64 -> word64 := @wplus _.
- Definition sub : word64 -> word64 -> word64 := @wminus _.
- Definition mul : word64 -> word64 -> word64 := @wmult _.
- Definition shl : word64 -> word64 -> word64 := @wordBin N.shiftl _.
- Definition shr : word64 -> word64 -> word64 := @wordBin N.shiftr _.
- Definition land : word64 -> word64 -> word64 := @wand _.
- Definition lor : word64 -> word64 -> word64 := @wor _.
- Definition neg (int_width : Z) : word64 -> word64 (* TODO: Is this right? *)
- := fun x => ZToWord64 (ModularBaseSystemListZOperations.neg int_width (word64ToZ x)).
- Definition cmovne : word64 -> word64 -> word64 -> word64 -> word64 (* TODO: Is this right? *)
- := fun x y z w => ZToWord64 (ModularBaseSystemListZOperations.cmovne (word64ToZ x) (word64ToZ y) (word64ToZ z) (word64ToZ w)).
- Definition cmovle : word64 -> word64 -> word64 -> word64 -> word64 (* TODO: Is this right? *)
- := fun x y z w => ZToWord64 (ModularBaseSystemListZOperations.cmovl (word64ToZ x) (word64ToZ y) (word64ToZ z) (word64ToZ w)).
-
- Infix "+" := add : word64_scope.
- Infix "-" := sub : word64_scope.
- Infix "*" := mul : word64_scope.
- Infix "<<" := shl : word64_scope.
- Infix ">>" := shr : word64_scope.
- Infix "&'" := land : word64_scope.
-
- (*Local*) Hint Resolve <- Z.log2_lt_pow2_alt : zarith.
- Local Hint Resolve eq_refl : zarith.
- Local Ltac w64ToZ_t :=
- intros;
- try match goal with
- | [ |- ?wordToZ ?op = _ ]
- => let op' := head op in
- cbv [wordToZ op'] in *
- end;
- autorewrite with push_Zto_N push_Zof_N push_wordToN; try reflexivity.
- Local Ltac w64ToZ_extra_t :=
- repeat first [ reflexivity
- | progress cbv [ModularBaseSystemListZOperations.neg ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.cmovl] in *
- | progress break_match
- | progress fold_Word64_Z
- | progress intros
- | progress autorewrite with push_Zto_N push_Zof_N push_wordToN push_word64ToZ ].
-
- Local Notation bounds_statement wop Zop
- := ((0 <= Zop -> Z.log2 Zop < Z.of_nat bit_width -> word64ToZ wop = Zop)%Z).
- Local Notation bounds_statement_tuple wop Zop
- := ((HList.hlist (fun v => 0 <= v /\ Z.log2 v < Z.of_nat bit_width) Zop -> Tuple.map word64ToZ wop = Zop)%Z).
- Local Notation bounds_1statement wop Zop
- := (forall x,
- bounds_statement (wop x) (Zop (word64ToZ x))).
- Local Notation bounds_2statement wop Zop
- := (forall x y,
- bounds_statement (wop x y) (Zop (word64ToZ x) (word64ToZ y))).
- Local Notation bounds_4statement wop Zop
- := (forall x y z w,
- bounds_statement (wop x y z w) (Zop (word64ToZ x) (word64ToZ y) (word64ToZ z) (word64ToZ w))).
-
- Lemma word64ToZ_add : bounds_2statement add Z.add. Proof. w64ToZ_t. Qed.
- Lemma word64ToZ_sub : bounds_2statement sub Z.sub. Proof. w64ToZ_t. Qed.
- Lemma word64ToZ_mul : bounds_2statement mul Z.mul. Proof. w64ToZ_t. Qed.
- Lemma word64ToZ_shl : bounds_2statement shl Z.shiftl.
- Proof.
- w64ToZ_t; w64ToZ_extra_t; unfold word64ToZ, wordBin.
- rewrite wordToN_NToWord_idempotent; [rewrite <- Z_inj_shiftl; reflexivity|].
- apply N2Z.inj_lt.
- rewrite Z_inj_shiftl.
- destruct (Z.lt_ge_cases 0 ((word64ToZ x) << (word64ToZ y)))%Z;
- [|eapply Z.le_lt_trans; [|apply N2Z.inj_lt, Npow2_gt0]; assumption].
- rewrite Npow2_N, N2Z.inj_pow.
- apply Z.log2_lt_pow2; assumption.
- Qed.
-
- Lemma word64ToZ_shr : bounds_2statement shr Z.shiftr.
- Proof.
- w64ToZ_t; w64ToZ_extra_t; unfold word64ToZ, wordBin.
- rewrite wordToN_NToWord_idempotent; [rewrite <- Z_inj_shiftr; reflexivity|].
- apply N2Z.inj_lt.
- rewrite Z_inj_shiftr.
- destruct (Z.lt_ge_cases 0 ((word64ToZ x) >> (word64ToZ y)))%Z;
- [|eapply Z.le_lt_trans; [|apply N2Z.inj_lt, Npow2_gt0]; assumption].
- rewrite Npow2_N, N2Z.inj_pow.
- apply Z.log2_lt_pow2; assumption.
- Qed.
-
- Lemma word64ToZ_land : bounds_2statement land Z.land.
- Proof. w64ToZ_t. Qed.
- Lemma word64ToZ_lor : bounds_2statement lor Z.lor.
- Proof. w64ToZ_t. Qed.
- Lemma word64ToZ_neg int_width : bounds_1statement (neg int_width) (ModularBaseSystemListZOperations.neg int_width).
- Proof. w64ToZ_t; w64ToZ_extra_t. Qed.
- Lemma word64ToZ_cmovne : bounds_4statement cmovne ModularBaseSystemListZOperations.cmovne.
- Proof. w64ToZ_t; w64ToZ_extra_t. Qed.
- Lemma word64ToZ_cmovle : bounds_4statement cmovle ModularBaseSystemListZOperations.cmovl.
- Proof. w64ToZ_t; w64ToZ_extra_t. Qed.
-
- Definition interp_base_type (t : base_type) : Type
- := match t with
- | TZ => word64
- end.
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
- | Add => fun xy => fst xy + snd xy
- | Sub => fun xy => fst xy - snd xy
- | Mul => fun xy => fst xy * snd xy
- | Shl => fun xy => fst xy << snd xy
- | Shr => fun xy => fst xy >> snd xy
- | Land => fun xy => land (fst xy) (snd xy)
- | Lor => fun xy => lor (fst xy) (snd xy)
- | Neg int_width => fun x => neg int_width x
- | Cmovne => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovne x y z w
- | Cmovle => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovle x y z w
- end%word64.
-
- Definition of_Z ty : Z.interp_base_type ty -> interp_base_type ty
- := match ty return Z.interp_base_type ty -> interp_base_type ty with
- | TZ => ZToWord64
- end.
- Definition to_Z ty : interp_base_type ty -> Z.interp_base_type ty
- := match ty return interp_base_type ty -> Z.interp_base_type ty with
- | TZ => word64ToZ
- end.
-
- Module Export Rewrites.
- Ltac word64_util_arith := omega.
-
- Hint Rewrite word64ToZ_add using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_add using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_sub using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_sub using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_mul using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_mul using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_shl using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_shl using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_shr using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_shr using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_land using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_land using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_lor using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_lor using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_neg using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_neg using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_cmovne using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_cmovne using word64_util_arith : pull_word64ToZ.
- Hint Rewrite word64ToZ_cmovle using word64_util_arith : push_word64ToZ.
- Hint Rewrite <- word64ToZ_cmovle using word64_util_arith : pull_word64ToZ.
- End Rewrites.
-End Word64.
-
-Module ZBounds.
- Record bounds := { lower : Z ; upper : Z }.
- Bind Scope bounds_scope with bounds.
- Definition t := option bounds. (* TODO?: Separate out the bounds computation from the overflow computation? e.g., have [safety := in_bounds | overflow] and [t := bounds * safety]? *)
- Bind Scope bounds_scope with t.
- Local Coercion Z.of_nat : nat >-> Z.
- Definition word64ToBounds (x : Word64.word64) : t
- := let v := Word64.word64ToZ x in Some {| lower := v ; upper := v |}.
- Definition SmartBuildBounds (l u : Z)
- := if ((0 <=? l) && (Z.log2 u <? Word64.bit_width))%Z%bool
- then Some {| lower := l ; upper := u |}
- else None.
- Definition t_map1 (f : bounds -> bounds) (x : t)
- := match x with
- | Some x
- => match f x with
- | Build_bounds l u
- => SmartBuildBounds l u
- end
- | _ => None
- end%Z.
- Definition t_map2 (f : bounds -> bounds -> bounds) (x y : t)
- := match x, y with
- | Some x, Some y
- => match f x y with
- | Build_bounds l u
- => SmartBuildBounds l u
- end
- | _, _ => None
- end%Z.
- Definition t_map4 (f : bounds -> bounds -> bounds -> bounds -> bounds) (x y z w : t)
- := match x, y, z, w with
- | Some x, Some y, Some z, Some w
- => match f x y z w with
- | Build_bounds l u
- => SmartBuildBounds l u
- end
- | _, _, _, _ => None
- end%Z.
- Definition add' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx + ly ; upper := ux + uy |}.
- Definition add : t -> t -> t := t_map2 add'.
- Definition sub' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx - uy ; upper := ux - ly |}.
- Definition sub : t -> t -> t := t_map2 sub'.
- Definition mul' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx * ly ; upper := ux * uy |}.
- Definition mul : t -> t -> t := t_map2 mul'.
- Definition shl' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx << ly ; upper := ux << uy |}.
- Definition shl : t -> t -> t := t_map2 shl'.
- Definition shr' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx >> uy ; upper := ux >> ly |}.
- Definition shr : t -> t -> t := t_map2 shr'.
- Definition land' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := 0 ; upper := Z.min ux uy |}.
- Definition land : t -> t -> t := t_map2 land'.
- Definition lor' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in
- {| lower := Z.max lx ly;
- upper := 2^(Z.max (Z.log2_up (ux+1)) (Z.log2_up (uy+1))) - 1 |}.
- Definition lor : t -> t -> t := t_map2 lor'.
- Definition neg' (int_width : Z) : bounds -> bounds
- := fun v
- => let (lb, ub) := v in
- let might_be_one := ((lb <=? 1) && (1 <=? ub))%Z%bool in
- let must_be_one := ((lb =? 1) && (ub =? 1))%Z%bool in
- if must_be_one
- then {| lower := Z.ones int_width ; upper := Z.ones int_width |}
- else if might_be_one
- then {| lower := 0 ; upper := Z.ones int_width |}
- else {| lower := 0 ; upper := 0 |}.
- Definition neg (int_width : Z) : t -> t
- := fun v
- => if ((0 <=? int_width) && (int_width <=? Word64.bit_width))%Z%bool
- then t_map1 (neg' int_width) v
- else None.
- Definition cmovne' (r1 r2 : bounds) : bounds
- := let (lr1, ur1) := r1 in let (lr2, ur2) := r2 in {| lower := Z.min lr1 lr2 ; upper := Z.max ur1 ur2 |}.
- Definition cmovne (x y r1 r2 : t) : t := t_map4 (fun _ _ => cmovne') x y r1 r2.
- Definition cmovle' (r1 r2 : bounds) : bounds
- := let (lr1, ur1) := r1 in let (lr2, ur2) := r2 in {| lower := Z.min lr1 lr2 ; upper := Z.max ur1 ur2 |}.
- Definition cmovle (x y r1 r2 : t) : t := t_map4 (fun _ _ => cmovle') x y r1 r2.
-
- Module Export Notations.
- Delimit Scope bounds_scope with bounds.
- Notation "b[ l ~> u ]" := {| lower := l ; upper := u |} : bounds_scope.
- Infix "+" := add : bounds_scope.
- Infix "-" := sub : bounds_scope.
- Infix "*" := mul : bounds_scope.
- Infix "<<" := shl : bounds_scope.
- Infix ">>" := shr : bounds_scope.
- Infix "&'" := land : bounds_scope.
- End Notations.
-
- Definition interp_base_type (ty : base_type) : Type
- := LiftOption.interp_base_type' bounds ty.
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
- | Add => fun xy => fst xy + snd xy
- | Sub => fun xy => fst xy - snd xy
- | Mul => fun xy => fst xy * snd xy
- | Shl => fun xy => fst xy << snd xy
- | Shr => fun xy => fst xy >> snd xy
- | Land => fun xy => land (fst xy) (snd xy)
- | Lor => fun xy => lor (fst xy) (snd xy)
- | Neg int_width => fun x => neg int_width x
- | Cmovne => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovne x y z w
- | Cmovle => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovle x y z w
- end%bounds.
-
- Definition of_word64 ty : Word64.interp_base_type ty -> interp_base_type ty
- := match ty return Word64.interp_base_type ty -> interp_base_type ty with
- | TZ => word64ToBounds
- end.
-
- Ltac inversion_bounds :=
- let lower := (eval cbv [lower] in (fun x => lower x)) in
- let upper := (eval cbv [upper] in (fun y => upper y)) in
- repeat match goal with
- | [ H : _ = _ :> bounds |- _ ]
- => pose proof (f_equal lower H); pose proof (f_equal upper H); clear H;
- cbv beta iota in *
- | [ H : _ = _ :> t |- _ ]
- => unfold t in H; inversion_option
- end.
-End ZBounds.
-
-Module BoundedWord64.
- Local Notation is_bounded_by value lower upper
- := ((0 <= lower /\ lower <= Word64.word64ToZ value <= upper /\ Z.log2 upper < Z.of_nat Word64.bit_width)%Z)
- (only parsing).
- Record BoundedWord :=
- { lower : Z ; value : Word64.word64 ; upper : Z ;
- in_bounds : is_bounded_by value lower upper }.
- Bind Scope bounded_word_scope with BoundedWord.
- Definition t := option BoundedWord.
- Bind Scope bounded_word_scope with t.
- Local Coercion Z.of_nat : nat >-> Z.
-
- Ltac inversion_BoundedWord :=
- repeat match goal with
- | _ => progress subst
- | [ H : _ = _ :> BoundedWord |- _ ]
- => pose proof (f_equal lower H);
- pose proof (f_equal upper H);
- pose proof (f_equal value H);
- clear H
- end.
-
- Definition interp_base_type (ty : base_type)
- := LiftOption.interp_base_type' BoundedWord ty.
-
- Definition word64ToBoundedWord (x : Word64.word64) : t.
- Proof.
- refine (let v := Word64.word64ToZ x in
- match Sumbool.sumbool_of_bool (0 <=? v)%Z, Sumbool.sumbool_of_bool (Z.log2 v <? Z.of_nat Word64.bit_width)%Z with
- | left Hl, left Hu
- => Some {| lower := Word64.word64ToZ x ; value := x ; upper := Word64.word64ToZ x |}
- | _, _ => None
- end).
- subst v.
- abstract (Z.ltb_to_lt; repeat split; (assumption || reflexivity)).
- Defined.
-
- Definition boundedWordToWord64 (x : t) : Word64.word64
- := match x with
- | Some x' => value x'
- | None => Word64.ZToWord64 0
- end.
-
- Definition of_word64 ty : Word64.interp_base_type ty -> interp_base_type ty
- := match ty return Word64.interp_base_type ty -> interp_base_type ty with
- | TZ => word64ToBoundedWord
- end.
- Definition to_word64 ty : interp_base_type ty -> Word64.interp_base_type ty
- := match ty return interp_base_type ty -> Word64.interp_base_type ty with
- | TZ => boundedWordToWord64
- end.
-
- (** XXX FIXME(jgross) This is going to break horribly if we need to support any types other than [Z] *)
- Definition to_word64' ty : BoundedWord -> Word64.interp_base_type ty
- := match ty return BoundedWord -> Word64.interp_base_type ty with
- | TZ => fun x => boundedWordToWord64 (Some x)
- end.
-
- Definition to_Z' ty : BoundedWord -> Z.interp_base_type ty
- := fun x => Word64.to_Z _ (to_word64' _ x).
-
- Definition of_Z ty : Z.interp_base_type ty -> interp_base_type ty
- := fun x => of_word64 _ (Word64.of_Z _ x).
- Definition to_Z ty : interp_base_type ty -> Z.interp_base_type ty
- := fun x => Word64.to_Z _ (to_word64 _ x).
-
- Definition BoundedWordToBounds (x : BoundedWord) : ZBounds.bounds
- := {| ZBounds.lower := lower x ; ZBounds.upper := upper x |}.
-
- Definition to_bounds' : t -> ZBounds.t
- := option_map BoundedWordToBounds.
-
- Definition to_bounds ty : interp_base_type ty -> ZBounds.interp_base_type ty
- := match ty return interp_base_type ty -> ZBounds.interp_base_type ty with
- | TZ => to_bounds'
- end.
-
- Definition t_map1
- (opW : Word64.word64 -> Word64.word64)
- (opB : ZBounds.t -> ZBounds.t)
- (pf : forall x l u,
- opB (Some (BoundedWordToBounds x))
- = Some {| ZBounds.lower := l ; ZBounds.upper := u |}
- -> let val := opW (value x) in
- is_bounded_by val l u)
- : t -> t
- := fun x : t
- => match x with
- | Some x
- => match opB (Some (BoundedWordToBounds x))
- as bop return opB (Some (BoundedWordToBounds x)) = bop -> t
- with
- | Some (ZBounds.Build_bounds l u)
- => fun Heq => Some {| lower := l ; value := opW (value x) ; upper := u;
- in_bounds := pf _ _ _ Heq |}
- | None => fun _ => None
- end eq_refl
- | _ => None
- end.
-
- Definition t_map2
- (opW : Word64.word64 -> Word64.word64 -> Word64.word64)
- (opB : ZBounds.t -> ZBounds.t -> ZBounds.t)
- (pf : forall x y l u,
- opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- = Some {| ZBounds.lower := l ; ZBounds.upper := u |}
- -> let val := opW (value x) (value y) in
- is_bounded_by val l u)
- : t -> t -> t
- := fun x y : t
- => match x, y with
- | Some x, Some y
- => match opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- as bop return opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = bop -> t
- with
- | Some (ZBounds.Build_bounds l u)
- => fun Heq => Some {| lower := l ; value := opW (value x) (value y) ; upper := u;
- in_bounds := pf _ _ _ _ Heq |}
- | None => fun _ => None
- end eq_refl
- | _, _ => None
- end.
-
- Definition t_map4
- (opW : Word64.word64 -> Word64.word64 -> Word64.word64 -> Word64.word64 -> Word64.word64)
- (opB : ZBounds.t -> ZBounds.t -> ZBounds.t -> ZBounds.t -> ZBounds.t)
- (pf : forall x y z w l u,
- opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- = Some {| ZBounds.lower := l ; ZBounds.upper := u |}
- -> let val := opW (value x) (value y) (value z) (value w) in
- is_bounded_by val l u)
- : t -> t -> t -> t -> t
- := fun x y z w : t
- => match x, y, z, w with
- | Some x, Some y, Some z, Some w
- => match opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- as bop return opB _ _ _ _ = bop -> t
- with
- | Some (ZBounds.Build_bounds l u)
- => fun Heq => Some {| lower := l ; value := opW (value x) (value y) (value z) (value w) ; upper := u;
- in_bounds := pf _ _ _ _ _ _ Heq |}
- | None => fun _ => None
- end eq_refl
- | _, _, _, _ => None
- end.
-
- Local Opaque Word64.bit_width.
- Hint Resolve Z.ones_nonneg : zarith.
- Local Ltac t_prestart :=
- repeat first [ match goal with
- | [ |- forall x l u, ?opB (Some (BoundedWordToBounds x)) = Some _ -> let val := ?opW (value x) in _ ]
- => let opB' := head opB in let opW' := head opW in progress (try unfold opB'; try unfold opW')
- | [ |- forall x y l u, ?opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = Some _ -> let val := ?opW (value x) (value y) in _ ]
- => progress (try unfold opB; try unfold opW)
- | [ |- forall x y z w l u, ?opB _ _ _ _ = Some _ -> let val := ?opW (value x) (value y) (value z) (value w) in _ ]
- => progress (try unfold opB; try unfold opW)
- | [ |- appcontext[ZBounds.t_map1 ?op] ] => let op' := head op in unfold op'
- | [ |- appcontext[ZBounds.t_map2 ?op] ] => let op' := head op in unfold op'
- | [ |- appcontext[?op (ZBounds.Build_bounds _ _)] ] => let op' := head op in unfold op'
- | [ |- appcontext[?op (ZBounds.Build_bounds _ _) (ZBounds.Build_bounds _ _)] ] => unfold op
- end
- | progress cbv [BoundedWordToBounds ZBounds.SmartBuildBounds cmovne cmovl ModularBaseSystemListZOperations.neg] in *
- | progress break_match
- | progress break_match_hyps
- | progress intros
- | progress subst
- | progress ZBounds.inversion_bounds
- | progress inversion_option
- | progress Word64.fold_Word64_Z
- | progress autorewrite with bool_congr_setoid in *
- | progress destruct_head' and
- | progress Z.ltb_to_lt
- | assumption
- | progress destruct_head' BoundedWord; simpl in *
- | progress autorewrite with push_word64ToZ
- | progress repeat apply conj
- | solve [ Word64.arith ]
- | progress destruct_head' or ].
- Local Ltac t_start :=
- repeat first [ progress t_prestart
- | match goal with
- | [ |- appcontext[Z.min ?x ?y] ]
- => apply (Z.min_case_strong x y)
- | [ |- appcontext[Z.max ?x ?y] ]
- => apply (Z.max_case_strong x y)
- end ].
-
- Local Hint Resolve Word64.bit_width_pos : zarith.
- Local Hint Extern 1 (Z.log2 _ < _)%Z => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono; eassumption | eassumption ] : zarith.
- (* Local *) Hint Resolve <- Z.log2_lt_pow2_alt : zarith.
-
-
- Definition add : t -> t -> t.
- Proof.
- refine (t_map2 Word64.add ZBounds.add _);
- abstract (t_start; eapply add_valid_update; eauto).
- Defined.
-
- Definition sub : t -> t -> t.
- Proof.
- refine (t_map2 Word64.sub ZBounds.sub _);
- abstract (t_start; eapply sub_valid_update; eauto).
- Defined.
-
- Definition mul : t -> t -> t.
- Proof.
- refine (t_map2 Word64.mul ZBounds.mul _);
- abstract (t_start; eapply mul_valid_update; eauto).
- Defined.
-
- Definition land : t -> t -> t.
- Proof.
- refine (t_map2 Word64.land ZBounds.land _);
- abstract (t_prestart; eapply land_valid_update; eauto).
- Defined.
-
- Definition lor : t -> t -> t.
- Proof.
- refine (t_map2 Word64.lor ZBounds.lor _);
- abstract (t_prestart; eapply lor_valid_update; eauto).
- Defined.
-
- Definition shl : t -> t -> t.
- Proof.
- refine (t_map2 Word64.shl ZBounds.shl _);
- abstract (t_start; eapply shl_valid_update; eauto).
- Defined.
-
- Definition shr : t -> t -> t.
- Proof.
- refine (t_map2 Word64.shr ZBounds.shr _);
- abstract (t_start; eapply shr_valid_update; eauto).
- Defined.
-
- Definition neg (int_width : Z) : t -> t.
- Proof. refine (t_map1 (Word64.neg int_width) (ZBounds.neg int_width) _); abstract t_start. Defined.
-
- Definition cmovne : t -> t -> t -> t -> t.
- Proof. refine (t_map4 Word64.cmovne ZBounds.cmovne _); abstract t_start. Defined.
-
- Definition cmovle : t -> t -> t -> t -> t.
- Proof. refine (t_map4 Word64.cmovle ZBounds.cmovle _); abstract t_start. Defined.
-
- Local Notation unop_correct op opW opB :=
- (forall x v, op (Some x) = Some v
- -> value v = opW (value x)
- /\ Some (BoundedWordToBounds v) = opB (Some (BoundedWordToBounds x)))
- (only parsing).
-
- Local Notation binop_correct op opW opB :=
- (forall x y v, op (Some x) (Some y) = Some v
- -> value v = opW (value x) (value y)
- /\ Some (BoundedWordToBounds v) = opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)))
- (only parsing).
-
- Local Notation op4_correct op opW opB :=
- (forall x y z w v, op (Some x) (Some y) (Some z) (Some w) = Some v
- -> value v = opW (value x) (value y) (value z) (value w)
- /\ Some (BoundedWordToBounds v) = opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w)))
- (only parsing).
-
- Lemma t_map1_correct opW opB pf
- : unop_correct (t_map1 opW opB pf) opW opB.
- Proof.
- intros ?? H.
- unfold t_map1 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map2_correct opW opB pf
- : binop_correct (t_map2 opW opB pf) opW opB.
- Proof.
- intros ??? H.
- unfold t_map2 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map4_correct opW opB pf
- : op4_correct (t_map4 opW opB pf) opW opB.
- Proof.
- intros ????? H.
- unfold t_map4 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Local Notation unop_correct_None op opB :=
- (forall x, op (Some x) = None -> opB (Some (BoundedWordToBounds x)) = None)
- (only parsing).
-
- Local Notation binop_correct_None op opB :=
- (forall x y, op (Some x) (Some y) = None -> opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = None)
- (only parsing).
-
- Local Notation op4_correct_None op opB :=
- (forall x y z w, op (Some x) (Some y) (Some z) (Some w) = None
- -> opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- = None)
- (only parsing).
-
- Lemma t_map1_correct_None opW opB pf
- : unop_correct_None (t_map1 opW opB pf) opB.
- Proof.
- intros ? H.
- unfold t_map1 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map2_correct_None opW opB pf
- : binop_correct_None (t_map2 opW opB pf) opB.
- Proof.
- intros ?? H.
- unfold t_map2 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map4_correct_None opW opB pf
- : op4_correct_None (t_map4 opW opB pf) opB.
- Proof.
- intros ???? H.
- unfold t_map4 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Module Export Notations.
- Delimit Scope bounded_word_scope with bounded_word.
- Notation "b[ l ~> u ]" := {| lower := l ; upper := u |} : bounded_word_scope.
- Infix "+" := add : bounded_word_scope.
- Infix "-" := sub : bounded_word_scope.
- Infix "*" := mul : bounded_word_scope.
- Infix "<<" := shl : bounded_word_scope.
- Infix ">>" := shr : bounded_word_scope.
- Infix "&'" := land : bounded_word_scope.
- End Notations.
-
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
- | Add => fun xy => fst xy + snd xy
- | Sub => fun xy => fst xy - snd xy
- | Mul => fun xy => fst xy * snd xy
- | Shl => fun xy => fst xy << snd xy
- | Shr => fun xy => fst xy >> snd xy
- | Land => fun xy => land (fst xy) (snd xy)
- | Lor => fun xy => lor (fst xy) (snd xy)
- | Neg int_width => fun x => neg int_width x
- | Cmovne => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovne x y z w
- | Cmovle => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovle x y z w
- end%bounded_word.
-End BoundedWord64.
-
-Module ZBoundsTuple.
- Definition interp_flat_type (t : flat_type base_type)
- := LiftOption.interp_flat_type ZBounds.bounds t.
-
- Definition of_ZBounds {ty} : Syntax.interp_flat_type ZBounds.interp_base_type ty -> interp_flat_type ty
- := @LiftOption.of' ZBounds.bounds ty.
- Definition to_ZBounds {ty} : interp_flat_type ty -> Syntax.interp_flat_type ZBounds.interp_base_type ty
- := @LiftOption.to' ZBounds.bounds ty.
-End ZBoundsTuple.
-
-Module BoundedWord64Tuple.
- Definition interp_flat_type (t : flat_type base_type)
- := LiftOption.interp_flat_type BoundedWord64.BoundedWord t.
-
- Definition of_BoundedWord64 {ty} : Syntax.interp_flat_type BoundedWord64.interp_base_type ty -> interp_flat_type ty
- := @LiftOption.of' BoundedWord64.BoundedWord ty.
- Definition to_BoundedWord64 {ty} : interp_flat_type ty -> Syntax.interp_flat_type BoundedWord64.interp_base_type ty
- := @LiftOption.to' BoundedWord64.BoundedWord ty.
-End BoundedWord64Tuple.
diff --git a/src/Reflection/Z/Interpretations128.v b/src/Reflection/Z/Interpretations128.v
new file mode 100644
index 000000000..fad3168ef
--- /dev/null
+++ b/src/Reflection/Z/Interpretations128.v
@@ -0,0 +1,13 @@
+(** * Interpretation of PHOAS syntax for expression trees on ℤ *)
+Require Import Coq.ZArith.ZArith.
+Require Export Crypto.Reflection.Z.InterpretationsGen.
+Export Reflection.Syntax.Notations.
+
+Module BitSize128 <: BitSize.
+ Definition bit_width : nat := 128%nat.
+ Lemma bit_width_pos : (0 < Z.of_nat bit_width)%Z.
+ Proof. simpl; omega. Qed.
+End BitSize128.
+
+Module Interpretations128 := InterpretationsGen BitSize128.
+Include Interpretations128.
diff --git a/src/Reflection/Z/Interpretations/Relations.v b/src/Reflection/Z/Interpretations128/Relations.v
index 281f4d7a7..5ce0df08e 100644
--- a/src/Reflection/Z/Interpretations/Relations.v
+++ b/src/Reflection/Z/Interpretations128/Relations.v
@@ -3,7 +3,8 @@ Require Import Coq.micromega.Psatz.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.InterpretationsGen.
+Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.ModularArithmetic.ModularBaseSystemListZOperationsProofs.
Require Import Crypto.Util.Option.
Require Import Crypto.Util.Bool.
@@ -13,34 +14,34 @@ Require Import Crypto.Util.Tactics.
Definition proj_eq_rel {A B} (proj : A -> B) (x : A) (y : B) : Prop
:= proj x = y.
-Definition related'_Z {t} (x : BoundedWord64.BoundedWord) (y : Z.interp_base_type t) : Prop
- := proj_eq_rel (BoundedWord64.to_Z' _) x y.
-Definition related_Z t : BoundedWord64.interp_base_type t -> Z.interp_base_type t -> Prop
+Definition related'_Z {t} (x : BoundedWordW.BoundedWord) (y : Z.interp_base_type t) : Prop
+ := proj_eq_rel (BoundedWordW.to_Z' _) x y.
+Definition related_Z t : BoundedWordW.interp_base_type t -> Z.interp_base_type t -> Prop
:= LiftOption.lift_relation (@related'_Z) t.
-Definition related'_word64 {t} (x : BoundedWord64.BoundedWord) (y : Word64.interp_base_type t) : Prop
- := proj_eq_rel (BoundedWord64.to_word64' _) x y.
-Definition related_word64 t : BoundedWord64.interp_base_type t -> Word64.interp_base_type t -> Prop
- := LiftOption.lift_relation (@related'_word64) t.
-Definition related_bounds t : BoundedWord64.interp_base_type t -> ZBounds.interp_base_type t -> Prop
- := LiftOption.lift_relation2 (proj_eq_rel BoundedWord64.BoundedWordToBounds) t.
-
-Definition related_word64_Z t : Word64.interp_base_type t -> Z.interp_base_type t -> Prop
- := proj_eq_rel (Word64.to_Z _).
-
-Definition related'_word64_bounds : Word64.word64 -> ZBounds.bounds -> Prop
- := fun value b => (0 <= ZBounds.lower b /\ ZBounds.lower b <= Word64.word64ToZ value <= ZBounds.upper b /\ Z.log2 (ZBounds.upper b) < Z.of_nat Word64.bit_width)%Z.
-Definition related_word64_bounds : Word64.word64 -> ZBounds.t -> Prop
+Definition related'_wordW {t} (x : BoundedWordW.BoundedWord) (y : WordW.interp_base_type t) : Prop
+ := proj_eq_rel (BoundedWordW.to_wordW' _) x y.
+Definition related_wordW t : BoundedWordW.interp_base_type t -> WordW.interp_base_type t -> Prop
+ := LiftOption.lift_relation (@related'_wordW) t.
+Definition related_bounds t : BoundedWordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
+ := LiftOption.lift_relation2 (proj_eq_rel BoundedWordW.BoundedWordToBounds) t.
+
+Definition related_wordW_Z t : WordW.interp_base_type t -> Z.interp_base_type t -> Prop
+ := proj_eq_rel (WordW.to_Z _).
+
+Definition related'_wordW_bounds : WordW.wordW -> ZBounds.bounds -> Prop
+ := fun value b => (0 <= Bounds.lower b /\ Bounds.lower b <= WordW.wordWToZ value <= Bounds.upper b /\ Z.log2 (Bounds.upper b) < Z.of_nat WordW.bit_width)%Z.
+Definition related_wordW_bounds : WordW.wordW -> ZBounds.t -> Prop
:= fun value b => match b with
- | Some b => related'_word64_bounds value b
+ | Some b => related'_wordW_bounds value b
| None => True
end.
-Definition related_word64_boundsi (t : base_type) : Word64.interp_base_type t -> ZBounds.interp_base_type t -> Prop
+Definition related_wordW_boundsi (t : base_type) : WordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
:= match t with
- | TZ => related_word64_bounds
+ | TZ => related_wordW_bounds
end.
-Definition related_word64_boundsi' (t : base_type) : ZBounds.bounds -> Word64.interp_base_type t -> Prop
- := match t return ZBounds.bounds -> Word64.interp_base_type t -> Prop with
- | TZ => fun x y => related'_word64_bounds y x
+Definition related_wordW_boundsi' (t : base_type) : ZBounds.bounds -> WordW.interp_base_type t -> Prop
+ := match t return ZBounds.bounds -> WordW.interp_base_type t -> Prop with
+ | TZ => fun x y => related'_wordW_bounds y x
end.
Local Notation related_op R interp_op1 interp_op2
@@ -57,21 +58,21 @@ Local Ltac related_const_t :=
let v := fresh in
let t := fresh in
intros t v; destruct t; intros; simpl in *; hnf; simpl;
- cbv [BoundedWord64.word64ToBoundedWord related'_Z LiftOption.of' related_Z related_word64 related'_word64 proj_eq_rel] in *;
+ cbv [BoundedWordW.wordWToBoundedWord related'_Z LiftOption.of' related_Z related_wordW related'_wordW proj_eq_rel] in *;
break_innermost_match; simpl;
first [ tauto
| Z.ltb_to_lt;
- pose proof (Word64.word64ToZ_log_bound v);
+ pose proof (WordW.wordWToZ_log_bound v);
try omega ].
-Lemma related_Z_const : related_const related_Z Word64.interp_base_type BoundedWord64.of_word64 Word64.to_Z.
+Lemma related_Z_const : related_const related_Z WordW.interp_base_type BoundedWordW.of_wordW WordW.to_Z.
Proof. related_const_t. Qed.
-Lemma related_bounds_const : related_const related_bounds Word64.interp_base_type BoundedWord64.of_word64 ZBounds.of_word64.
+Lemma related_bounds_const : related_const related_bounds WordW.interp_base_type BoundedWordW.of_wordW ZBounds.of_wordW.
Proof. related_const_t. Qed.
-Lemma related_word64_const : related_const related_word64 Word64.interp_base_type BoundedWord64.of_word64 (fun _ x => x).
+Lemma related_wordW_const : related_const related_wordW WordW.interp_base_type BoundedWordW.of_wordW (fun _ x => x).
Proof. related_const_t. Qed.
-Local Ltac related_word64_op_t_step :=
+Local Ltac related_wordW_op_t_step :=
first [ exact I
| reflexivity
| progress intros
@@ -82,12 +83,12 @@ Local Ltac related_word64_op_t_step :=
| progress destruct_head' prod
| progress destruct_head' and
| progress destruct_head' option
- | progress destruct_head' BoundedWord64.BoundedWord
- | progress cbv [related_word64 related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWord64.BoundedWordToBounds BoundedWord64.to_bounds'] in *
+ | progress destruct_head' BoundedWordW.BoundedWord
+ | progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds'] in *
| progress simpl @fst in *
| progress simpl @snd in *
- | progress simpl @BoundedWord64.upper in *
- | progress simpl @BoundedWord64.lower in *
+ | progress simpl @BoundedWord.upper in *
+ | progress simpl @BoundedWord.lower in *
| progress break_match
| progress break_match_hyps
| congruence
@@ -95,53 +96,53 @@ Local Ltac related_word64_op_t_step :=
| [ H : ?op _ = Some _ |- _ ]
=> let H' := fresh in
rename H into H';
- first [ pose proof (@BoundedWord64.t_map1_correct _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWord64.t_map2_correct _ _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWord64.t_map4_correct _ _ _ _ _ _ H') as H; clear H' ];
+ first [ pose proof (@BoundedWordW.t_map1_correct _ _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map2_correct _ _ _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map4_correct _ _ _ _ _ _ H') as H; clear H' ];
simpl in H
| [ H : ?op _ = None |- _ ]
=> let H' := fresh in
rename H into H';
- first [ pose proof (@BoundedWord64.t_map1_correct_None _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWord64.t_map2_correct_None _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWord64.t_map4_correct_None _ _ _ _ _ H') as H; clear H' ];
+ first [ pose proof (@BoundedWordW.t_map1_correct_None _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map2_correct_None _ _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map4_correct_None _ _ _ _ _ H') as H; clear H' ];
simpl in H
end
- | progress cbv [related'_word64 proj_eq_rel BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value] in *
+ | progress cbv [related'_wordW proj_eq_rel BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value] in *
| match goal with
| [ H : ?op None _ = Some _ |- _ ] => progress simpl in H
| [ H : ?op _ None = Some _ |- _ ] => progress simpl in H
| [ H : ?op (Some _) (Some _) = Some _ |- _ ] => progress simpl in H
| [ H : ?op (Some _) (Some _) = None |- _ ] => progress simpl in H
end ].
-Local Ltac related_word64_op_t := repeat related_word64_op_t_step.
+Local Ltac related_wordW_op_t := repeat related_wordW_op_t_step.
-Lemma related_word64_t_map1 opW opB pf
+Lemma related_wordW_t_map1 opW opB pf
sv1 sv2
- : interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_word64 sv1 sv2
- -> @related_word64 TZ (BoundedWord64.t_map1 opW opB pf sv1) (opW sv2).
+ : interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_wordW sv1 sv2
+ -> @related_wordW TZ (BoundedWordW.t_map1 opW opB pf sv1) (opW sv2).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
- related_word64_op_t.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
Qed.
-Lemma related_word64_t_map2 opW opB pf
+Lemma related_wordW_t_map2 opW opB pf
sv1 sv2
- : interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_word64 sv1 sv2
- -> @related_word64 TZ (BoundedWord64.t_map2 opW opB pf (fst sv1) (snd sv1)) (opW (fst sv2) (snd sv2)).
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_wordW sv1 sv2
+ -> @related_wordW TZ (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opW (fst sv2) (snd sv2)).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
- related_word64_op_t.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
Qed.
-Lemma related_word64_t_map4 opW opB pf
+Lemma related_wordW_t_map4 opW opB pf
sv1 sv2
- : interp_flat_type_rel_pointwise2 (t:=Prod (Prod (Prod (Tbase TZ) (Tbase TZ)) (Tbase TZ)) (Tbase TZ)) related_word64 sv1 sv2
- -> @related_word64 TZ (BoundedWord64.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Prod (Prod (Tbase TZ) (Tbase TZ)) (Tbase TZ)) (Tbase TZ)) related_wordW sv1 sv2
+ -> @related_wordW TZ (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
(opW (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
- related_word64_op_t.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
Qed.
Lemma related_tuples_None_left
@@ -325,24 +326,24 @@ Qed.
Local Arguments LiftOption.lift_relation _ _ _ _ !_ _ / .
Local Arguments LiftOption.of' _ _ !_ / .
-Local Arguments BoundedWord64.BoundedWordToBounds !_ / .
+Local Arguments BoundedWordW.BoundedWordToBounds !_ / .
-Lemma related_word64_op : related_op related_word64 (@BoundedWord64.interp_op) (@Word64.interp_op).
+Lemma related_wordW_op : related_op related_wordW (@BoundedWordW.interp_op) (@WordW.interp_op).
Proof.
(let op := fresh in intros ?? op; destruct op; simpl);
- try first [ apply related_word64_t_map1
- | apply related_word64_t_map2
- | apply related_word64_t_map4 ].
+ try first [ apply related_wordW_t_map1
+ | apply related_wordW_t_map2
+ | apply related_wordW_t_map4 ].
Qed.
Lemma related_bounds_t_map1 opW opB pf
(HN : opB None = None)
sv1 sv2
: interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_bounds sv1 sv2
- -> @related_bounds TZ (BoundedWord64.t_map1 opW opB pf sv1) (opB sv2).
+ -> @related_bounds TZ (BoundedWordW.t_map1 opW opB pf sv1) (opB sv2).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
- related_word64_op_t.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
Qed.
Lemma related_bounds_t_map2 opW opB pf
@@ -350,10 +351,10 @@ Lemma related_bounds_t_map2 opW opB pf
(HN1 : forall v, opB v None = None)
sv1 sv2
: interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_bounds sv1 sv2
- -> @related_bounds TZ (BoundedWord64.t_map2 opW opB pf (fst sv1) (snd sv1)) (opB (fst sv2) (snd sv2)).
+ -> @related_bounds TZ (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opB (fst sv2) (snd sv2)).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
- related_word64_op_t.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
Qed.
Lemma related_bounds_t_map4 opW opB pf
@@ -363,17 +364,17 @@ Lemma related_bounds_t_map4 opW opB pf
(HN3 : forall x y z, opB x y z None = None)
sv1 sv2
: interp_flat_type_rel_pointwise2 (t:=Prod (Prod (Prod (Tbase TZ) (Tbase TZ)) (Tbase TZ)) (Tbase TZ)) related_bounds sv1 sv2
- -> @related_bounds TZ (BoundedWord64.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
+ -> @related_bounds TZ (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
(opB (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
destruct_head prod.
intros; destruct_head' prod.
- progress cbv [related_word64 related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWord64.BoundedWordToBounds BoundedWord64.to_bounds' proj_eq_rel] in *.
+ progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds' proj_eq_rel] in *.
destruct_head' option; destruct_head_hnf' and; destruct_head_hnf' False; subst;
try solve [ simpl; rewrite ?HN0, ?HN1, ?HN2, ?HN3; tauto ];
[].
- related_word64_op_t.
+ related_wordW_op_t.
Qed.
Local Arguments Tuple.lift_option : simpl never.
@@ -382,7 +383,7 @@ Local Arguments Tuple.map : simpl never.
Local Arguments Tuple.map2 : simpl never.
Local Arguments ZBounds.SmartBuildBounds _ _ / .
-Lemma related_bounds_op : related_op related_bounds (@BoundedWord64.interp_op) (@ZBounds.interp_op).
+Lemma related_bounds_op : related_op related_bounds (@BoundedWordW.interp_op) (@ZBounds.interp_op).
Proof.
let op := fresh in intros ?? op; destruct op; simpl.
{ apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
@@ -397,7 +398,7 @@ Proof.
{ apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
Qed.
-Local Ltac Word64.Rewrites.word64_util_arith ::=
+Local Ltac WordW.Rewrites.wordW_util_arith ::=
solve [ autorewrite with Zshift_to_pow; omega
| autorewrite with Zshift_to_pow; nia
| autorewrite with Zshift_to_pow; auto with zarith
@@ -407,7 +408,7 @@ Local Ltac Word64.Rewrites.word64_util_arith ::=
| nia
| etransitivity; [ eapply Z.div_le_mono | eapply Z.div_le_compat_l ];
auto with zarith ]
- | apply Z.land_nonneg; Word64.Rewrites.word64_util_arith
+ | apply Z.land_nonneg; WordW.Rewrites.wordW_util_arith
| eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eassumption ];
instantiate; apply Z.min_case_strong; intros;
first [ etransitivity; [ apply Z.land_upper_bound_l | ]; omega
@@ -416,20 +417,20 @@ Local Ltac Word64.Rewrites.word64_util_arith ::=
apply Z.max_case_strong; intro;
(eapply Z.le_lt_trans; [ eapply Z.log2_le_mono; eassumption | assumption ])
| eapply Z.le_lt_trans; [ eapply Z.log2_le_mono, neg_upperbound | ];
- Word64.Rewrites.word64_util_arith
+ WordW.Rewrites.wordW_util_arith
| (progress unfold ModularBaseSystemListZOperations.cmovne, ModularBaseSystemListZOperations.cmovl, ModularBaseSystemListZOperations.neg); break_match;
- Word64.Rewrites.word64_util_arith ].
+ WordW.Rewrites.wordW_util_arith ].
Local Ltac related_Z_op_t_step :=
- first [ progress related_word64_op_t_step
- | progress cbv [related'_Z proj_eq_rel BoundedWord64.to_Z' BoundedWord64.to_word64' Word64.to_Z BoundedWord64.boundedWordToWord64 BoundedWord64.value] in *
- | autorewrite with push_word64ToZ ].
+ first [ progress related_wordW_op_t_step
+ | progress cbv [related'_Z proj_eq_rel BoundedWordW.to_Z' BoundedWordW.to_wordW' WordW.to_Z BoundedWordW.boundedWordToWordW BoundedWord.value] in *
+ | autorewrite with push_wordWToZ ].
Local Ltac related_Z_op_t := repeat related_Z_op_t_step.
Local Notation is_bounded_by value lower upper
- := ((0 <= lower /\ lower <= Word64.word64ToZ value <= upper /\ Z.log2 upper < Z.of_nat Word64.bit_width)%Z)
+ := ((0 <= lower /\ lower <= WordW.wordWToZ value <= upper /\ Z.log2 upper < Z.of_nat WordW.bit_width)%Z)
(only parsing).
Local Notation is_in_bounds value bounds
- := (is_bounded_by value (ZBounds.lower bounds) (ZBounds.upper bounds))
+ := (is_bounded_by value (Bounds.lower bounds) (Bounds.upper bounds))
(only parsing).
Lemma related_Z_t_map1 opZ opW opB pf
@@ -437,12 +438,12 @@ Lemma related_Z_t_map1 opZ opW opB pf
Some brs = opB (Some bxs)
-> is_in_bounds x bxs
-> is_in_bounds (opW x) brs
- -> Word64.word64ToZ (opW x) = (opZ (Word64.word64ToZ x)))
+ -> WordW.wordWToZ (opW x) = (opZ (WordW.wordWToZ x)))
sv1 sv2
: interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_Z sv1 sv2
- -> @related_Z TZ (BoundedWord64.t_map1 opW opB pf sv1) (opZ sv2).
+ -> @related_Z TZ (BoundedWordW.t_map1 opW opB pf sv1) (opZ sv2).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
related_Z_op_t.
eapply H; eauto.
Qed.
@@ -453,12 +454,12 @@ Lemma related_Z_t_map2 opZ opW opB pf
-> is_in_bounds x bxs
-> is_in_bounds y bys
-> is_in_bounds (opW x y) brs
- -> Word64.word64ToZ (opW x y) = (opZ (Word64.word64ToZ x) (Word64.word64ToZ y)))
+ -> WordW.wordWToZ (opW x y) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y)))
sv1 sv2
: interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_Z sv1 sv2
- -> @related_Z TZ (BoundedWord64.t_map2 opW opB pf (fst sv1) (snd sv1)) (opZ (fst sv2) (snd sv2)).
+ -> @related_Z TZ (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opZ (fst sv2) (snd sv2)).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
related_Z_op_t.
eapply H; eauto.
Qed.
@@ -471,13 +472,13 @@ Lemma related_Z_t_map4 opZ opW opB pf
-> is_in_bounds z bzs
-> is_in_bounds w bws
-> is_in_bounds (opW x y z w) brs
- -> Word64.word64ToZ (opW x y z w) = (opZ (Word64.word64ToZ x) (Word64.word64ToZ y) (Word64.word64ToZ z) (Word64.word64ToZ w)))
+ -> WordW.wordWToZ (opW x y z w) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y) (WordW.wordWToZ z) (WordW.wordWToZ w)))
sv1 sv2
: interp_flat_type_rel_pointwise2 (t:=(Tbase TZ * Tbase TZ * Tbase TZ * Tbase TZ)%ctype) related_Z sv1 sv2
- -> @related_Z TZ (BoundedWord64.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
+ -> @related_Z TZ (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
(opZ (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
Proof.
- cbv [interp_flat_type BoundedWord64.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
related_Z_op_t.
eapply H; eauto.
Qed.
@@ -488,15 +489,17 @@ Local Arguments related'_Z _ _ _ / .
Local Ltac related_Z_op_fin_t_step :=
first [ progress subst
+ | progress inversion_option
+ | progress ZBounds.inversion_bounds
+ | progress destruct_head' Bounds.bounds
| progress destruct_head' ZBounds.bounds
| progress destruct_head' and
- | progress ZBounds.inversion_bounds
| progress simpl in * |-
| progress break_match_hyps
| congruence
| progress inversion_option
| intro
- | progress autorewrite with push_word64ToZ
+ | progress autorewrite with push_wordWToZ
| match goal with
| [ H : andb _ _ = true |- _ ] => rewrite Bool.andb_true_iff in H
| [ H : context[Tuple.lift_option (Tuple.push_option _)] |- _ ]
@@ -505,11 +508,11 @@ Local Ltac related_Z_op_fin_t_step :=
| progress Z.ltb_to_lt ].
Local Ltac related_Z_op_fin_t := repeat related_Z_op_fin_t_step.
-Local Opaque Word64.bit_width.
+Local Opaque WordW.bit_width.
Local Arguments ZBounds.neg _ !_ / .
-Lemma related_Z_op : related_op related_Z (@BoundedWord64.interp_op) (@Z.interp_op).
+Lemma related_Z_op : related_op related_Z (@BoundedWordW.interp_op) (@Z.interp_op).
Proof.
let op := fresh in intros ?? op; destruct op; simpl.
{ apply related_Z_t_map2; related_Z_op_fin_t. }
@@ -525,4 +528,4 @@ Proof.
Qed.
Create HintDb interp_related discriminated.
-Hint Resolve related_Z_op related_bounds_op related_word64_op related_Z_const related_bounds_const related_word64_const : interp_related.
+Hint Resolve related_Z_op related_bounds_op related_wordW_op related_Z_const related_bounds_const related_wordW_const : interp_related.
diff --git a/src/Reflection/Z/Interpretations128/RelationsCombinations.v b/src/Reflection/Z/Interpretations128/RelationsCombinations.v
new file mode 100644
index 000000000..4fde4d69c
--- /dev/null
+++ b/src/Reflection/Z/Interpretations128/RelationsCombinations.v
@@ -0,0 +1,359 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.Z.InterpretationsGen.
+Require Import Crypto.Reflection.Z.Interpretations128.
+Require Import Crypto.Reflection.Z.Interpretations128.Relations.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Prod.
+Require Import Crypto.Util.Tactics.
+
+Module Relations.
+ Section lift.
+ Context {interp_base_type1 interp_base_type2 : base_type -> Type}
+ (R : forall t, interp_base_type1 t -> interp_base_type2 t -> Prop).
+
+ Definition interp_type_rel_pointwise2_uncurried
+ {t : type base_type}
+ := match t return interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> _ with
+ | Tflat T => fun f g => interp_flat_type_rel_pointwise2 (t:=T) R f g
+ | Arrow A B
+ => fun f g
+ => forall x y, interp_flat_type_rel_pointwise2 R x y
+ -> interp_flat_type_rel_pointwise2 R (ApplyInterpedAll f x) (ApplyInterpedAll g y)
+ end.
+
+ Lemma uncurry_interp_type_rel_pointwise2
+ {t f g}
+ : interp_type_rel_pointwise2 (t:=t) R f g
+ <-> interp_type_rel_pointwise2_uncurried (t:=t) f g.
+ Proof.
+ unfold interp_type_rel_pointwise2_uncurried.
+ induction t as [|A B IHt]; [ reflexivity | ].
+ { simpl; unfold Morphisms.respectful_hetero in *; destruct B.
+ { reflexivity. }
+ { setoid_rewrite IHt; clear IHt.
+ split; intro H; intros.
+ { simpl in *; intuition. }
+ { eapply (H (_, _) (_, _)); simpl in *; intuition. } } }
+ Qed.
+ End lift.
+
+ Section proj.
+ Context {interp_base_type1 interp_base_type2}
+ (proj : forall t : base_type, interp_base_type1 t -> interp_base_type2 t).
+
+ Let R {t : flat_type base_type} f g :=
+ SmartVarfMap (t:=t) proj f = g.
+
+ Definition interp_type_rel_pointwise2_uncurried_proj
+ {t : type base_type}
+ : interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop
+ := match t return interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop with
+ | Tflat T => @R _
+ | Arrow A B
+ => fun f g
+ => forall x : interp_flat_type interp_base_type1 (all_binders_for (Arrow A B)),
+ let y := SmartVarfMap proj x in
+ let fx := ApplyInterpedAll f x in
+ let gy := ApplyInterpedAll g y in
+ @R _ fx gy
+ end.
+
+ Lemma uncurry_interp_type_rel_pointwise2_proj
+ {t : type base_type}
+ {f : interp_type interp_base_type1 t}
+ {g}
+ : interp_type_rel_pointwise2 (t:=t) (fun t => @R _) f g
+ -> interp_type_rel_pointwise2_uncurried_proj (t:=t) f g.
+ Proof.
+ unfold interp_type_rel_pointwise2_uncurried_proj.
+ induction t as [t|A B IHt]; simpl; unfold Morphisms.respectful_hetero in *.
+ { induction t as [t|A IHA B IHB]; simpl; [ solve [ trivial | reflexivity ] | ].
+ intros [HA HB].
+ specialize (IHA _ _ HA); specialize (IHB _ _ HB).
+ unfold R in *.
+ repeat first [ progress destruct_head_hnf' prod
+ | progress simpl in *
+ | progress subst
+ | reflexivity ]. }
+ { destruct B; intros H ?; apply IHt, H; clear IHt;
+ repeat first [ reflexivity
+ | progress simpl in *
+ | progress unfold R, LiftOption.of' in *
+ | progress break_match ]. }
+ Qed.
+ End proj.
+
+ Section proj_option.
+ Context {interp_base_type1 : Type} {interp_base_type2 : base_type -> Type}
+ (proj_option : forall t, interp_base_type1 -> interp_base_type2 t).
+
+ Let R {t : flat_type base_type} f g :=
+ let f' := LiftOption.of' (t:=t) f in
+ match f' with
+ | Some f' => SmartVarfMap proj_option f' = g
+ | None => True
+ end.
+
+ Definition interp_type_rel_pointwise2_uncurried_proj_option
+ {t : type base_type}
+ : interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop
+ := match t return interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop with
+ | Tflat T => @R _
+ | Arrow A B
+ => fun f g
+ => forall x : interp_flat_type (fun _ => interp_base_type1) (all_binders_for (Arrow A B)),
+ let y := SmartVarfMap proj_option x in
+ let fx := ApplyInterpedAll f (LiftOption.to' (Some x)) in
+ let gy := ApplyInterpedAll g y in
+ @R _ fx gy
+ end.
+
+ Lemma uncurry_interp_type_rel_pointwise2_proj_option
+ {t : type base_type}
+ {f : interp_type (LiftOption.interp_base_type' interp_base_type1) t}
+ {g}
+ : interp_type_rel_pointwise2 (t:=t) (fun t => @R _) f g
+ -> interp_type_rel_pointwise2_uncurried_proj_option (t:=t) f g.
+ Proof.
+ unfold interp_type_rel_pointwise2_uncurried_proj_option.
+ induction t as [t|A B IHt]; simpl; unfold Morphisms.respectful_hetero in *.
+ { induction t as [t|A IHA B IHB]; simpl; [ solve [ trivial | reflexivity ] | ].
+ intros [HA HB].
+ specialize (IHA _ _ HA); specialize (IHB _ _ HB).
+ unfold R in *.
+ repeat first [ progress destruct_head_hnf' prod
+ | progress simpl in *
+ | progress unfold LiftOption.of' in *
+ | progress break_match
+ | progress break_match_hyps
+ | progress inversion_prod
+ | progress inversion_option
+ | reflexivity
+ | progress intuition subst ]. }
+ { destruct B; intros H ?; apply IHt, H; clear IHt.
+ { repeat first [ progress simpl in *
+ | progress unfold R, LiftOption.of' in *
+ | progress break_match
+ | reflexivity ]. }
+ { simpl in *; break_match; reflexivity. } }
+ Qed.
+ End proj_option.
+
+ Section proj_option2.
+ Context {interp_base_type1 : Type} {interp_base_type2 : Type}
+ (proj : interp_base_type1 -> interp_base_type2).
+
+ Let R {t : flat_type base_type} f g :=
+ let f' := LiftOption.of' (t:=t) f in
+ let g' := LiftOption.of' (t:=t) g in
+ match f', g' with
+ | Some f', Some g' => SmartVarfMap (fun _ => proj) f' = g'
+ | None, None => True
+ | Some _, _ => False
+ | None, _ => False
+ end.
+
+ Definition interp_type_rel_pointwise2_uncurried_proj_option2
+ {t : type base_type}
+ : interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type (LiftOption.interp_base_type' interp_base_type2) t -> Prop
+ := match t return interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type (LiftOption.interp_base_type' interp_base_type2) t -> Prop with
+ | Tflat T => @R _
+ | Arrow A B
+ => fun f g
+ => forall x : interp_flat_type (fun _ => interp_base_type1) (all_binders_for (Arrow A B)),
+ let y := SmartVarfMap (fun _ => proj) x in
+ let fx := ApplyInterpedAll f (LiftOption.to' (Some x)) in
+ let gy := ApplyInterpedAll g (LiftOption.to' (Some y)) in
+ @R _ fx gy
+ end.
+
+ Lemma uncurry_interp_type_rel_pointwise2_proj_option2
+ {t : type base_type}
+ {f : interp_type (LiftOption.interp_base_type' interp_base_type1) t}
+ {g : interp_type (LiftOption.interp_base_type' interp_base_type2) t}
+ : interp_type_rel_pointwise2 (t:=t) (fun t => @R _) f g
+ -> interp_type_rel_pointwise2_uncurried_proj_option2 (t:=t) f g.
+ Proof.
+ unfold interp_type_rel_pointwise2_uncurried_proj_option2.
+ induction t as [t|A B IHt]; simpl; unfold Morphisms.respectful_hetero in *.
+ { induction t as [t|A IHA B IHB]; simpl; [ solve [ trivial | reflexivity ] | ].
+ intros [HA HB].
+ specialize (IHA _ _ HA); specialize (IHB _ _ HB).
+ unfold R in *.
+ repeat first [ progress destruct_head_hnf' prod
+ | progress simpl in *
+ | progress unfold LiftOption.of' in *
+ | progress break_match
+ | progress break_match_hyps
+ | progress inversion_prod
+ | progress inversion_option
+ | reflexivity
+ | progress intuition subst ]. }
+ { destruct B; intros H ?; apply IHt, H; clear IHt.
+ { repeat first [ progress simpl in *
+ | progress unfold R, LiftOption.of' in *
+ | progress break_match
+ | reflexivity ]. }
+ { simpl in *; break_match; reflexivity. } }
+ Qed.
+ End proj_option2.
+
+ Section proj_from_option2.
+ Context {interp_base_type0 : Type} {interp_base_type1 interp_base_type2 : base_type -> Type}
+ (proj01 : forall t, interp_base_type0 -> interp_base_type1 t)
+ (proj02 : forall t, interp_base_type0 -> interp_base_type2 t)
+ (proj : forall t, interp_base_type1 t -> interp_base_type2 t).
+
+ Let R {t : flat_type base_type} f g :=
+ proj_eq_rel (SmartVarfMap proj (t:=t)) f g.
+
+ Definition interp_type_rel_pointwise2_uncurried_proj_from_option2
+ {t : type base_type}
+ : interp_type (LiftOption.interp_base_type' interp_base_type0) t -> interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop
+ := match t return interp_type _ t -> interp_type _ t -> interp_type _ t -> Prop with
+ | Tflat T => fun f0 f g => match LiftOption.of' f0 with
+ | Some _ => True
+ | None => False
+ end -> @R _ f g
+ | Arrow A B
+ => fun f0 f g
+ => forall x : interp_flat_type (fun _ => interp_base_type0) (all_binders_for (Arrow A B)),
+ let x' := SmartVarfMap proj01 x in
+ let y' := SmartVarfMap proj x' in
+ let fx := ApplyInterpedAll f x' in
+ let gy := ApplyInterpedAll g y' in
+ let f0x := LiftOption.of' (ApplyInterpedAll f0 (LiftOption.to' (Some x))) in
+ match f0x with
+ | Some _ => True
+ | None => False
+ end
+ -> @R _ fx gy
+ end.
+
+ Lemma uncurry_interp_type_rel_pointwise2_proj_from_option2
+ {t : type base_type}
+ {f0}
+ {f : interp_type interp_base_type1 t}
+ {g : interp_type interp_base_type2 t}
+ (proj012 : forall t x, proj t (proj01 t x) = proj02 t x)
+ : interp_type_rel_pointwise2 (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj01 t))) f0 f
+ -> interp_type_rel_pointwise2 (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj02 t))) f0 g
+ -> interp_type_rel_pointwise2_uncurried_proj_from_option2 (t:=t) f0 f g.
+ Proof.
+ unfold interp_type_rel_pointwise2_uncurried_proj_from_option2.
+ induction t as [t|A B IHt]; simpl; unfold Morphisms.respectful_hetero in *.
+ { induction t as [t|A IHA B IHB]; simpl.
+ { cbv [LiftOption.lift_relation proj_eq_rel R].
+ break_match; try tauto; intros; subst.
+ apply proj012. }
+ { intros [HA HB] [HA' HB'] Hrel.
+ specialize (IHA _ _ _ HA HA'); specialize (IHB _ _ _ HB HB').
+ unfold R, proj_eq_rel in *.
+ repeat first [ progress destruct_head_hnf' prod
+ | progress simpl in *
+ | progress unfold LiftOption.of' in *
+ | progress break_match
+ | progress break_match_hyps
+ | progress inversion_prod
+ | progress inversion_option
+ | reflexivity
+ | progress intuition subst ]. } }
+ { destruct B; intros H0 H1 ?; apply IHt; clear IHt; first [ apply H0 | apply H1 ];
+ repeat first [ progress simpl in *
+ | progress unfold R, LiftOption.of', proj_eq_rel, LiftOption.lift_relation in *
+ | break_match; rewrite <- ?proj012; reflexivity ]. }
+ Qed.
+ End proj_from_option2.
+ Global Arguments uncurry_interp_type_rel_pointwise2_proj_from_option2 {_ _ _ _ _} proj {t f0 f g} _ _ _.
+
+ Section proj1_from_option2.
+ Context {interp_base_type0 interp_base_type1 : Type} {interp_base_type2 : base_type -> Type}
+ (proj01 : interp_base_type0 -> interp_base_type1)
+ (proj02 : forall t, interp_base_type0 -> interp_base_type2 t)
+ (R : forall t, interp_base_type1 -> interp_base_type2 t -> Prop).
+
+ Definition interp_type_rel_pointwise2_uncurried_proj1_from_option2
+ {t : type base_type}
+ : interp_type (LiftOption.interp_base_type' interp_base_type0) t -> interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop
+ := match t return interp_type _ t -> interp_type _ t -> interp_type _ t -> Prop with
+ | Tflat T => fun f0 f g => match LiftOption.of' f0 with
+ | Some _ => True
+ | None => False
+ end -> match LiftOption.of' f with
+ | Some f' => interp_flat_type_rel_pointwise2 (@R) f' g
+ | None => True
+ end
+ | Arrow A B
+ => fun f0 f g
+ => forall x : interp_flat_type (fun _ => interp_base_type0) (all_binders_for (Arrow A B)),
+ let x' := SmartVarfMap (fun _ => proj01) x in
+ let y' := SmartVarfMap proj02 x in
+ let fx := LiftOption.of' (ApplyInterpedAll f (LiftOption.to' (Some x'))) in
+ let gy := ApplyInterpedAll g y' in
+ let f0x := LiftOption.of' (ApplyInterpedAll f0 (LiftOption.to' (Some x))) in
+ match f0x with
+ | Some _ => True
+ | None => False
+ end
+ -> match fx with
+ | Some fx' => interp_flat_type_rel_pointwise2 (@R) fx' gy
+ | None => True
+ end
+ end.
+
+ Lemma uncurry_interp_type_rel_pointwise2_proj1_from_option2
+ {t : type base_type}
+ {f0}
+ {f : interp_type (LiftOption.interp_base_type' interp_base_type1) t}
+ {g : interp_type interp_base_type2 t}
+ (proj012R : forall t x, @R _ (proj01 x) (proj02 t x))
+ : interp_type_rel_pointwise2 (t:=t) (LiftOption.lift_relation2 (proj_eq_rel proj01)) f0 f
+ -> interp_type_rel_pointwise2 (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj02 t))) f0 g
+ -> interp_type_rel_pointwise2_uncurried_proj1_from_option2 (t:=t) f0 f g.
+ Proof.
+ unfold interp_type_rel_pointwise2_uncurried_proj1_from_option2.
+ induction t as [t|A B IHt]; simpl; unfold Morphisms.respectful_hetero in *.
+ { induction t as [t|A IHA B IHB]; simpl.
+ { cbv [LiftOption.lift_relation proj_eq_rel LiftOption.lift_relation2].
+ break_match; try tauto; intros; subst.
+ apply proj012R. }
+ { intros [HA HB] [HA' HB'] Hrel.
+ specialize (IHA _ _ _ HA HA'); specialize (IHB _ _ _ HB HB').
+ unfold proj_eq_rel in *.
+ repeat first [ progress destruct_head_hnf' prod
+ | progress simpl in *
+ | progress unfold LiftOption.of' in *
+ | progress break_match
+ | progress break_match_hyps
+ | progress inversion_prod
+ | progress inversion_option
+ | reflexivity
+ | progress intuition subst ]. } }
+ { destruct B; intros H0 H1 ?; apply IHt; clear IHt; first [ apply H0 | apply H1 ];
+ repeat first [ progress simpl in *
+ | progress unfold R, LiftOption.of', proj_eq_rel, LiftOption.lift_relation in *
+ | break_match; reflexivity ]. }
+ Qed.
+ End proj1_from_option2.
+ Global Arguments uncurry_interp_type_rel_pointwise2_proj1_from_option2 {_ _ _ _ _} R {t f0 f g} _ _ _.
+
+ Section combine_related.
+ Lemma related_flat_transitivity {interp_base_type1 interp_base_type2 interp_base_type3}
+ {R1 : forall t : base_type, interp_base_type1 t -> interp_base_type2 t -> Prop}
+ {R2 : forall t : base_type, interp_base_type1 t -> interp_base_type3 t -> Prop}
+ {R3 : forall t : base_type, interp_base_type2 t -> interp_base_type3 t -> Prop}
+ {t x y z}
+ : (forall t a b c, (R1 t a b : Prop) -> (R2 t a c : Prop) -> (R3 t b c : Prop))
+ -> interp_flat_type_rel_pointwise2 (t:=t) R1 x y
+ -> interp_flat_type_rel_pointwise2 (t:=t) R2 x z
+ -> interp_flat_type_rel_pointwise2 (t:=t) R3 y z.
+ Proof.
+ intro HRel; induction t; simpl; intuition eauto.
+ Qed.
+ End combine_related.
+End Relations.
diff --git a/src/Reflection/Z/Interpretations64.v b/src/Reflection/Z/Interpretations64.v
new file mode 100644
index 000000000..712cc96e0
--- /dev/null
+++ b/src/Reflection/Z/Interpretations64.v
@@ -0,0 +1,13 @@
+(** * Interpretation of PHOAS syntax for expression trees on ℤ *)
+Require Import Coq.ZArith.ZArith.
+Require Export Crypto.Reflection.Z.InterpretationsGen.
+Export Reflection.Syntax.Notations.
+
+Module BitSize64 <: BitSize.
+ Definition bit_width : nat := 64%nat.
+ Lemma bit_width_pos : (0 < Z.of_nat bit_width)%Z.
+ Proof. simpl; omega. Qed.
+End BitSize64.
+
+Module Interpretations64 := InterpretationsGen BitSize64.
+Include Interpretations64.
diff --git a/src/Reflection/Z/Interpretations64/Relations.v b/src/Reflection/Z/Interpretations64/Relations.v
new file mode 100644
index 000000000..80f3d139c
--- /dev/null
+++ b/src/Reflection/Z/Interpretations64/Relations.v
@@ -0,0 +1,531 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Coq.micromega.Psatz.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.Z.InterpretationsGen.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemListZOperationsProofs.
+Require Import Crypto.Util.Option.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Prod.
+Require Import Crypto.Util.Tactics.
+
+Definition proj_eq_rel {A B} (proj : A -> B) (x : A) (y : B) : Prop
+ := proj x = y.
+Definition related'_Z {t} (x : BoundedWordW.BoundedWord) (y : Z.interp_base_type t) : Prop
+ := proj_eq_rel (BoundedWordW.to_Z' _) x y.
+Definition related_Z t : BoundedWordW.interp_base_type t -> Z.interp_base_type t -> Prop
+ := LiftOption.lift_relation (@related'_Z) t.
+Definition related'_wordW {t} (x : BoundedWordW.BoundedWord) (y : WordW.interp_base_type t) : Prop
+ := proj_eq_rel (BoundedWordW.to_wordW' _) x y.
+Definition related_wordW t : BoundedWordW.interp_base_type t -> WordW.interp_base_type t -> Prop
+ := LiftOption.lift_relation (@related'_wordW) t.
+Definition related_bounds t : BoundedWordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
+ := LiftOption.lift_relation2 (proj_eq_rel BoundedWordW.BoundedWordToBounds) t.
+
+Definition related_wordW_Z t : WordW.interp_base_type t -> Z.interp_base_type t -> Prop
+ := proj_eq_rel (WordW.to_Z _).
+
+Definition related'_wordW_bounds : WordW.wordW -> ZBounds.bounds -> Prop
+ := fun value b => (0 <= Bounds.lower b /\ Bounds.lower b <= WordW.wordWToZ value <= Bounds.upper b /\ Z.log2 (Bounds.upper b) < Z.of_nat WordW.bit_width)%Z.
+Definition related_wordW_bounds : WordW.wordW -> ZBounds.t -> Prop
+ := fun value b => match b with
+ | Some b => related'_wordW_bounds value b
+ | None => True
+ end.
+Definition related_wordW_boundsi (t : base_type) : WordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
+ := match t with
+ | TZ => related_wordW_bounds
+ end.
+Definition related_wordW_boundsi' (t : base_type) : ZBounds.bounds -> WordW.interp_base_type t -> Prop
+ := match t return ZBounds.bounds -> WordW.interp_base_type t -> Prop with
+ | TZ => fun x y => related'_wordW_bounds y x
+ end.
+
+Local Notation related_op R interp_op1 interp_op2
+ := (forall (src dst : flat_type base_type) (op : op src dst)
+ (sv1 : interp_flat_type _ src) (sv2 : interp_flat_type _ src),
+ interp_flat_type_rel_pointwise2 R sv1 sv2 ->
+ interp_flat_type_rel_pointwise2 R (interp_op1 _ _ op sv1) (interp_op2 _ _ op sv2))
+ (only parsing).
+Local Notation related_const R interp f g
+ := (forall (t : base_type) (v : interp t), R t (f t v) (g t v))
+ (only parsing).
+
+Local Ltac related_const_t :=
+ let v := fresh in
+ let t := fresh in
+ intros t v; destruct t; intros; simpl in *; hnf; simpl;
+ cbv [BoundedWordW.wordWToBoundedWord related'_Z LiftOption.of' related_Z related_wordW related'_wordW proj_eq_rel] in *;
+ break_innermost_match; simpl;
+ first [ tauto
+ | Z.ltb_to_lt;
+ pose proof (WordW.wordWToZ_log_bound v);
+ try omega ].
+
+Lemma related_Z_const : related_const related_Z WordW.interp_base_type BoundedWordW.of_wordW WordW.to_Z.
+Proof. related_const_t. Qed.
+Lemma related_bounds_const : related_const related_bounds WordW.interp_base_type BoundedWordW.of_wordW ZBounds.of_wordW.
+Proof. related_const_t. Qed.
+Lemma related_wordW_const : related_const related_wordW WordW.interp_base_type BoundedWordW.of_wordW (fun _ x => x).
+Proof. related_const_t. Qed.
+
+Local Ltac related_wordW_op_t_step :=
+ first [ exact I
+ | reflexivity
+ | progress intros
+ | progress inversion_option
+ | progress ZBounds.inversion_bounds
+ | progress subst
+ | progress destruct_head' False
+ | progress destruct_head' prod
+ | progress destruct_head' and
+ | progress destruct_head' option
+ | progress destruct_head' BoundedWordW.BoundedWord
+ | progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds'] in *
+ | progress simpl @fst in *
+ | progress simpl @snd in *
+ | progress simpl @BoundedWord.upper in *
+ | progress simpl @BoundedWord.lower in *
+ | progress break_match
+ | progress break_match_hyps
+ | congruence
+ | match goal with
+ | [ H : ?op _ = Some _ |- _ ]
+ => let H' := fresh in
+ rename H into H';
+ first [ pose proof (@BoundedWordW.t_map1_correct _ _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map2_correct _ _ _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map4_correct _ _ _ _ _ _ H') as H; clear H' ];
+ simpl in H
+ | [ H : ?op _ = None |- _ ]
+ => let H' := fresh in
+ rename H into H';
+ first [ pose proof (@BoundedWordW.t_map1_correct_None _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map2_correct_None _ _ _ _ _ H') as H; clear H'
+ | pose proof (@BoundedWordW.t_map4_correct_None _ _ _ _ _ H') as H; clear H' ];
+ simpl in H
+ end
+ | progress cbv [related'_wordW proj_eq_rel BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value] in *
+ | match goal with
+ | [ H : ?op None _ = Some _ |- _ ] => progress simpl in H
+ | [ H : ?op _ None = Some _ |- _ ] => progress simpl in H
+ | [ H : ?op (Some _) (Some _) = Some _ |- _ ] => progress simpl in H
+ | [ H : ?op (Some _) (Some _) = None |- _ ] => progress simpl in H
+ end ].
+Local Ltac related_wordW_op_t := repeat related_wordW_op_t_step.
+
+Lemma related_wordW_t_map1 opW opB pf
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_wordW sv1 sv2
+ -> @related_wordW TZ (BoundedWordW.t_map1 opW opB pf sv1) (opW sv2).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
+Qed.
+
+Lemma related_wordW_t_map2 opW opB pf
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_wordW sv1 sv2
+ -> @related_wordW TZ (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opW (fst sv2) (snd sv2)).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
+Qed.
+
+Lemma related_wordW_t_map4 opW opB pf
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Prod (Prod (Tbase TZ) (Tbase TZ)) (Tbase TZ)) (Tbase TZ)) related_wordW sv1 sv2
+ -> @related_wordW TZ (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
+ (opW (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
+Qed.
+
+Lemma related_tuples_None_left
+ n T interp_base_type'
+ (R : forall t, LiftOption.interp_base_type' T t -> interp_base_type' t -> Prop)
+ (RNone : forall v, R TZ None v)
+ (v : interp_flat_type interp_base_type' (tuple (Tbase TZ) (S n)))
+ : interp_flat_type_rel_pointwise2
+ R
+ (flat_interp_untuple' (T:=Tbase TZ) (Tuple.push_option (n:=S n) None))
+ v.
+Proof.
+ induction n; simpl; intuition.
+Qed.
+
+Lemma related_tuples_Some_left
+ n T interp_base_type'
+ (R : forall t, T -> interp_base_type' t -> Prop)
+ u
+ (v : interp_flat_type interp_base_type' (tuple (Tbase TZ) (S n)))
+ : interp_flat_type_rel_pointwise2
+ R
+ (flat_interp_untuple' (T:=Tbase TZ) u)
+ v
+ <-> interp_flat_type_rel_pointwise2
+ (LiftOption.lift_relation R)
+ (flat_interp_untuple' (T:=Tbase TZ) (Tuple.push_option (n:=S n) (Some u)))
+ v.
+Proof.
+ induction n; [ reflexivity | ].
+ simpl in *; rewrite <- IHn; clear IHn.
+ reflexivity.
+Qed.
+
+Lemma related_tuples_Some_left_ext
+ {n T interp_base_type'}
+ {R : forall t, T -> interp_base_type' t -> Prop}
+ {u v u'}
+ (H : Tuple.lift_option (flat_interp_tuple (T:=Tbase TZ) (n:=S n) u) = Some u')
+ : interp_flat_type_rel_pointwise2
+ R
+ (flat_interp_untuple' (T:=Tbase TZ) u') v
+ <-> interp_flat_type_rel_pointwise2
+ (LiftOption.lift_relation R)
+ u v.
+Proof.
+ induction n.
+ { simpl in *; subst; reflexivity. }
+ { destruct_head_hnf' prod.
+ simpl in H; break_match_hyps; inversion_option; inversion_prod; subst.
+ simpl; rewrite <- IHn by eassumption; clear IHn.
+ reflexivity. }
+Qed.
+
+Lemma related_tuples_proj_eq_rel_untuple
+ {n T interp_base_type'}
+ {proj : forall t, T -> interp_base_type' t}
+ {u : Tuple.tuple _ (S n)} {v : Tuple.tuple _ (S n)}
+ : interp_flat_type_rel_pointwise2
+ (fun t => proj_eq_rel (proj t))
+ (flat_interp_untuple' (T:=Tbase TZ) u)
+ (flat_interp_untuple' (T:=Tbase TZ) v)
+ <-> (Tuple.map (proj _) u = v).
+Proof.
+ induction n; [ reflexivity | ].
+ destruct_head_hnf' prod.
+ simpl @Tuple.tuple.
+ rewrite !Tuple.map_S, path_prod_uncurried_iff, <- prod_iff_and; unfold fst, snd.
+ rewrite <- IHn.
+ reflexivity.
+Qed.
+
+Lemma related_tuples_proj_eq_rel_tuple
+ {n T interp_base_type'}
+ {proj : forall t, T -> interp_base_type' t}
+ {u v}
+ : interp_flat_type_rel_pointwise2
+ (fun t => proj_eq_rel (proj t))
+ u v
+ <-> (Tuple.map (proj _) (flat_interp_tuple (n:=S n) (T:=Tbase TZ) u)
+ = flat_interp_tuple (T:=Tbase TZ) v).
+Proof.
+ rewrite <- related_tuples_proj_eq_rel_untuple, !flat_interp_untuple'_tuple; reflexivity.
+Qed.
+
+Local Arguments LiftOption.lift_relation2 _ _ _ _ !_ !_ / .
+Lemma related_tuples_lift_relation2_untuple'
+ n T U
+ (R : T -> U -> Prop)
+ (t : option (Tuple.tuple T (S n)))
+ (u : option (Tuple.tuple U (S n)))
+ : interp_flat_type_rel_pointwise2
+ (LiftOption.lift_relation2 R)
+ (flat_interp_untuple' (T:=Tbase TZ) (Tuple.push_option t))
+ (flat_interp_untuple' (T:=Tbase TZ) (Tuple.push_option u))
+ <-> LiftOption.lift_relation2
+ (interp_flat_type_rel_pointwise2 (fun _ => R))
+ TZ
+ (option_map (flat_interp_untuple' (interp_base_type:=fun _ => T) (T:=Tbase TZ)) t)
+ (option_map (flat_interp_untuple' (interp_base_type:=fun _ => U) (T:=Tbase TZ)) u).
+Proof.
+ induction n.
+ { destruct_head' option; reflexivity. }
+ { specialize (IHn (option_map (@fst _ _) t) (option_map (@fst _ _) u)).
+ destruct_head' option;
+ destruct_head_hnf' prod;
+ simpl @option_map in *;
+ simpl @LiftOption.lift_relation2 in *;
+ try (rewrite <- IHn; reflexivity);
+ try (simpl @interp_flat_type_rel_pointwise2; tauto). }
+Qed.
+
+Lemma related_tuples_lift_relation2_untuple'_ext
+ {n T U}
+ {R : T -> U -> Prop}
+ {t u}
+ (H : (exists v, Tuple.lift_option (n:=S n) (flat_interp_tuple (T:=Tbase TZ) t) = Some v)
+ \/ (exists v, Tuple.lift_option (n:=S n) (flat_interp_tuple (T:=Tbase TZ) u) = Some v))
+ : interp_flat_type_rel_pointwise2
+ (LiftOption.lift_relation2 R)
+ t u
+ <-> LiftOption.lift_relation2
+ (interp_flat_type_rel_pointwise2 (fun _ => R))
+ TZ
+ (option_map (flat_interp_untuple' (interp_base_type:=fun _ => T) (T:=Tbase TZ)) (Tuple.lift_option (flat_interp_tuple (T:=Tbase TZ) t)))
+ (option_map (flat_interp_untuple' (interp_base_type:=fun _ => U) (T:=Tbase TZ)) (Tuple.lift_option (flat_interp_tuple (T:=Tbase TZ) u))).
+Proof.
+ induction n.
+ { destruct_head_hnf' option; reflexivity. }
+ { specialize (IHn (fst t) (fst u)).
+ lazymatch type of IHn with
+ | ?T -> _ => let H := fresh in assert (H : T); [ | specialize (IHn H); clear H ]
+ end.
+ { destruct_head' or; [ left | right ]; destruct_head' ex; destruct_head_hnf' prod; eexists;
+ (etransitivity;
+ [ | first [ refine (f_equal (option_map (@fst _ _)) (_ : _ = Some (_, _))); eassumption
+ | refine (f_equal (option_map (@snd _ _)) (_ : _ = Some (_, _))); eassumption ] ]);
+ instantiate; simpl in *; break_match; simpl in *; congruence. }
+ destruct_head_hnf' prod;
+ destruct_head_hnf' option;
+ simpl @fst in *; simpl @snd in *;
+ (etransitivity; [ simpl @interp_flat_type_rel_pointwise2 | reflexivity ]);
+ try solve [ repeat first [ progress simpl in *
+ | tauto
+ | congruence
+ | progress destruct_head ex
+ | progress destruct_head or
+ | progress break_match ] ]. }
+Qed.
+
+Lemma lift_option_flat_interp_tuple'
+ {n T x y}
+ : (Tuple.lift_option (n:=S n) (A:=T) (flat_interp_tuple' (interp_base_type:=LiftOption.interp_base_type' _) (T:=Tbase TZ) x) = Some y)
+ <-> (x = flat_interp_untuple' (T:=Tbase TZ) (n:=n) (Tuple.push_option (n:=S n) (Some y))).
+Proof.
+ rewrite Tuple.push_lift_option; generalize (Tuple.push_option (Some y)); intro.
+ split; intro; subst;
+ rewrite ?flat_interp_tuple'_untuple', ?flat_interp_untuple'_tuple';
+ reflexivity.
+Qed.
+
+Lemma lift_option_None_interp_flat_type_rel_pointwise2_1
+ T U n R x y
+ (H : interp_flat_type_rel_pointwise2 (LiftOption.lift_relation2 R) x y)
+ (HNone : Tuple.lift_option (A:=T) (n:=S n) (flat_interp_tuple' (T:=Tbase TZ) (n:=n) x) = None)
+ : Tuple.lift_option (A:=U) (n:=S n) (flat_interp_tuple' (T:=Tbase TZ) (n:=n) y) = None.
+Proof.
+ induction n; [ | specialize (IHn (fst x) (fst y) (proj1 H)) ];
+ repeat first [ progress destruct_head_hnf' False
+ | reflexivity
+ | progress inversion_option
+ | progress simpl in *
+ | progress subst
+ | progress specialize_by congruence
+ | progress destruct_head_hnf' prod
+ | progress destruct_head_hnf' and
+ | progress destruct_head_hnf' option
+ | progress break_match
+ | progress break_match_hyps ].
+Qed.
+
+Local Arguments LiftOption.lift_relation _ _ _ _ !_ _ / .
+Local Arguments LiftOption.of' _ _ !_ / .
+Local Arguments BoundedWordW.BoundedWordToBounds !_ / .
+
+Lemma related_wordW_op : related_op related_wordW (@BoundedWordW.interp_op) (@WordW.interp_op).
+Proof.
+ (let op := fresh in intros ?? op; destruct op; simpl);
+ try first [ apply related_wordW_t_map1
+ | apply related_wordW_t_map2
+ | apply related_wordW_t_map4 ].
+Qed.
+
+Lemma related_bounds_t_map1 opW opB pf
+ (HN : opB None = None)
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_bounds sv1 sv2
+ -> @related_bounds TZ (BoundedWordW.t_map1 opW opB pf sv1) (opB sv2).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
+Qed.
+
+Lemma related_bounds_t_map2 opW opB pf
+ (HN0 : forall v, opB None v = None)
+ (HN1 : forall v, opB v None = None)
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_bounds sv1 sv2
+ -> @related_bounds TZ (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opB (fst sv2) (snd sv2)).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_wordW_op_t.
+Qed.
+
+Lemma related_bounds_t_map4 opW opB pf
+ (HN0 : forall x y z, opB None x y z = None)
+ (HN1 : forall x y z, opB x None y z = None)
+ (HN2 : forall x y z, opB x y None z = None)
+ (HN3 : forall x y z, opB x y z None = None)
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Prod (Prod (Tbase TZ) (Tbase TZ)) (Tbase TZ)) (Tbase TZ)) related_bounds sv1 sv2
+ -> @related_bounds TZ (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
+ (opB (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ destruct_head prod.
+ intros; destruct_head' prod.
+ progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds' proj_eq_rel] in *.
+ destruct_head' option; destruct_head_hnf' and; destruct_head_hnf' False; subst;
+ try solve [ simpl; rewrite ?HN0, ?HN1, ?HN2, ?HN3; tauto ];
+ [].
+ related_wordW_op_t.
+Qed.
+
+Local Arguments Tuple.lift_option : simpl never.
+Local Arguments Tuple.push_option : simpl never.
+Local Arguments Tuple.map : simpl never.
+Local Arguments Tuple.map2 : simpl never.
+
+Local Arguments ZBounds.SmartBuildBounds _ _ / .
+Lemma related_bounds_op : related_op related_bounds (@BoundedWordW.interp_op) (@ZBounds.interp_op).
+Proof.
+ let op := fresh in intros ?? op; destruct op; simpl.
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map1; intros; destruct_head' option; unfold ZBounds.neg; break_match; reflexivity. }
+ { apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
+ { apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
+Qed.
+
+Local Ltac WordW.Rewrites.wordW_util_arith ::=
+ solve [ autorewrite with Zshift_to_pow; omega
+ | autorewrite with Zshift_to_pow; nia
+ | autorewrite with Zshift_to_pow; auto with zarith
+ | eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eassumption ];
+ autorewrite with Zshift_to_pow; auto using Z.mul_le_mono_nonneg with zarith;
+ solve [ omega
+ | nia
+ | etransitivity; [ eapply Z.div_le_mono | eapply Z.div_le_compat_l ];
+ auto with zarith ]
+ | apply Z.land_nonneg; WordW.Rewrites.wordW_util_arith
+ | eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eassumption ];
+ instantiate; apply Z.min_case_strong; intros;
+ first [ etransitivity; [ apply Z.land_upper_bound_l | ]; omega
+ | etransitivity; [ apply Z.land_upper_bound_r | ]; omega ]
+ | rewrite Z.log2_lor by omega;
+ apply Z.max_case_strong; intro;
+ (eapply Z.le_lt_trans; [ eapply Z.log2_le_mono; eassumption | assumption ])
+ | eapply Z.le_lt_trans; [ eapply Z.log2_le_mono, neg_upperbound | ];
+ WordW.Rewrites.wordW_util_arith
+ | (progress unfold ModularBaseSystemListZOperations.cmovne, ModularBaseSystemListZOperations.cmovl, ModularBaseSystemListZOperations.neg); break_match;
+ WordW.Rewrites.wordW_util_arith ].
+Local Ltac related_Z_op_t_step :=
+ first [ progress related_wordW_op_t_step
+ | progress cbv [related'_Z proj_eq_rel BoundedWordW.to_Z' BoundedWordW.to_wordW' WordW.to_Z BoundedWordW.boundedWordToWordW BoundedWord.value] in *
+ | autorewrite with push_wordWToZ ].
+Local Ltac related_Z_op_t := repeat related_Z_op_t_step.
+
+Local Notation is_bounded_by value lower upper
+ := ((0 <= lower /\ lower <= WordW.wordWToZ value <= upper /\ Z.log2 upper < Z.of_nat WordW.bit_width)%Z)
+ (only parsing).
+Local Notation is_in_bounds value bounds
+ := (is_bounded_by value (Bounds.lower bounds) (Bounds.upper bounds))
+ (only parsing).
+
+Lemma related_Z_t_map1 opZ opW opB pf
+ (H : forall x bxs brs,
+ Some brs = opB (Some bxs)
+ -> is_in_bounds x bxs
+ -> is_in_bounds (opW x) brs
+ -> WordW.wordWToZ (opW x) = (opZ (WordW.wordWToZ x)))
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Tbase TZ) related_Z sv1 sv2
+ -> @related_Z TZ (BoundedWordW.t_map1 opW opB pf sv1) (opZ sv2).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_Z_op_t.
+ eapply H; eauto.
+Qed.
+
+Lemma related_Z_t_map2 opZ opW opB pf
+ (H : forall x y bxs bys brs,
+ Some brs = opB (Some bxs) (Some bys)
+ -> is_in_bounds x bxs
+ -> is_in_bounds y bys
+ -> is_in_bounds (opW x y) brs
+ -> WordW.wordWToZ (opW x y) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y)))
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=Prod (Tbase TZ) (Tbase TZ)) related_Z sv1 sv2
+ -> @related_Z TZ (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opZ (fst sv2) (snd sv2)).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_Z_op_t.
+ eapply H; eauto.
+Qed.
+
+Lemma related_Z_t_map4 opZ opW opB pf
+ (H : forall x y z w bxs bys bzs bws brs,
+ Some brs = opB (Some bxs) (Some bys) (Some bzs) (Some bws)
+ -> is_in_bounds x bxs
+ -> is_in_bounds y bys
+ -> is_in_bounds z bzs
+ -> is_in_bounds w bws
+ -> is_in_bounds (opW x y z w) brs
+ -> WordW.wordWToZ (opW x y z w) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y) (WordW.wordWToZ z) (WordW.wordWToZ w)))
+ sv1 sv2
+ : interp_flat_type_rel_pointwise2 (t:=(Tbase TZ * Tbase TZ * Tbase TZ * Tbase TZ)%ctype) related_Z sv1 sv2
+ -> @related_Z TZ (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
+ (opZ (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
+Proof.
+ cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise2 interp_flat_type_rel_pointwise2_gen_Prop] in *.
+ related_Z_op_t.
+ eapply H; eauto.
+Qed.
+
+Local Arguments related_Z _ !_ _ / .
+
+Local Arguments related'_Z _ _ _ / .
+
+Local Ltac related_Z_op_fin_t_step :=
+ first [ progress subst
+ | progress inversion_option
+ | progress ZBounds.inversion_bounds
+ | progress destruct_head' Bounds.bounds
+ | progress destruct_head' ZBounds.bounds
+ | progress destruct_head' and
+ | progress simpl in * |-
+ | progress break_match_hyps
+ | congruence
+ | progress inversion_option
+ | intro
+ | progress autorewrite with push_wordWToZ
+ | match goal with
+ | [ H : andb _ _ = true |- _ ] => rewrite Bool.andb_true_iff in H
+ | [ H : context[Tuple.lift_option (Tuple.push_option _)] |- _ ]
+ => rewrite Tuple.lift_push_option in H
+ end
+ | progress Z.ltb_to_lt ].
+Local Ltac related_Z_op_fin_t := repeat related_Z_op_fin_t_step.
+
+Local Opaque WordW.bit_width.
+
+Local Arguments ZBounds.neg _ !_ / .
+
+Lemma related_Z_op : related_op related_Z (@BoundedWordW.interp_op) (@Z.interp_op).
+Proof.
+ let op := fresh in intros ?? op; destruct op; simpl.
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map2; related_Z_op_fin_t. }
+ { apply related_Z_t_map1; related_Z_op_fin_t. }
+ { apply related_Z_t_map4; related_Z_op_fin_t. }
+ { apply related_Z_t_map4; related_Z_op_fin_t. }
+Qed.
+
+Create HintDb interp_related discriminated.
+Hint Resolve related_Z_op related_bounds_op related_wordW_op related_Z_const related_bounds_const related_wordW_const : interp_related.
diff --git a/src/Reflection/Z/Interpretations/RelationsCombinations.v b/src/Reflection/Z/Interpretations64/RelationsCombinations.v
index e8ba22e00..b1287c7a7 100644
--- a/src/Reflection/Z/Interpretations/RelationsCombinations.v
+++ b/src/Reflection/Z/Interpretations64/RelationsCombinations.v
@@ -2,8 +2,9 @@ Require Import Coq.ZArith.ZArith.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Import Crypto.Reflection.Z.Interpretations.Relations.
+Require Import Crypto.Reflection.Z.InterpretationsGen.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Import Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Util.Option.
Require Import Crypto.Util.Bool.
Require Import Crypto.Util.ZUtil.
diff --git a/src/Reflection/Z/InterpretationsGen.v b/src/Reflection/Z/InterpretationsGen.v
index 6805be6a0..c24506cc1 100644
--- a/src/Reflection/Z/InterpretationsGen.v
+++ b/src/Reflection/Z/InterpretationsGen.v
@@ -25,6 +25,18 @@ Module Type BitSize.
Axiom bit_width_pos : (0 < Z.of_nat bit_width)%Z.
End BitSize.
+Module Import Bounds.
+ Record bounds := { lower : Z ; upper : Z }.
+End Bounds.
+Module Import BoundedWord.
+ Record BoundedWordGen wordW (is_bounded_by : _ -> _ -> _ -> Prop) :=
+ { lower : Z ; value : wordW ; upper : Z ;
+ in_bounds : is_bounded_by value lower upper }.
+ Global Arguments lower {_ _} _.
+ Global Arguments value {_ _} _.
+ Global Arguments upper {_ _} _.
+ Global Arguments in_bounds {_ _} _.
+End BoundedWord.
Module InterpretationsGen (Bit : BitSize).
Module Z.
Definition interp_base_type (t : base_type) : Type := interp_base_type t.
@@ -94,7 +106,7 @@ Module InterpretationsGen (Bit : BitSize).
End LiftOption.
Module WordW.
- Include BitSize.
+ Include Bit.
Definition wordW := word bit_width.
Delimit Scope wordW_scope with wordW.
Bind Scope wordW_scope with wordW.
@@ -304,7 +316,8 @@ Module InterpretationsGen (Bit : BitSize).
End WordW.
Module ZBounds.
- Record bounds := { lower : Z ; upper : Z }.
+ Export Bounds.
+ Definition bounds := bounds.
Bind Scope bounds_scope with bounds.
Definition t := option bounds. (* TODO?: Separate out the bounds computation from the overflow computation? e.g., have [safety := in_bounds | overflow] and [t := bounds * safety]? *)
Bind Scope bounds_scope with t.
@@ -432,12 +445,11 @@ Module InterpretationsGen (Bit : BitSize).
End ZBounds.
Module BoundedWordW.
+ Export BoundedWord.
Local Notation is_bounded_by value lower upper
:= ((0 <= lower /\ lower <= WordW.wordWToZ value <= upper /\ Z.log2 upper < Z.of_nat WordW.bit_width)%Z)
(only parsing).
- Record BoundedWord :=
- { lower : Z ; value : WordW.wordW ; upper : Z ;
- in_bounds : is_bounded_by value lower upper }.
+ Definition BoundedWord := BoundedWordGen WordW.wordW (fun value lower upper => is_bounded_by value lower upper).
Bind Scope bounded_word_scope with BoundedWord.
Definition t := option BoundedWord.
Bind Scope bounded_word_scope with t.
@@ -498,7 +510,7 @@ Module InterpretationsGen (Bit : BitSize).
:= fun x => WordW.to_Z _ (to_wordW _ x).
Definition BoundedWordToBounds (x : BoundedWord) : ZBounds.bounds
- := {| ZBounds.lower := lower x ; ZBounds.upper := upper x |}.
+ := {| Bounds.lower := lower x ; Bounds.upper := upper x |}.
Definition to_bounds' : t -> ZBounds.t
:= option_map BoundedWordToBounds.
@@ -513,7 +525,7 @@ Module InterpretationsGen (Bit : BitSize).
(opB : ZBounds.t -> ZBounds.t)
(pf : forall x l u,
opB (Some (BoundedWordToBounds x))
- = Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ = Some {| Bounds.lower := l ; Bounds.upper := u |}
-> let val := opW (value x) in
is_bounded_by val l u)
: t -> t
@@ -523,7 +535,7 @@ Module InterpretationsGen (Bit : BitSize).
=> match opB (Some (BoundedWordToBounds x))
as bop return opB (Some (BoundedWordToBounds x)) = bop -> t
with
- | Some (ZBounds.Build_bounds l u)
+ | Some (Bounds.Build_bounds l u)
=> fun Heq => Some {| lower := l ; value := opW (value x) ; upper := u;
in_bounds := pf _ _ _ Heq |}
| None => fun _ => None
@@ -536,7 +548,7 @@ Module InterpretationsGen (Bit : BitSize).
(opB : ZBounds.t -> ZBounds.t -> ZBounds.t)
(pf : forall x y l u,
opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- = Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ = Some {| Bounds.lower := l ; Bounds.upper := u |}
-> let val := opW (value x) (value y) in
is_bounded_by val l u)
: t -> t -> t
@@ -546,7 +558,7 @@ Module InterpretationsGen (Bit : BitSize).
=> match opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
as bop return opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = bop -> t
with
- | Some (ZBounds.Build_bounds l u)
+ | Some (Bounds.Build_bounds l u)
=> fun Heq => Some {| lower := l ; value := opW (value x) (value y) ; upper := u;
in_bounds := pf _ _ _ _ Heq |}
| None => fun _ => None
@@ -559,7 +571,7 @@ Module InterpretationsGen (Bit : BitSize).
(opB : ZBounds.t -> ZBounds.t -> ZBounds.t -> ZBounds.t -> ZBounds.t)
(pf : forall x y z w l u,
opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- = Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ = Some {| Bounds.lower := l ; Bounds.upper := u |}
-> let val := opW (value x) (value y) (value z) (value w) in
is_bounded_by val l u)
: t -> t -> t -> t -> t
@@ -570,7 +582,7 @@ Module InterpretationsGen (Bit : BitSize).
(Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
as bop return opB _ _ _ _ = bop -> t
with
- | Some (ZBounds.Build_bounds l u)
+ | Some (Bounds.Build_bounds l u)
=> fun Heq => Some {| lower := l ; value := opW (value x) (value y) (value z) (value w) ; upper := u;
in_bounds := pf _ _ _ _ _ _ Heq |}
| None => fun _ => None
@@ -590,8 +602,8 @@ Module InterpretationsGen (Bit : BitSize).
=> progress (try unfold opB; try unfold opW)
| [ |- appcontext[ZBounds.t_map1 ?op] ] => let op' := head op in unfold op'
| [ |- appcontext[ZBounds.t_map2 ?op] ] => let op' := head op in unfold op'
- | [ |- appcontext[?op (ZBounds.Build_bounds _ _)] ] => let op' := head op in unfold op'
- | [ |- appcontext[?op (ZBounds.Build_bounds _ _) (ZBounds.Build_bounds _ _)] ] => unfold op
+ | [ |- appcontext[?op (Bounds.Build_bounds _ _)] ] => let op' := head op in unfold op'
+ | [ |- appcontext[?op (Bounds.Build_bounds _ _) (Bounds.Build_bounds _ _)] ] => unfold op
end
| progress cbv [BoundedWordToBounds ZBounds.SmartBuildBounds cmovne cmovl ModularBaseSystemListZOperations.neg] in *
| progress break_match
diff --git a/src/Specific/GF25519Bounded.v b/src/Specific/GF25519Bounded.v
index 778c3b3ed..3936d70c2 100644
--- a/src/Specific/GF25519Bounded.v
+++ b/src/Specific/GF25519Bounded.v
@@ -78,7 +78,7 @@ Definition postfreezeW : fe25519W -> fe25519W :=
(num_limbs := length_fe25519)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF25519.int_width)
+ (Interpretations64.WordW.neg GF25519.int_width)
).
Definition freezeW (f : fe25519W) : fe25519W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations64.WordW.wordWToZ with word64ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations64.WordW.neg].
+ change word64ToZ with Interpretations64.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word64ToZ (_ ^- (Interpretations64.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
diff --git a/src/Specific/GF25519Reflective.v b/src/Specific/GF25519Reflective.v
index 5e306d61d..a7ca684fb 100644
--- a/src/Specific/GF25519Reflective.v
+++ b/src/Specific/GF25519Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.Specific.GF25519BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.Specific.GF25519Reflective.Common.
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519W curry_unop_fe25519W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,10 +61,10 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519W curry_unop_fe25519W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
diff --git a/src/Specific/GF25519Reflective/Common.v b/src/Specific/GF25519Reflective/Common.v
index 8b500b42b..c6750fa3c 100644
--- a/src/Specific/GF25519Reflective/Common.v
+++ b/src/Specific/GF25519Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.Specific.GF25519.
Require Export Crypto.Specific.GF25519BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.fe25519W
- := fun e => curry_binop_fe25519W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe25519W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.fe25519W
- := fun e => curry_unop_fe25519W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.word64
- := fun e => curry_unop_fe25519W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe25519W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> Specific.GF25519BoundedCommon.wire_digitsW -> Specific.GF25519BoundedCommon.fe25519W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe25519W) (H : is_bounded (fe25519WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe25519W * fe25519W)
(H : is_bounded (fe25519WToZ (fst x)) = true)
(H' : is_bounded (fe25519WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe25519WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word64ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/Specific/GF25519Reflective/CommonBinOp.v b/src/Specific/GF25519Reflective/CommonBinOp.v
index 32dae16e6..b325a56eb 100644
--- a/src/Specific/GF25519Reflective/CommonBinOp.v
+++ b/src/Specific/GF25519Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.Specific.GF25519Reflective.Common.
Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/Specific/GF25519Reflective/CommonUnOp.v b/src/Specific/GF25519Reflective/CommonUnOp.v
index eac381b50..10390798b 100644
--- a/src/Specific/GF25519Reflective/CommonUnOp.v
+++ b/src/Specific/GF25519Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.Specific.GF25519Reflective.Common.
Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe25519WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe25519W x)
(Hx : is_bounded (fe25519WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v b/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v
index 2fa1d2ee3..2531e6184 100644
--- a/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v
+++ b/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.Specific.GF25519Reflective.Common.
Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe25519WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe25519W x)
(Hx : is_bounded (fe25519WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v b/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v
index 7528cfc2d..06d53a1e1 100644
--- a/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v
+++ b/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.Specific.GF25519Reflective.Common.
Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe25519WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe25519W x)
(Hx : is_bounded (fe25519WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v b/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v
index d61807413..d0c46d1ed 100644
--- a/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v
+++ b/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.Specific.GF25519Reflective.Common.
Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Bounded.v b/src/SpecificGen/GF2213_32Bounded.v
index 7aeb5fd45..2947aacfc 100644
--- a/src/SpecificGen/GF2213_32Bounded.v
+++ b/src/SpecificGen/GF2213_32Bounded.v
@@ -78,7 +78,7 @@ Definition postfreezeW : fe2213_32W -> fe2213_32W :=
(num_limbs := length_fe2213_32)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF2213_32.int_width)
+ (Interpretations64.WordW.neg GF2213_32.int_width)
).
Definition freezeW (f : fe2213_32W) : fe2213_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations64.WordW.wordWToZ with word64ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations64.WordW.neg].
+ change word64ToZ with Interpretations64.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word64ToZ (_ ^- (Interpretations64.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
diff --git a/src/SpecificGen/GF2213_32Reflective.v b/src/SpecificGen/GF2213_32Reflective.v
index 840cfe090..1fbf1a1c9 100644
--- a/src/SpecificGen/GF2213_32Reflective.v
+++ b/src/SpecificGen/GF2213_32Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Common.
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2213_32W curry_unop_fe2213_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,10 +61,10 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2213_32W curry_unop_fe2213_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
diff --git a/src/SpecificGen/GF2213_32Reflective/Common.v b/src/SpecificGen/GF2213_32Reflective/Common.v
index c01957df8..ff4999bcc 100644
--- a/src/SpecificGen/GF2213_32Reflective/Common.v
+++ b/src/SpecificGen/GF2213_32Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.SpecificGen.GF2213_32.
Require Export Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- := fun e => curry_binop_fe2213_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe2213_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- := fun e => curry_unop_fe2213_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe2213_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.word64
- := fun e => curry_unop_fe2213_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe2213_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe2213_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe2213_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF2213_32BoundedCommon.wire_digitsW -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe2213_32W) (H : is_bounded (fe2213_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe2213_32W * fe2213_32W)
(H : is_bounded (fe2213_32WToZ (fst x)) = true)
(H' : is_bounded (fe2213_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe2213_32WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word64ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v b/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
index be248d708..23becba34 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
index 3d52bcf78..250a64400 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe2213_32W x)
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
index a08efcfd8..2dcae7edd 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe2213_32W x)
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
index 1f3a5f8a0..ef997e3f9 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe2213_32W x)
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
index 727936a13..4342ef865 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Bounded.v b/src/SpecificGen/GF2519_32Bounded.v
index b49228967..ca448e458 100644
--- a/src/SpecificGen/GF2519_32Bounded.v
+++ b/src/SpecificGen/GF2519_32Bounded.v
@@ -78,7 +78,7 @@ Definition postfreezeW : fe2519_32W -> fe2519_32W :=
(num_limbs := length_fe2519_32)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF2519_32.int_width)
+ (Interpretations64.WordW.neg GF2519_32.int_width)
).
Definition freezeW (f : fe2519_32W) : fe2519_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations64.WordW.wordWToZ with word64ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations64.WordW.neg].
+ change word64ToZ with Interpretations64.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word64ToZ (_ ^- (Interpretations64.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
diff --git a/src/SpecificGen/GF2519_32Reflective.v b/src/SpecificGen/GF2519_32Reflective.v
index 964d0f608..ca7b86e5f 100644
--- a/src/SpecificGen/GF2519_32Reflective.v
+++ b/src/SpecificGen/GF2519_32Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Common.
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2519_32W curry_unop_fe2519_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,10 +61,10 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2519_32W curry_unop_fe2519_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
diff --git a/src/SpecificGen/GF2519_32Reflective/Common.v b/src/SpecificGen/GF2519_32Reflective/Common.v
index 16a7dd72c..aac5d8fb9 100644
--- a/src/SpecificGen/GF2519_32Reflective/Common.v
+++ b/src/SpecificGen/GF2519_32Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.SpecificGen.GF2519_32.
Require Export Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- := fun e => curry_binop_fe2519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe2519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- := fun e => curry_unop_fe2519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe2519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.word64
- := fun e => curry_unop_fe2519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe2519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe2519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe2519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF2519_32BoundedCommon.wire_digitsW -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe2519_32W) (H : is_bounded (fe2519_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe2519_32W * fe2519_32W)
(H : is_bounded (fe2519_32WToZ (fst x)) = true)
(H' : is_bounded (fe2519_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe2519_32WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word64ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v b/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
index d2bce7c3f..095c7a8c6 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
index d651e6118..f2e091570 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe2519_32W x)
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
index 10dc0866e..cf8a85d6f 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe2519_32W x)
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
index ec00c6233..7d70c1234 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe2519_32W x)
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
index 6ca61afe9..22e68fcf5 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Bounded.v b/src/SpecificGen/GF25519_32Bounded.v
index 264d7c3d3..7dddfe2c5 100644
--- a/src/SpecificGen/GF25519_32Bounded.v
+++ b/src/SpecificGen/GF25519_32Bounded.v
@@ -78,7 +78,7 @@ Definition postfreezeW : fe25519_32W -> fe25519_32W :=
(num_limbs := length_fe25519_32)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF25519_32.int_width)
+ (Interpretations64.WordW.neg GF25519_32.int_width)
).
Definition freezeW (f : fe25519_32W) : fe25519_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations64.WordW.wordWToZ with word64ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations64.WordW.neg].
+ change word64ToZ with Interpretations64.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word64ToZ (_ ^- (Interpretations64.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
diff --git a/src/SpecificGen/GF25519_32Reflective.v b/src/SpecificGen/GF25519_32Reflective.v
index 0459f10ab..e4e17b47f 100644
--- a/src/SpecificGen/GF25519_32Reflective.v
+++ b/src/SpecificGen/GF25519_32Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Common.
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_32W curry_unop_fe25519_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,10 +61,10 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_32W curry_unop_fe25519_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
diff --git a/src/SpecificGen/GF25519_32Reflective/Common.v b/src/SpecificGen/GF25519_32Reflective/Common.v
index 4105b4173..96958f330 100644
--- a/src/SpecificGen/GF25519_32Reflective/Common.v
+++ b/src/SpecificGen/GF25519_32Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.SpecificGen.GF25519_32.
Require Export Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- := fun e => curry_binop_fe25519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe25519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- := fun e => curry_unop_fe25519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.word64
- := fun e => curry_unop_fe25519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe25519_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF25519_32BoundedCommon.wire_digitsW -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe25519_32W) (H : is_bounded (fe25519_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe25519_32W * fe25519_32W)
(H : is_bounded (fe25519_32WToZ (fst x)) = true)
(H' : is_bounded (fe25519_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe25519_32WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word64ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v b/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
index 837229900..f254a2d3a 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
index 864efb92b..246dcbf70 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe25519_32W x)
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
index 68efe5b86..a940a0c8b 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe25519_32W x)
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
index 6b9864d6f..b7ec2c6a5 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe25519_32W x)
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
index 09ac72cdc..34e0896f1 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Bounded.v b/src/SpecificGen/GF25519_64Bounded.v
index 35ad2ee01..71b859d32 100644
--- a/src/SpecificGen/GF25519_64Bounded.v
+++ b/src/SpecificGen/GF25519_64Bounded.v
@@ -46,7 +46,7 @@ Local Ltac define_unop_WireToFE f opW blem :=
abstract bounded_wire_digits_t opW blem.
Local Opaque Let_In.
-Local Opaque Z.add Z.sub Z.mul Z.shiftl Z.shiftr Z.land Z.lor Z.eqb NToWord64.
+Local Opaque Z.add Z.sub Z.mul Z.shiftl Z.shiftr Z.land Z.lor Z.eqb NToWord128.
Local Arguments interp_radd / _ _.
Local Arguments interp_rsub / _ _.
Local Arguments interp_rmul / _ _.
@@ -60,12 +60,12 @@ Definition subW (f g : fe25519_64W) : fe25519_64W := Eval simpl in interp_rsub f
Definition mulW (f g : fe25519_64W) : fe25519_64W := Eval simpl in interp_rmul f g.
Definition oppW (f : fe25519_64W) : fe25519_64W := Eval simpl in interp_ropp f.
Definition prefreezeW (f : fe25519_64W) : fe25519_64W := Eval simpl in interp_rprefreeze f.
-Definition ge_modulusW (f : fe25519_64W) : word64 := Eval simpl in interp_rge_modulus f.
+Definition ge_modulusW (f : fe25519_64W) : word128 := Eval simpl in interp_rge_modulus f.
Definition packW (f : fe25519_64W) : wire_digitsW := Eval simpl in interp_rpack f.
Definition unpackW (f : wire_digitsW) : fe25519_64W := Eval simpl in interp_runpack f.
Definition modulusW :=
- Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z length_fe25519_64 GF25519_64.modulus_digits_)).
+ Eval cbv - [ZToWord128] in (Tuple.map ZToWord128 (Tuple.from_list_default 0%Z length_fe25519_64 GF25519_64.modulus_digits_)).
Definition postfreeze : GF25519_64.fe25519_64 -> GF25519_64.fe25519_64 :=
GF25519_64.postfreeze.
@@ -78,7 +78,7 @@ Definition postfreezeW : fe25519_64W -> fe25519_64W :=
(num_limbs := length_fe25519_64)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF25519_64.int_width)
+ (Interpretations128.WordW.neg GF25519_64.int_width)
).
Definition freezeW (f : fe25519_64W) : fe25519_64W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations128.WordW.wordWToZ with word128ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations128.WordW.neg].
+ change word128ToZ with Interpretations128.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word128ToZ (_ ^- (Interpretations128.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord128 with Interpretations128.WordW.ZToWordW in *.
+ rewrite !Interpretations128.WordW.wordWToZ_sub;
+ rewrite !Interpretations128.WordW.wordWToZ_land;
+ rewrite !Interpretations128.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord128 with Interpretations128.WordW.ZToWordW in *.
+ rewrite !Interpretations128.WordW.wordWToZ_sub;
+ rewrite !Interpretations128.WordW.wordWToZ_land;
+ rewrite !Interpretations128.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
@@ -235,7 +235,7 @@ Proof.
hnf in f, g; destruct_head' prod.
eexists.
cbv [GF25519_64.fieldwiseb fe25519_64WToZ].
- rewrite ?word64eqb_Zeqb.
+ rewrite ?word128eqb_Zeqb.
reflexivity.
Defined.
@@ -288,7 +288,7 @@ Qed.
Definition sqrt_m1W' : fe25519_64W :=
Eval vm_compute in fe25519_64ZToW sqrt_m1.
-Definition sqrt_m1W := Eval cbv [sqrt_m1W' fe25519_64W_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_64W_word64ize sqrt_m1W'.
+Definition sqrt_m1W := Eval cbv [sqrt_m1W' fe25519_64W_word128ize word128ize andb opt.word128ToZ opt.word128ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_64W_word128ize sqrt_m1W'.
Definition GF25519_64sqrt x : GF25519_64.fe25519_64 :=
dlet powx := powW (fe25519_64ZToW x) (chain GF25519_64.sqrt_ec) in
@@ -366,7 +366,7 @@ Definition opp (f : fe25519_64) : fe25519_64.
Proof. define_unop f oppW oppW_correct_and_bounded. Defined.
Definition freeze (f : fe25519_64) : fe25519_64.
Proof. define_unop f freezeW freezeW_correct_and_bounded. Defined.
-Definition ge_modulus (f : fe25519_64) : word64.
+Definition ge_modulus (f : fe25519_64) : word128.
Proof. define_unop_FEToZ f ge_modulusW. Defined.
Definition pack (f : fe25519_64) : wire_digits.
Proof. define_unop_FEToWire f packW packW_correct_and_bounded. Defined.
@@ -407,7 +407,7 @@ Lemma opp_correct (f : fe25519_64) : proj1_fe25519_64 (opp f) = carry_opp (proj1
Proof. op_correct_t opp oppW_correct_and_bounded. Qed.
Lemma freeze_correct (f : fe25519_64) : proj1_fe25519_64 (freeze f) = GF25519_64.freeze (proj1_fe25519_64 f).
Proof. op_correct_t freeze freezeW_correct_and_bounded. Qed.
-Lemma ge_modulus_correct (f : fe25519_64) : word64ToZ (ge_modulus f) = GF25519_64.ge_modulus (proj1_fe25519_64 f).
+Lemma ge_modulus_correct (f : fe25519_64) : word128ToZ (ge_modulus f) = GF25519_64.ge_modulus (proj1_fe25519_64 f).
Proof. op_correct_t ge_modulus ge_modulusW_correct. Qed.
Lemma pack_correct (f : fe25519_64) : proj1_wire_digits (pack f) = GF25519_64.pack (proj1_fe25519_64 f).
Proof. op_correct_t pack packW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF25519_64BoundedCommon.v b/src/SpecificGen/GF25519_64BoundedCommon.v
index e6a3b6c20..b14612f07 100644
--- a/src/SpecificGen/GF25519_64BoundedCommon.v
+++ b/src/SpecificGen/GF25519_64BoundedCommon.v
@@ -38,50 +38,50 @@ Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
(* END common curve-specific definitions *)
(* BEGIN aliases for word extraction *)
-Definition word64 := Word.word bit_width.
-Coercion word64ToZ (x : word64) : Z := Z.of_N (wordToN x).
-Coercion ZToWord64 (x : Z) : word64 := NToWord _ (Z.to_N x).
-Definition NToWord64 : N -> word64 := NToWord _.
-Definition word64ize (x : word64) : word64
- := Eval cbv [wordToN N.succ_double N.double] in NToWord64 (wordToN x).
-Definition w64eqb (x y : word64) := weqb x y.
-
-Global Arguments NToWord64 : simpl never.
-Arguments word64 : simpl never.
+Definition word128 := Word.word bit_width.
+Coercion word128ToZ (x : word128) : Z := Z.of_N (wordToN x).
+Coercion ZToWord128 (x : Z) : word128 := NToWord _ (Z.to_N x).
+Definition NToWord128 : N -> word128 := NToWord _.
+Definition word128ize (x : word128) : word128
+ := Eval cbv [wordToN N.succ_double N.double] in NToWord128 (wordToN x).
+Definition w128eqb (x y : word128) := weqb x y.
+
+Global Arguments NToWord128 : simpl never.
+Arguments word128 : simpl never.
Arguments bit_width : simpl never.
-Global Opaque word64.
+Global Opaque word128.
Global Opaque bit_width.
(* END aliases for word extraction *)
(* BEGIN basic types *)
Module Type WordIsBounded.
- Parameter is_boundedT : forall (lower upper : Z), word64 -> bool.
- Parameter Build_is_boundedT : forall {lower upper} {proj_word : word64},
+ Parameter is_boundedT : forall (lower upper : Z), word128 -> bool.
+ Parameter Build_is_boundedT : forall {lower upper} {proj_word : word128},
andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true -> is_boundedT lower upper proj_word = true.
- Parameter project_is_boundedT : forall {lower upper} {proj_word : word64},
+ Parameter project_is_boundedT : forall {lower upper} {proj_word : word128},
is_boundedT lower upper proj_word = true -> andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true.
End WordIsBounded.
Module Import WordIsBoundedDefault : WordIsBounded.
- Definition is_boundedT : forall (lower upper : Z), word64 -> bool
+ Definition is_boundedT : forall (lower upper : Z), word128 -> bool
:= fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
- Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ Definition Build_is_boundedT {lower upper} {proj_word : word128}
: andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true -> is_boundedT lower upper proj_word = true
:= fun x => x.
- Definition project_is_boundedT {lower upper} {proj_word : word64}
+ Definition project_is_boundedT {lower upper} {proj_word : word128}
: is_boundedT lower upper proj_word = true -> andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true
:= fun x => x.
End WordIsBoundedDefault.
Definition bounded_word (lower upper : Z)
- := { proj_word : word64 | is_boundedT lower upper proj_word = true }.
+ := { proj_word : word128 | is_boundedT lower upper proj_word = true }.
Local Notation word_of exp := (bounded_word (fst (b_of exp)) (snd (b_of exp))).
Local Notation unbounded_word sz := (bounded_word 0 (2^sz-1)%Z).
-Local Opaque word64.
-Definition fe25519_64W := Eval cbv (*-[word64]*) in (tuple word64 length_fe25519_64).
-Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
+Local Opaque word128.
+Definition fe25519_64W := Eval cbv (*-[word128]*) in (tuple word128 length_fe25519_64).
+Definition wire_digitsW := Eval cbv (*-[word128]*) in (tuple word128 (length wire_widths)).
Definition fe25519_64 :=
Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
hlist (fun e => word_of e) bounds_exp.
@@ -203,25 +203,25 @@ Defined.
Definition eta_wire_digitsW (x : wire_digitsW) : wire_digitsW
:= Eval cbv [proj1_sig eta_wire_digitsW_sig] in proj1_sig (eta_wire_digitsW_sig x).
-Local Transparent word64.
-Lemma word64ize_id x : word64ize x = x.
+Local Transparent word128.
+Lemma word128ize_id x : word128ize x = x.
Proof. apply NToWord_wordToN. Qed.
-Local Opaque word64.
+Local Opaque word128.
-Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb x y.
+Lemma word128eqb_Zeqb x y : (word128ToZ x =? word128ToZ y)%Z = w128eqb x y.
Proof. apply wordeqb_Zeqb. Qed.
Local Arguments Z.pow_pos !_ !_ / .
-Lemma word64ToZ_ZToWord64 x : 0 <= x < 2^Z.of_nat bit_width -> word64ToZ (ZToWord64 x) = x.
+Lemma word128ToZ_ZToWord128 x : 0 <= x < 2^Z.of_nat bit_width -> word128ToZ (ZToWord128 x) = x.
Proof.
- intros; unfold word64ToZ, ZToWord64.
+ intros; unfold word128ToZ, ZToWord128.
rewrite wordToN_NToWord_idempotent, Z2N.id
by (omega || apply N2Z.inj_lt; rewrite <- ?(N_nat_Z (Npow2 _)), ?Npow2_nat, ?Zpow_pow2, ?N2Z.id, ?Z2N.id, ?Z2Nat.id by omega; omega).
reflexivity.
Qed.
-Lemma ZToWord64_word64ToZ x : ZToWord64 (word64ToZ x) = x.
+Lemma ZToWord128_word128ToZ x : ZToWord128 (word128ToZ x) = x.
Proof.
- intros; unfold word64ToZ, ZToWord64.
+ intros; unfold word128ToZ, ZToWord128.
rewrite N2Z.id, NToWord_wordToN; reflexivity.
Qed.
@@ -236,7 +236,7 @@ Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_wo
Arguments proj_word {_ _} _.
Arguments word_bounded {_ _} _.
Arguments Build_bounded_word' {_ _} _ _.
-Definition Build_bounded_word {lower upper} (proj_word : word64) (word_bounded : andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true)
+Definition Build_bounded_word {lower upper} (proj_word : word128) (word_bounded : andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true)
: bounded_word lower upper
:= Build_bounded_word'
proj_word
@@ -244,13 +244,13 @@ Definition Build_bounded_word {lower upper} (proj_word : word64) (word_bounded :
| true => fun _ => eq_refl
| false => fun x => x
end word_bounded).
-Lemma word_to_unbounded_helper {x e : nat} : (x < pow2 e)%nat -> (Z.of_nat e <= Z.of_nat bit_width)%Z -> ((0 <=? word64ToZ (ZToWord64 (Z.of_nat x))) && (word64ToZ (ZToWord64 (Z.of_nat x)) <=? 2 ^ (Z.of_nat e) - 1))%bool = true.
+Lemma word_to_unbounded_helper {x e : nat} : (x < pow2 e)%nat -> (Z.of_nat e <= Z.of_nat bit_width)%Z -> ((0 <=? word128ToZ (ZToWord128 (Z.of_nat x))) && (word128ToZ (ZToWord128 (Z.of_nat x)) <=? 2 ^ (Z.of_nat e) - 1))%bool = true.
Proof.
rewrite pow2_id; intro H; apply Nat2Z.inj_lt in H; revert H.
rewrite Z.pow_Zpow; simpl Z.of_nat.
intros H H'.
assert (2^Z.of_nat e <= 2^Z.of_nat bit_width) by auto with zarith.
- rewrite ?word64ToZ_ZToWord64 by omega.
+ rewrite ?word128ToZ_ZToWord128 by omega.
match goal with
| [ |- context[andb ?x ?y] ]
=> destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
@@ -267,26 +267,26 @@ Proof. apply (word_to_unbounded_word x); reflexivity. Defined.
Definition word31_to_unbounded_word (x : word 31) : unbounded_word 31.
Proof. apply (word_to_unbounded_word x); reflexivity. Defined.
-Local Opaque word64.
+Local Opaque word128.
Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe25519_64W app_fe25519_64 HList.mapt HList.mapt' Tuple.map on_tuple List.map List.app length_fe25519_64 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
Definition fe25519_64WToZ (x : fe25519_64W) : SpecificGen.GF25519_64.fe25519_64
:= Eval app_tuple_map in
- app_fe25519_64W x (Tuple.map (fun v : word64 => v : Z)).
+ app_fe25519_64W x (Tuple.map (fun v : word128 => v : Z)).
Definition fe25519_64ZToW (x : SpecificGen.GF25519_64.fe25519_64) : fe25519_64W
:= Eval app_tuple_map in
- app_fe25519_64W x (Tuple.map (fun v : Z => v : word64)).
+ app_fe25519_64W x (Tuple.map (fun v : Z => v : word128)).
Definition wire_digitsWToZ (x : wire_digitsW) : SpecificGen.GF25519_64.wire_digits
:= Eval app_tuple_map in
- app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
+ app_wire_digitsW x (Tuple.map (fun v : word128 => v : Z)).
Definition wire_digitsZToW (x : SpecificGen.GF25519_64.wire_digits) : wire_digitsW
:= Eval app_tuple_map in
- app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
-Definition fe25519_64W_word64ize (x : fe25519_64W) : fe25519_64W
+ app_wire_digitsW x (Tuple.map (fun v : Z => v : word128)).
+Definition fe25519_64W_word128ize (x : fe25519_64W) : fe25519_64W
:= Eval app_tuple_map in
- app_fe25519_64W x (Tuple.map word64ize).
-Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ app_fe25519_64W x (Tuple.map word128ize).
+Definition wire_digitsW_word128ize (x : wire_digitsW) : wire_digitsW
:= Eval app_tuple_map in
- app_wire_digitsW x (Tuple.map word64ize).
+ app_wire_digitsW x (Tuple.map word128ize).
(** TODO: Turn this into a lemma to speed up proofs *)
Ltac unfold_is_bounded_in H :=
@@ -300,17 +300,17 @@ Ltac unfold_is_bounded :=
rewrite ?Bool.andb_true_iff.
Local Transparent bit_width.
-Definition Pow2_64 := Eval compute in 2^Z.of_nat bit_width.
-Definition unfold_Pow2_64 : 2^Z.of_nat bit_width = Pow2_64 := eq_refl.
+Definition Pow2_128 := Eval compute in 2^Z.of_nat bit_width.
+Definition unfold_Pow2_128 : 2^Z.of_nat bit_width = Pow2_128 := eq_refl.
Local Opaque bit_width.
Local Ltac prove_lt_bit_width :=
- rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
+ rewrite unfold_Pow2_128; cbv [Pow2_128]; omega.
Lemma fe25519_64ZToW_WToZ (x : fe25519_64W) : fe25519_64ZToW (fe25519_64WToZ x) = x.
Proof.
hnf in x; destruct_head' prod; cbv [fe25519_64WToZ fe25519_64ZToW].
- rewrite !ZToWord64_word64ToZ; reflexivity.
+ rewrite !ZToWord128_word128ToZ; reflexivity.
Qed.
Lemma fe25519_64WToZ_ZToW x : is_bounded x = true -> fe25519_64WToZ (fe25519_64ZToW x) = x.
@@ -319,36 +319,36 @@ Proof.
intro H.
unfold_is_bounded_in H; destruct_head' and.
Z.ltb_to_lt.
- rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ rewrite !word128ToZ_ZToWord128 by prove_lt_bit_width.
reflexivity.
Qed.
-Lemma fe25519_64W_word64ize_id x : fe25519_64W_word64ize x = x.
+Lemma fe25519_64W_word128ize_id x : fe25519_64W_word128ize x = x.
Proof.
hnf in x; destruct_head' prod.
- cbv [fe25519_64W_word64ize];
- repeat apply f_equal2; apply word64ize_id.
+ cbv [fe25519_64W_word128ize];
+ repeat apply f_equal2; apply word128ize_id.
Qed.
-Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Lemma wire_digitsW_word128ize_id x : wire_digitsW_word128ize x = x.
Proof.
hnf in x; destruct_head' prod.
- cbv [wire_digitsW_word64ize];
- repeat apply f_equal2; apply word64ize_id.
+ cbv [wire_digitsW_word128ize];
+ repeat apply f_equal2; apply word128ize_id.
Qed.
Definition uncurry_unop_fe25519_64W {T} (op : fe25519_64W -> T)
- := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe25519_64) op.
+ := Eval cbv (*-[word128]*) in Tuple.uncurry (n:=length_fe25519_64) op.
Definition curry_unop_fe25519_64W {T} op : fe25519_64W -> T
- := Eval cbv (*-[word64]*) in fun f => app_fe25519_64W f (Tuple.curry (n:=length_fe25519_64) op).
+ := Eval cbv (*-[word128]*) in fun f => app_fe25519_64W f (Tuple.curry (n:=length_fe25519_64) op).
Definition uncurry_binop_fe25519_64W {T} (op : fe25519_64W -> fe25519_64W -> T)
- := Eval cbv (*-[word64]*) in uncurry_unop_fe25519_64W (fun f => uncurry_unop_fe25519_64W (op f)).
+ := Eval cbv (*-[word128]*) in uncurry_unop_fe25519_64W (fun f => uncurry_unop_fe25519_64W (op f)).
Definition curry_binop_fe25519_64W {T} op : fe25519_64W -> fe25519_64W -> T
- := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe25519_64W (curry_unop_fe25519_64W op f)).
+ := Eval cbv (*-[word128]*) in appify2 (fun f => curry_unop_fe25519_64W (curry_unop_fe25519_64W op f)).
Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
- := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length wire_widths) op.
+ := Eval cbv (*-[word128]*) in Tuple.uncurry (n:=length wire_widths) op.
Definition curry_unop_wire_digitsW {T} op : wire_digitsW -> T
- := Eval cbv (*-[word64]*) in fun f => app_wire_digitsW f (Tuple.curry (n:=length wire_widths) op).
+ := Eval cbv (*-[word128]*) in fun f => app_wire_digitsW f (Tuple.curry (n:=length wire_widths) op).
Definition proj1_fe25519_64W (x : fe25519_64) : fe25519_64W
@@ -417,7 +417,7 @@ Local Ltac make_exist'' x exist_W ZToW :=
unfold_is_bounded_in H;
destruct_head' and; simpl in *;
Z.ltb_to_lt;
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ rewrite ?word128ToZ_ZToWord128 by prove_lt_bit_width;
assumption
).
Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
@@ -459,7 +459,7 @@ Proof.
unfold_is_bounded_in pf.
destruct_head' and.
Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ rewrite ?word128ToZ_ZToWord128 by prove_lt_bit_width.
reflexivity.
Qed.
@@ -491,18 +491,18 @@ Proof.
unfold_is_bounded_in pf.
destruct_head' and.
Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ rewrite ?word128ToZ_ZToWord128 by prove_lt_bit_width.
reflexivity.
Qed.
Module opt.
- Definition word64ToZ := Eval vm_compute in word64ToZ.
- Definition word64ToN := Eval vm_compute in @wordToN bit_width.
- Definition NToWord64 := Eval vm_compute in NToWord64.
+ Definition word128ToZ := Eval vm_compute in word128ToZ.
+ Definition word128ToN := Eval vm_compute in @wordToN bit_width.
+ Definition NToWord128 := Eval vm_compute in NToWord128.
Definition bit_width := Eval vm_compute in bit_width.
Definition Zleb := Eval cbv [Z.leb] in Z.leb.
Definition andb := Eval vm_compute in andb.
- Definition word64ize := Eval vm_compute in word64ize.
+ Definition word128ize := Eval vm_compute in word128ize.
End opt.
Local Transparent bit_width.
@@ -512,35 +512,35 @@ Local Ltac do_change lem :=
=> let x' := (eval vm_compute in x) in
let z' := (eval vm_compute in z) in
lazymatch y with
- | word64ToZ (word64ize ?v)
- => let y' := constr:(opt.word64ToZ (opt.word64ize v)) in
+ | word128ToZ (word128ize ?v)
+ => let y' := constr:(opt.word128ToZ (opt.word128ize v)) in
let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
do_change L'
end
| _ => lem
end.
-Definition fe25519_64_word64ize (x : fe25519_64) : fe25519_64.
+Definition fe25519_64_word128ize (x : fe25519_64) : fe25519_64.
Proof.
set (x' := x).
hnf in x; destruct_head' prod.
- let lem := constr:(exist_fe25519_64W (fe25519_64W_word64ize (proj1_fe25519_64W x'))) in
- let lem := (eval cbv [proj1_fe25519_64W x' fe25519_64W_word64ize proj_word exist_fe25519_64W Build_bounded_word' Build_bounded_word] in lem) in
+ let lem := constr:(exist_fe25519_64W (fe25519_64W_word128ize (proj1_fe25519_64W x'))) in
+ let lem := (eval cbv [proj1_fe25519_64W x' fe25519_64W_word128ize proj_word exist_fe25519_64W Build_bounded_word' Build_bounded_word] in lem) in
let lem := do_change lem in
refine (lem _);
- change (is_bounded (fe25519_64WToZ (fe25519_64W_word64ize (proj1_fe25519_64W x'))) = true);
- abstract (rewrite fe25519_64W_word64ize_id; apply is_bounded_proj1_fe25519_64).
+ change (is_bounded (fe25519_64WToZ (fe25519_64W_word128ize (proj1_fe25519_64W x'))) = true);
+ abstract (rewrite fe25519_64W_word128ize_id; apply is_bounded_proj1_fe25519_64).
Defined.
-Definition wire_digits_word64ize (x : wire_digits) : wire_digits.
+Definition wire_digits_word128ize (x : wire_digits) : wire_digits.
Proof.
set (x' := x).
hnf in x; destruct_head' prod.
- let lem := constr:(exist_wire_digitsW (wire_digitsW_word64ize (proj1_wire_digitsW x'))) in
- let lem := (eval cbv [proj1_wire_digitsW x' wire_digitsW_word64ize proj_word exist_wire_digitsW Build_bounded_word Build_bounded_word'] in lem) in
+ let lem := constr:(exist_wire_digitsW (wire_digitsW_word128ize (proj1_wire_digitsW x'))) in
+ let lem := (eval cbv [proj1_wire_digitsW x' wire_digitsW_word128ize proj_word exist_wire_digitsW Build_bounded_word Build_bounded_word'] in lem) in
let lem := do_change lem in
- let lem := (eval cbv [word64ize opt.word64ize andb Z.leb Z.compare CompOpp Pos.compare] in lem) in
+ let lem := (eval cbv [word128ize opt.word128ize andb Z.leb Z.compare CompOpp Pos.compare] in lem) in
refine (lem _);
- change (wire_digits_is_bounded (wire_digitsWToZ (wire_digitsW_word64ize (proj1_wire_digitsW x'))) = true);
- abstract (rewrite wire_digitsW_word64ize_id; apply is_bounded_proj1_wire_digits).
+ change (wire_digits_is_bounded (wire_digitsWToZ (wire_digitsW_word128ize (proj1_wire_digitsW x'))) = true);
+ abstract (rewrite wire_digitsW_word128ize_id; apply is_bounded_proj1_wire_digits).
Defined.
Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
@@ -566,10 +566,10 @@ Qed.
(* Precompute constants *)
Definition one' := Eval vm_compute in exist_fe25519_64 SpecificGen.GF25519_64.one_ eq_refl.
-Definition one := Eval cbv [one' fe25519_64_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_64_word64ize one'.
+Definition one := Eval cbv [one' fe25519_64_word128ize word128ize andb opt.word128ToZ opt.word128ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_64_word128ize one'.
Definition zero' := Eval vm_compute in exist_fe25519_64 SpecificGen.GF25519_64.zero_ eq_refl.
-Definition zero := Eval cbv [zero' fe25519_64_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_64_word64ize zero'.
+Definition zero := Eval cbv [zero' fe25519_64_word128ize word128ize andb opt.word128ToZ opt.word128ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_64_word128ize zero'.
Lemma fold_chain_opt_gen {A B} (F : A -> B) is_bounded ls id' op' id op chain
(Hid_bounded : is_bounded (F id') = true)
@@ -642,7 +642,7 @@ Proof.
unfold_is_bounded_in pf.
destruct_head' and.
Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ rewrite ?word128ToZ_ZToWord128 by prove_lt_bit_width.
reflexivity.
Qed.
@@ -654,7 +654,7 @@ Proof.
unfold_is_bounded_in pf.
destruct_head' and.
Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ rewrite ?word128ToZ_ZToWord128 by prove_lt_bit_width.
reflexivity.
Qed.
@@ -678,7 +678,7 @@ Notation iunop_correct_and_bounded irop op
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe25519_64WToZ x) = true
- -> word64ToZ (irop x) = op (fe25519_64WToZ x)) (only parsing).
+ -> word128ToZ (irop x) = op (fe25519_64WToZ x)) (only parsing).
Notation iunop_FEToWire_correct_and_bounded irop op
:= (forall x,
is_bounded (fe25519_64WToZ x) = true
diff --git a/src/SpecificGen/GF25519_64Reflective.v b/src/SpecificGen/GF25519_64Reflective.v
index 095def252..69d3ea34a 100644
--- a/src/SpecificGen/GF25519_64Reflective.v
+++ b/src/SpecificGen/GF25519_64Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations128.
+Require Crypto.Reflection.Z.Interpretations128.Relations.
+Require Import Crypto.Reflection.Z.Interpretations128.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Common.
@@ -41,8 +41,8 @@ Definition rge_modulus : ExprUnOpFEToZ := Eval vm_compute in rge_modulusW.
Definition rpack : ExprUnOpFEToWire := Eval vm_compute in rpackW.
Definition runpack : ExprUnOpWireToFE := Eval vm_compute in runpackW.
-Definition rword64ize {t} (x : Expr t) : Expr t
- := MapInterp (fun t => match t with TZ => word64ize end) x.
+Definition rword128ize {t} (x : Expr t) : Expr t
+ := MapInterp (fun t => match t with TZ => word128ize end) x.
Declare Reduction asm_interp
:= cbv beta iota delta
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_64W curry_unop_fe25519_64W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word128ize rword128ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,44 +61,44 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_64W curry_unop_fe25519_64W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word128ize rword128ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Definition interp_radd : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := Eval asm_interp in interp_bexpr (rword64ize radd).
+ := Eval asm_interp in interp_bexpr (rword128ize radd).
(*Print interp_radd.*)
Definition interp_radd_correct : interp_radd = interp_bexpr radd := eq_refl.
Definition interp_rsub : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := Eval asm_interp in interp_bexpr (rword64ize rsub).
+ := Eval asm_interp in interp_bexpr (rword128ize rsub).
(*Print interp_rsub.*)
Definition interp_rsub_correct : interp_rsub = interp_bexpr rsub := eq_refl.
Definition interp_rmul : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := Eval asm_interp in interp_bexpr (rword64ize rmul).
+ := Eval asm_interp in interp_bexpr (rword128ize rmul).
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
Definition interp_ropp : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := Eval asm_interp in interp_uexpr (rword64ize ropp).
+ := Eval asm_interp in interp_uexpr (rword128ize ropp).
(*Print interp_ropp.*)
Definition interp_ropp_correct : interp_ropp = interp_uexpr ropp := eq_refl.
Definition interp_rprefreeze : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := Eval asm_interp in interp_uexpr (rword64ize rprefreeze).
+ := Eval asm_interp in interp_uexpr (rword128ize rprefreeze).
(*Print interp_rprefreeze.*)
Definition interp_rprefreeze_correct : interp_rprefreeze = interp_uexpr rprefreeze := eq_refl.
-Definition interp_rge_modulus : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.word64
- := Eval asm_interp in interp_uexpr_FEToZ (rword64ize rge_modulus).
+Definition interp_rge_modulus : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.word128
+ := Eval asm_interp in interp_uexpr_FEToZ (rword128ize rge_modulus).
Definition interp_rge_modulus_correct : interp_rge_modulus = interp_uexpr_FEToZ rge_modulus := eq_refl.
Definition interp_rpack : SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.wire_digitsW
- := Eval asm_interp in interp_uexpr_FEToWire (rword64ize rpack).
+ := Eval asm_interp in interp_uexpr_FEToWire (rword128ize rpack).
Definition interp_rpack_correct : interp_rpack = interp_uexpr_FEToWire rpack := eq_refl.
Definition interp_runpack : SpecificGen.GF25519_64BoundedCommon.wire_digitsW -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := Eval asm_interp in interp_uexpr_WireToFE (rword64ize runpack).
+ := Eval asm_interp in interp_uexpr_WireToFE (rword128ize runpack).
Definition interp_runpack_correct : interp_runpack = interp_uexpr_WireToFE runpack := eq_refl.
Lemma radd_correct_and_bounded : binop_correct_and_bounded radd carry_add.
diff --git a/src/SpecificGen/GF25519_64Reflective/Common.v b/src/SpecificGen/GF25519_64Reflective/Common.v
index 444edcde4..0811a71cc 100644
--- a/src/SpecificGen/GF25519_64Reflective/Common.v
+++ b/src/SpecificGen/GF25519_64Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.SpecificGen.GF25519_64.
Require Export Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations128.
+Require Crypto.Reflection.Z.Interpretations128.Relations.
+Require Import Crypto.Reflection.Z.Interpretations128.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := fun e => curry_binop_fe25519_64W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe25519_64W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := fun e => curry_unop_fe25519_64W (Interp (@Word64.interp_op) e).
-Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.word64
- := fun e => curry_unop_fe25519_64W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519_64W (Interp (@WordW.interp_op) e).
+Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.word128
+ := fun e => curry_unop_fe25519_64W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe25519_64W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe25519_64W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF25519_64BoundedCommon.wire_digitsW -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe25519_64W) (H : is_bounded (fe25519_64WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe25519_64W * fe25519_64W)
(H : is_bounded (fe25519_64WToZ (fst x)) = true)
(H' : is_bounded (fe25519_64WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe25519_64WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word128ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word128ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v b/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
index cd4a5031b..c8872efc5 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
index bfc3c3cc1..f6a71740a 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe25519_64W x)
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
index f615067c6..7bd176749 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe25519_64W x)
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
index 079a0729c..d6b8bb2c7 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe25519_64W x)
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
index c58f70e91..baadad3c3 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Bounded.v b/src/SpecificGen/GF41417_32Bounded.v
index 53ee754f2..be7cf5714 100644
--- a/src/SpecificGen/GF41417_32Bounded.v
+++ b/src/SpecificGen/GF41417_32Bounded.v
@@ -78,7 +78,7 @@ Definition postfreezeW : fe41417_32W -> fe41417_32W :=
(num_limbs := length_fe41417_32)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF41417_32.int_width)
+ (Interpretations64.WordW.neg GF41417_32.int_width)
).
Definition freezeW (f : fe41417_32W) : fe41417_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations64.WordW.wordWToZ with word64ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations64.WordW.neg].
+ change word64ToZ with Interpretations64.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word64ToZ (_ ^- (Interpretations64.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
diff --git a/src/SpecificGen/GF41417_32Reflective.v b/src/SpecificGen/GF41417_32Reflective.v
index d85e22870..d89b5d87b 100644
--- a/src/SpecificGen/GF41417_32Reflective.v
+++ b/src/SpecificGen/GF41417_32Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Common.
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe41417_32W curry_unop_fe41417_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,10 +61,10 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe41417_32W curry_unop_fe41417_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
diff --git a/src/SpecificGen/GF41417_32Reflective/Common.v b/src/SpecificGen/GF41417_32Reflective/Common.v
index 517c66d8a..68de90117 100644
--- a/src/SpecificGen/GF41417_32Reflective/Common.v
+++ b/src/SpecificGen/GF41417_32Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.SpecificGen.GF41417_32.
Require Export Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- := fun e => curry_binop_fe41417_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe41417_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- := fun e => curry_unop_fe41417_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe41417_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.word64
- := fun e => curry_unop_fe41417_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe41417_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe41417_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe41417_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF41417_32BoundedCommon.wire_digitsW -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe41417_32W) (H : is_bounded (fe41417_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe41417_32W * fe41417_32W)
(H : is_bounded (fe41417_32WToZ (fst x)) = true)
(H' : is_bounded (fe41417_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe41417_32WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word64ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v b/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
index 93a83cf51..3542dc9cf 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
index c616b1a77..7d86a5e95 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe41417_32W x)
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
index 7391e98bf..a008c50bb 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe41417_32W x)
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
index 6d991dbb4..700581bad 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe41417_32W x)
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
index bb91ea89f..d04d44d32 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Bounded.v b/src/SpecificGen/GF5211_32Bounded.v
index 74480a0d1..c879a3b19 100644
--- a/src/SpecificGen/GF5211_32Bounded.v
+++ b/src/SpecificGen/GF5211_32Bounded.v
@@ -78,7 +78,7 @@ Definition postfreezeW : fe5211_32W -> fe5211_32W :=
(num_limbs := length_fe5211_32)
modulusW
ge_modulusW
- (Interpretations.Word64.neg GF5211_32.int_width)
+ (Interpretations64.WordW.neg GF5211_32.int_width)
).
Definition freezeW (f : fe5211_32W) : fe5211_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
@@ -117,7 +117,7 @@ Ltac lower_bound_minus_ge_modulus :=
cbv [ge_modulus Let_In ModularBaseSystemListZOperations.cmovl ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.neg];
repeat break_if; Z.ltb_to_lt; subst; try omega;
rewrite ?Z.land_0_l; auto;
- change Interpretations.Word64.word64ToZ with word64ToZ;
+ change Interpretations64.WordW.wordWToZ with word64ToZ;
etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
@@ -136,8 +136,8 @@ Proof.
destruct_head' and.
Z.ltb_to_lt.
cbv [postfreezeW].
- cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
- change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ cbv [conditional_subtract_modulusW Interpretations64.WordW.neg].
+ change word64ToZ with Interpretations64.WordW.wordWToZ in *.
rewrite Hgm.
cbv [modulusW Tuple.map].
@@ -148,12 +148,12 @@ Proof.
split.
{ match goal with
- |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
+ |- (_,word64ToZ (_ ^- (Interpretations64.WordW.ZToWordW ?x) ^& _)) = (_,_ - (?y &' _)) => assert (x = y) as Hxy by reflexivity; repeat rewrite <-Hxy; clear Hxy end.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
| |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
@@ -170,10 +170,10 @@ Proof.
unfold_is_bounded.
- change ZToWord64 with Interpretations.Word64.ZToWord64 in *.
- rewrite !Interpretations.Word64.word64ToZ_sub;
- rewrite !Interpretations.Word64.word64ToZ_land;
- rewrite !Interpretations.Word64.word64ToZ_ZToWord64;
+ change ZToWord64 with Interpretations64.WordW.ZToWordW in *.
+ rewrite !Interpretations64.WordW.wordWToZ_sub;
+ rewrite !Interpretations64.WordW.wordWToZ_land;
+ rewrite !Interpretations64.WordW.wordWToZ_ZToWordW;
repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
try match goal with
| |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply ModularBaseSystemListZOperationsProofs.neg_range; omega
diff --git a/src/SpecificGen/GF5211_32Reflective.v b/src/SpecificGen/GF5211_32Reflective.v
index 332d5c7e9..415fcad74 100644
--- a/src/SpecificGen/GF5211_32Reflective.v
+++ b/src/SpecificGen/GF5211_32Reflective.v
@@ -10,9 +10,9 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.MapInterp.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Import Crypto.Reflection.Z.Syntax.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Common.
@@ -50,10 +50,10 @@ Declare Reduction asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe5211_32W curry_unop_fe5211_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
Ltac asm_interp
:= cbv beta iota delta
@@ -61,10 +61,10 @@ Ltac asm_interp
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe5211_32W curry_unop_fe5211_32W curry_unop_wire_digitsW
- Word64.interp_op Word64.interp_base_type
+ WordW.interp_op WordW.interp_base_type
Z.interp_op Z.interp_base_type
Z.Syntax.interp_op Z.Syntax.interp_base_type
- mapf_interp_flat_type map_interp Word64.interp_base_type MapInterp mapf_interp word64ize rword64ize
+ mapf_interp_flat_type map_interp WordW.interp_base_type MapInterp mapf_interp word64ize rword64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
diff --git a/src/SpecificGen/GF5211_32Reflective/Common.v b/src/SpecificGen/GF5211_32Reflective/Common.v
index a2b9a7c63..ec776afd9 100644
--- a/src/SpecificGen/GF5211_32Reflective/Common.v
+++ b/src/SpecificGen/GF5211_32Reflective/Common.v
@@ -4,9 +4,9 @@ Require Export Crypto.SpecificGen.GF5211_32.
Require Export Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Reify.
Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.Z.Interpretations.
-Require Crypto.Reflection.Z.Interpretations.Relations.
-Require Import Crypto.Reflection.Z.Interpretations.RelationsCombinations.
+Require Import Crypto.Reflection.Z.Interpretations64.
+Require Crypto.Reflection.Z.Interpretations64.Relations.
+Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
@@ -19,7 +19,7 @@ Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
-Notation Expr := (Expr base_type Word64.interp_base_type op).
+Notation Expr := (Expr base_type WordW.interp_base_type op).
Local Ltac make_type_from uncurried_op :=
let T := (type of uncurried_op) in
@@ -45,9 +45,9 @@ Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |})
+ | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
| (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ constr:((Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs))
end.
Local Ltac make_bounds ls :=
@@ -66,17 +66,16 @@ Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZB
Proof. make_bounds (Tuple.to_list _ bounds). Defined.
Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
-
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- := fun e => curry_binop_fe5211_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_binop_fe5211_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- := fun e => curry_unop_fe5211_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe5211_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.word64
- := fun e => curry_unop_fe5211_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe5211_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe5211_32W (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_fe5211_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF5211_32BoundedCommon.wire_digitsW -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- := fun e => curry_unop_wire_digitsW (Interp (@Word64.interp_op) e).
+ := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
:= (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
@@ -117,13 +116,13 @@ Notation correct_and_bounded_genT ropW'v ropZ_sigv
:= (let ropW' := ropW'v in
let ropZ_sig := ropZ_sigv in
let ropW := MapInterp (fun _ x => x) ropW' in
- let ropZ := MapInterp Word64.to_Z ropW' in
- let ropBounds := MapInterp ZBounds.of_word64 ropW' in
- let ropBoundedWord64 := MapInterp BoundedWord64.of_word64 ropW' in
+ let ropZ := MapInterp WordW.to_Z ropW' in
+ let ropBounds := MapInterp ZBounds.of_wordW ropW' in
+ let ropBoundedWordW := MapInterp BoundedWordW.of_wordW ropW' in
ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_word64 (Interp (@BoundedWord64.interp_op) ropBoundedWord64) (Interp (@Word64.interp_op) ropW))
+ /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
+ /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
+ /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
(only parsing).
Ltac app_tuples x y :=
@@ -138,7 +137,7 @@ Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' Word64.word64 n)
+ (v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
(pf : List.fold_right
andb true
@@ -149,7 +148,7 @@ Fixpoint args_to_bounded_helperT {n}
(fun bounds v =>
let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
bounds
- (Tuple.map (n:=S n) Word64.word64ToZ v))) = true)
+ (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
(res : Type)
{struct n}
: Type.
@@ -161,9 +160,9 @@ Proof.
(Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
-> Type)
with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat Word64.bit_width)%Z, res
+ | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
| S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat Word64.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
+ forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
end v bounds pf).
{ clear -pf0.
abstract (
@@ -178,15 +177,15 @@ Defined.
Fixpoint args_to_bounded_helper {n} res
{struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
+ : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWord64.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord64.lower := fst bounds ; BoundedWord64.value := v ; BoundedWord64.upper := snd bounds |}
+ refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
+ | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
| S n'
=> fun v bounds pf f pf'
=> @args_to_bounded_helper
n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord64.lower := fst (snd bounds) ; BoundedWord64.value := snd v ; BoundedWord64.upper := snd (snd bounds) |}))
+ (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
end.
{ clear -pf pf'.
unfold Tuple.map2, Tuple.map in pf; simpl in *.
@@ -242,16 +241,16 @@ Local Ltac args_to_bounded x H :=
).
Definition unop_args_to_bounded (x : fe5211_32W) (H : is_bounded (fe5211_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
Proof. args_to_bounded x H. Defined.
Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprUnOpWireToFET).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe5211_32W * fe5211_32W)
(H : is_bounded (fe5211_32WToZ (fst x)) = true)
(H' : is_bounded (fe5211_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWord64.BoundedWord) (all_binders_for ExprBinOpT).
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
Proof.
let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
exact v.
@@ -328,13 +327,13 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
change interp_base_type with (@Z.interp_base_type) in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 Word64.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_word64_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
| reflexivity
| assumption
| progress destruct_head' base_type
- | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' BoundedWordW.BoundedWord
| progress destruct_head' and
| progress repeat apply conj ];
specialize (Hbounds_left args H0);
@@ -350,7 +349,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWord64.to_word64' BoundedWord64.boundedWordToWord64 BoundedWord64.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_word64_boundsi' Relations.related'_word64_bounds ZBounds.upper ZBounds.lower Application.remove_all_binders Word64.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
match goal with
| [ |- fe5211_32WToZ ?x = _ /\ _ ]
=> destruct x; destruct_head_hnf' prod
@@ -359,7 +358,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- change word64ToZ with Word64.word64ToZ in *;
+ change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
@@ -371,7 +370,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
unfold_is_bounded;
destruct_head' and;
Z.ltb_to_lt;
- change Word64.word64ToZ with word64ToZ in *;
+ change WordW.wordWToZ with word64ToZ in *;
repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
@@ -387,10 +386,10 @@ Ltac rexpr_correct :=
[ | apply @RelWfMapInterp, wf_ropW ].. ];
auto with interp_related.
-Notation rword_of_Z rexprZ_sig := (MapInterp Word64.of_Z (proj1_sig rexprZ_sig)) (only parsing).
+Notation rword_of_Z rexprZ_sig := (MapInterp WordW.of_Z (proj1_sig rexprZ_sig)) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_wordW) opW)) bounds)
(only parsing).
@@ -407,7 +406,7 @@ Module Export PrettyPrinting.
:= fun t : base_type
=> match t return ZBounds.interp_base_type t -> match t with TZ => bounds_on end with
| TZ => fun x => match x with
- | Some {| ZBounds.lower := l ; ZBounds.upper := u |}
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
=> in_range l u
| None
=> overflow
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v b/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
index a6a6e0f28..ccaefeb38 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -18,7 +18,7 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -31,9 +31,9 @@ Lemma ExprBinOp_correct_and_bounded
let Hx := let (Hx, Hy) := Hxy in Hx in
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
index 9a1d5bdef..4c8c024ef 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOp_correct_and_bounded
(x := eta_fe5211_32W x)
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
index 14f3d3dd7..4abf5e85f 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(x := eta_fe5211_32W x)
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
index ab44167a6..821f6529c 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(x := eta_fe5211_32W x)
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
index 9948ea3a0..9284bf40f 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
@@ -1,6 +1,6 @@
Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.MapInterp.
@@ -15,7 +15,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWord64.interp_op) (MapInterp BoundedWord64.of_word64 ropW))
+ (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) (MapInterp BoundedWordW.of_wordW ropW))
(LiftOption.to' (Some args)))
with
| Some _ => True
@@ -25,9 +25,9 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(x := eta_wire_digitsW x)
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWord64.BoundedWordToBounds) args in
+ let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_wordW ropW)) (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/copy_bounds.sh b/src/SpecificGen/copy_bounds.sh
index ec7821d19..0c73ac268 100755
--- a/src/SpecificGen/copy_bounds.sh
+++ b/src/SpecificGen/copy_bounds.sh
@@ -6,6 +6,10 @@ cd "$DIR"
FILENAME="$1"
V_FILE_STEM="${FILENAME%.*}"
+BIT_WIDTH=64
+case "$V_FILE_STEM" in
+ *_64) BIT_WIDTH=128;;
+esac
if [ -z "$V_FILE_STEM" ]; then
echo "USAGE: $0 JSON_FILE"
@@ -17,5 +21,5 @@ for ORIG in $(git ls-files "../Specific/**GF25519*.v" | grep -v "../Specific/GF2
echo "$NEW"
NEW_DIR="$(dirname "$NEW")"
mkdir -p "$NEW_DIR" || exit $?
- cat "$ORIG" | sed s'/Specific/SpecificGen/g' | sed s"/25519/${V_FILE_STEM}/g" > "$NEW" || exit $?
+ cat "$ORIG" | sed s"/64/${BIT_WIDTH}/g" | sed s'/Specific/SpecificGen/g' | sed s"/25519/${V_FILE_STEM}/g" > "$NEW" || exit $?
done