aboutsummaryrefslogtreecommitdiff
path: root/src/SpecificGen
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2016-11-13 17:00:44 -0500
committerGravatar Jason Gross <jgross@mit.edu>2016-11-13 17:00:44 -0500
commit646a21fc7271316880edc4e627923e7bdd93065b (patch)
treee4c4ba26f442f5e705a90c7325e1d622845344e1 /src/SpecificGen
parent3d3b942308e09a678641005eabdc2f3761f0edae (diff)
Add SpecificGen/GF*
For bounds analysis
Diffstat (limited to 'src/SpecificGen')
-rw-r--r--src/SpecificGen/GF2213_32Bounded.v462
-rw-r--r--src/SpecificGen/GF2213_32BoundedCommon.v693
-rw-r--r--src/SpecificGen/GF2213_32Reflective.v119
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Common.v438
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonBinOp.v50
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOp.v44
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v44
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v44
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v44
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified.v13
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/Add.v12
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/CarryAdd.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/CarryOpp.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/CarrySub.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/GeModulus.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/Mul.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/Opp.v12
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/Pack.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v17
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/Sub.v12
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/Unpack.v17
-rw-r--r--src/SpecificGen/GF2519_32Bounded.v462
-rw-r--r--src/SpecificGen/GF2519_32BoundedCommon.v693
-rw-r--r--src/SpecificGen/GF2519_32Reflective.v119
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Common.v438
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonBinOp.v50
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOp.v44
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v44
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v44
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v44
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified.v13
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/Add.v12
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/CarryAdd.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/CarryOpp.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/CarrySub.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/GeModulus.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/Mul.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/Opp.v12
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/Pack.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v17
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/Sub.v12
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/Unpack.v17
-rw-r--r--src/SpecificGen/GF25519_32Bounded.v462
-rw-r--r--src/SpecificGen/GF25519_32BoundedCommon.v693
-rw-r--r--src/SpecificGen/GF25519_32Reflective.v119
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Common.v438
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonBinOp.v50
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOp.v44
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v44
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v44
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v44
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified.v13
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/Add.v12
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/CarryAdd.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/CarryOpp.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/CarrySub.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/GeModulus.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/Mul.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/Opp.v12
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/Pack.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v17
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/Sub.v12
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/Unpack.v17
-rw-r--r--src/SpecificGen/GF25519_64Bounded.v462
-rw-r--r--src/SpecificGen/GF25519_64BoundedCommon.v693
-rw-r--r--src/SpecificGen/GF25519_64Reflective.v119
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Common.v438
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonBinOp.v50
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOp.v44
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v44
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v44
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v44
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified.v13
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/Add.v12
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/CarryAdd.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/CarryOpp.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/CarrySub.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/GeModulus.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/Mul.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/Opp.v12
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/Pack.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v17
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/Sub.v12
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/Unpack.v17
-rw-r--r--src/SpecificGen/GF41417_32Bounded.v462
-rw-r--r--src/SpecificGen/GF41417_32BoundedCommon.v693
-rw-r--r--src/SpecificGen/GF41417_32Reflective.v119
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Common.v438
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonBinOp.v50
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOp.v44
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v44
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v44
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v44
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified.v13
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/Add.v12
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/CarryAdd.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/CarryOpp.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/CarrySub.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/GeModulus.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/Mul.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/Opp.v12
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/Pack.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v17
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/Sub.v12
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/Unpack.v17
-rw-r--r--src/SpecificGen/GF5211_32Bounded.v462
-rw-r--r--src/SpecificGen/GF5211_32BoundedCommon.v693
-rw-r--r--src/SpecificGen/GF5211_32Reflective.v119
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Common.v438
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonBinOp.v50
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOp.v44
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v44
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v44
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v44
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified.v13
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/Add.v12
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/CarryAdd.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/CarryOpp.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/CarrySub.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/GeModulus.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/Mul.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/Opp.v12
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/Pack.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v17
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/Sub.v12
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/Unpack.v17
-rwxr-xr-xsrc/SpecificGen/copy_bounds.sh16
127 files changed, 12754 insertions, 0 deletions
diff --git a/src/SpecificGen/GF2213_32Bounded.v b/src/SpecificGen/GF2213_32Bounded.v
new file mode 100644
index 000000000..9a35ce124
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Bounded.v
@@ -0,0 +1,462 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF2213_32.
+Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
+Require Import Crypto.SpecificGen.GF2213_32Reflective.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+
+Local Ltac bounded_t opW blem :=
+ apply blem; apply is_bounded_proj1_fe2213_32.
+Local Ltac bounded_wire_digits_t opW blem :=
+ apply blem; apply is_bounded_proj1_wire_digits.
+
+Local Ltac define_binop f g opW blem :=
+ refine (exist_fe2213_32W (opW (proj1_fe2213_32W f) (proj1_fe2213_32W g)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop f opW blem :=
+ refine (exist_fe2213_32W (opW (proj1_fe2213_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_FEToZ f opW :=
+ refine (opW (proj1_fe2213_32W f)).
+Local Ltac define_unop_FEToWire f opW blem :=
+ refine (exist_wire_digitsW (opW (proj1_fe2213_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_WireToFE f opW blem :=
+ refine (exist_fe2213_32W (opW (proj1_wire_digitsW f)) _);
+ 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 Arguments interp_radd / _ _.
+Local Arguments interp_rsub / _ _.
+Local Arguments interp_rmul / _ _.
+Local Arguments interp_ropp / _.
+Local Arguments interp_rprefreeze / _.
+Local Arguments interp_rge_modulus / _.
+Local Arguments interp_rpack / _.
+Local Arguments interp_runpack / _.
+Definition addW (f g : fe2213_32W) : fe2213_32W := Eval simpl in interp_radd f g.
+Definition subW (f g : fe2213_32W) : fe2213_32W := Eval simpl in interp_rsub f g.
+Definition mulW (f g : fe2213_32W) : fe2213_32W := Eval simpl in interp_rmul f g.
+Definition oppW (f : fe2213_32W) : fe2213_32W := Eval simpl in interp_ropp f.
+Definition prefreezeW (f : fe2213_32W) : fe2213_32W := Eval simpl in interp_rprefreeze f.
+Definition ge_modulusW (f : fe2213_32W) : word64 := Eval simpl in interp_rge_modulus f.
+Definition packW (f : fe2213_32W) : wire_digitsW := Eval simpl in interp_rpack f.
+Definition unpackW (f : wire_digitsW) : fe2213_32W := Eval simpl in interp_runpack f.
+
+Require Import ModularBaseSystemWord.
+Definition modulusW :=
+ Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z 10 GF2213_32.modulus_digits_)).
+
+Definition postfreeze : GF2213_32.fe2213_32 -> GF2213_32.fe2213_32 :=
+ GF2213_32.postfreeze.
+
+Lemma freeze_prepost_freeze : forall x, postfreeze (prefreeze x) = GF2213_32.freeze x. Proof. reflexivity. Qed.
+
+Definition postfreezeW : fe2213_32W -> fe2213_32W :=
+ (conditional_subtract_modulusW
+ (num_limbs := 10)
+ modulusW
+ ge_modulusW
+ (Interpretations.Word64.neg GF2213_32.int_width)
+ ).
+
+Definition freezeW (f : fe2213_32W) : fe2213_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
+
+Local Transparent Let_In.
+Definition powW (f : fe2213_32W) chain := fold_chain_opt (proj1_fe2213_32W one) mulW chain [f].
+Definition invW (f : fe2213_32W) : fe2213_32W
+ := Eval cbv -[Let_In fe2213_32W mulW] in powW f (chain inv_ec).
+
+Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
+ change opW with (interp_rop);
+ rewrite pre_rewrite;
+ intros; apply rop_cb; assumption.
+
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Proof. port_correct_and_bounded interp_radd_correct addW interp_radd radd_correct_and_bounded. Qed.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW carry_sub.
+Proof. port_correct_and_bounded interp_rsub_correct subW interp_rsub rsub_correct_and_bounded. Qed.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW mul.
+Proof. port_correct_and_bounded interp_rmul_correct mulW interp_rmul rmul_correct_and_bounded. Qed.
+Lemma oppW_correct_and_bounded : iunop_correct_and_bounded oppW carry_opp.
+Proof. port_correct_and_bounded interp_ropp_correct oppW interp_ropp ropp_correct_and_bounded. Qed.
+Lemma prefreezeW_correct_and_bounded : iunop_correct_and_bounded prefreezeW prefreeze.
+Proof. port_correct_and_bounded interp_rprefreeze_correct prefreezeW interp_rprefreeze rprefreeze_correct_and_bounded. Qed.
+Lemma ge_modulusW_correct : iunop_FEToZ_correct ge_modulusW ge_modulus.
+Proof. port_correct_and_bounded interp_rge_modulus_correct ge_modulusW interp_rge_modulus rge_modulus_correct_and_bounded. Qed.
+Lemma packW_correct_and_bounded : iunop_FEToWire_correct_and_bounded packW pack.
+Proof. port_correct_and_bounded interp_rpack_correct packW interp_rpack rpack_correct_and_bounded. Qed.
+Lemma unpackW_correct_and_bounded : iunop_WireToFE_correct_and_bounded unpackW unpack.
+Proof. port_correct_and_bounded interp_runpack_correct unpackW interp_runpack runpack_correct_and_bounded. Qed.
+
+(* TODO : move *)
+Lemma neg_range : forall x y, 0 <= x ->
+ 0 <= ModularBaseSystemListZOperations.neg x y < 2 ^ x.
+Proof.
+ intros.
+ split; auto using ModularBaseSystemListZOperationsProofs.neg_nonneg.
+ eapply Z.le_lt_trans; eauto using ModularBaseSystemListZOperationsProofs.neg_upperbound.
+ rewrite Z.ones_equiv.
+ omega.
+Qed.
+
+Ltac lower_bound_minus_ge_modulus :=
+ apply Z.le_0_sub;
+ 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;
+ etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
+ apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
+
+Ltac upper_bound_minus_ge_modulus :=
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | ];
+ eapply Z.le_lt_trans; [ eassumption | ];
+ instantiate; vm_compute; reflexivity.
+
+Lemma postfreezeW_correct_and_bounded : iunop_correct_and_bounded postfreezeW postfreeze.
+Proof.
+ intros x H.
+ pose proof (ge_modulusW_correct x H) as Hgm.
+ destruct_head_hnf' prod.
+ unfold_is_bounded_in H.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ cbv [postfreezeW].
+ cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
+ change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ rewrite Hgm.
+
+ cbv [modulusW Tuple.map].
+ cbv [on_tuple List.map to_list to_list' from_list from_list'
+ Tuple.map2 on_tuple2 ListUtil.map2 fe2213_32WToZ].
+ cbv [postfreeze GF2213_32.postfreeze].
+ cbv [Let_In].
+
+ split.
+ { match goal with
+ |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?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;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ | |- (_,_) = (_,_) => reflexivity
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ end. }
+
+
+ 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;
+ repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ try match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ | |- _ - _ <= _ => etransitivity; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | instantiate; assumption ]
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- ModularBaseSystemListZOperations.neg _ _ < _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- _ => vm_compute; (discriminate || reflexivity)
+ end.
+Qed.
+
+Lemma freezeW_correct_and_bounded : iunop_correct_and_bounded freezeW freeze.
+Proof.
+ intros f H; rewrite <- freeze_prepost_freeze.
+ change (freezeW f) with (postfreezeW (prefreezeW f)).
+ destruct (prefreezeW_correct_and_bounded f H) as [H0 H1].
+ destruct (postfreezeW_correct_and_bounded _ H1) as [H0' H1'].
+ rewrite H1', H0', H0; split; reflexivity.
+Qed.
+
+Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
+Proof.
+ cbv [powW pow].
+ intro x; intros; apply (fold_chain_opt_gen fe2213_32WToZ is_bounded [x]).
+ { reflexivity. }
+ { reflexivity. }
+ { intros; progress rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption.
+ apply mulW_correct_and_bounded; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption; reflexivity. }
+ { intros [|?]; autorewrite with simpl_nth_default;
+ (assumption || reflexivity). }
+Qed.
+
+Lemma invW_correct_and_bounded : iunop_correct_and_bounded invW inv.
+Proof.
+ intro f.
+ assert (H : forall f, invW f = powW f (chain inv_ec))
+ by abstract (cbv -[Let_In fe2213_32W mulW]; reflexivity).
+ rewrite H.
+ rewrite inv_correct.
+ cbv [inv_opt].
+ rewrite <- pow_correct.
+ apply powW_correct_and_bounded.
+Qed.
+
+Definition fieldwisebW_sig (f g : fe2213_32W)
+ : { b | b = GF2213_32.fieldwiseb (fe2213_32WToZ f) (fe2213_32WToZ g) }.
+Proof.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ cbv [GF2213_32.fieldwiseb fe2213_32WToZ].
+ rewrite ?word64eqb_Zeqb.
+ reflexivity.
+Defined.
+
+Definition fieldwisebW (f g : fe2213_32W) : bool :=
+ Eval cbv [proj1_sig fieldwisebW_sig appify2 app_fe2213_32W] in
+ appify2 (fun f g => proj1_sig (fieldwisebW_sig f g)) f g.
+
+Lemma fieldwisebW_correct f g
+ : fieldwisebW f g = GF2213_32.fieldwiseb (fe2213_32WToZ f) (fe2213_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (fieldwisebW_sig f' g')).
+Qed.
+
+Local Arguments freezeW : simpl never.
+Local Arguments fe2213_32WToZ !_ / .
+Local Opaque freezeW.
+
+Definition eqbW_sig (f g : fe2213_32W)
+ : { b | is_bounded (fe2213_32WToZ f) = true
+ -> is_bounded (fe2213_32WToZ g) = true
+ -> b = GF2213_32.eqb (fe2213_32WToZ f) (fe2213_32WToZ g) }.
+Proof.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded f pf)) as frf.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded g pf)) as frg.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ unfold GF2213_32.eqb.
+ simpl @fe2213_32WToZ in *; cbv beta iota.
+ intros.
+ rewrite <- frf, <- frg by assumption.
+ rewrite <- fieldwisebW_correct.
+ reflexivity.
+Defined.
+
+Definition eqbW (f g : fe2213_32W) : bool :=
+ Eval cbv [proj1_sig eqbW_sig appify2 app_fe2213_32W] in
+ appify2 (fun f g => proj1_sig (eqbW_sig f g)) f g.
+
+Lemma eqbW_correct f g
+ : is_bounded (fe2213_32WToZ f) = true
+ -> is_bounded (fe2213_32WToZ g) = true
+ -> eqbW f g = GF2213_32.eqb (fe2213_32WToZ f) (fe2213_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (eqbW_sig f' g')).
+Qed.
+
+(* TODO(jgross): use NToWord or such for this constant too *)
+Definition sqrt_m1W : fe2213_32W :=
+ Eval vm_compute in fe2213_32ZToW sqrt_m1.
+
+Definition GF2213_32sqrt x : GF2213_32.fe2213_32 :=
+ dlet powx := powW (fe2213_32ZToW x) (chain GF2213_32.sqrt_ec) in
+ GF2213_32.sqrt (fe2213_32WToZ powx) (fe2213_32WToZ (mulW powx powx)) x.
+
+Definition sqrtW_sig
+ : { sqrtW | iunop_correct_and_bounded sqrtW GF2213_32sqrt }.
+Proof.
+ eexists.
+ unfold GF2213_32sqrt, GF2213_32.sqrt.
+ intros.
+ rewrite !fe2213_32ZToW_WToZ.
+ split.
+ { etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody_eq; intros.
+ set_evars.
+ match goal with (* unfold the first dlet ... in, but only if it's binding a var *)
+ | [ |- ?x = dlet y := fe2213_32WToZ ?z in ?f ]
+ => is_var z; change (x = match fe2213_32WToZ z with y => f end)
+ end.
+ change sqrt_m1 with (fe2213_32WToZ sqrt_m1W).
+ rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded sqrt_m1W a X Y)), <- eqbW_correct, (pull_bool_if fe2213_32WToZ)
+ by repeat match goal with
+ | _ => progress subst
+ | [ |- is_bounded (fe2213_32WToZ ?op) = true ]
+ => lazymatch op with
+ | mulW _ _ => apply mulW_correct_and_bounded
+ | powW _ _ => apply powW_correct_and_bounded
+ | sqrt_m1W => vm_compute; reflexivity
+ | _ => assumption
+ end
+ end.
+ subst_evars; reflexivity.
+ } Unfocus.
+ lazymatch goal with
+ | [ |- context G[dlet x := ?v in fe2213_32WToZ (@?f x)] ]
+ => let G' := context G[fe2213_32WToZ (dlet x := v in f x)] in
+ cut G'; cbv beta;
+ [ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
+ end.
+ reflexivity. }
+ { cbv [Let_In].
+ break_if.
+ { apply powW_correct_and_bounded; assumption. }
+ { apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ].
+ apply powW_correct_and_bounded; assumption. } }
+Defined.
+
+Definition sqrtW (f : fe2213_32W) : fe2213_32W :=
+ Eval cbv [proj1_sig sqrtW_sig app_fe2213_32W] in
+ app_fe2213_32W f (proj1_sig sqrtW_sig).
+
+Lemma sqrtW_correct_and_bounded : iunop_correct_and_bounded sqrtW GF2213_32sqrt.
+Proof.
+ intro f.
+ set (f' := f).
+ hnf in f; destruct_head' prod.
+ assert (H : sqrtW f' = proj1_sig sqrtW_sig f')
+ by (subst f'; cbv beta iota delta [proj1_sig sqrtW_sig sqrtW]; reflexivity).
+ rewrite H.
+ exact (proj2_sig sqrtW_sig f').
+Qed.
+
+
+
+Definition add (f g : fe2213_32) : fe2213_32.
+Proof. define_binop f g addW addW_correct_and_bounded. Defined.
+Definition sub (f g : fe2213_32) : fe2213_32.
+Proof. define_binop f g subW subW_correct_and_bounded. Defined.
+Definition mul (f g : fe2213_32) : fe2213_32.
+Proof. define_binop f g mulW mulW_correct_and_bounded. Defined.
+Definition opp (f : fe2213_32) : fe2213_32.
+Proof. define_unop f oppW oppW_correct_and_bounded. Defined.
+Definition freeze (f : fe2213_32) : fe2213_32.
+Proof. define_unop f freezeW freezeW_correct_and_bounded. Defined.
+Definition ge_modulus (f : fe2213_32) : word64.
+Proof. define_unop_FEToZ f ge_modulusW. Defined.
+Definition pack (f : fe2213_32) : wire_digits.
+Proof. define_unop_FEToWire f packW packW_correct_and_bounded. Defined.
+Definition unpack (f : wire_digits) : fe2213_32.
+Proof. define_unop_WireToFE f unpackW unpackW_correct_and_bounded. Defined.
+
+Definition pow (f : fe2213_32) (chain : list (nat * nat)) : fe2213_32.
+Proof. define_unop f (fun x => powW x chain) powW_correct_and_bounded. Defined.
+Definition inv (f : fe2213_32) : fe2213_32.
+Proof. define_unop f invW invW_correct_and_bounded. Defined.
+Definition sqrt (f : fe2213_32) : fe2213_32.
+Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
+
+Local Ltac op_correct_t op opW_correct_and_bounded :=
+ cbv [op];
+ lazymatch goal with
+ | [ |- context[proj1_fe2213_32 (exist_fe2213_32W _ _)] ]
+ => rewrite proj1_fe2213_32_exist_fe2213_32W
+ | [ |- context[proj1_wire_digits (exist_wire_digitsW _ _)] ]
+ => rewrite proj1_wire_digits_exist_wire_digitsW
+ | _ => idtac
+ end;
+ apply opW_correct_and_bounded;
+ lazymatch goal with
+ | [ |- is_bounded _ = true ]
+ => apply is_bounded_proj1_fe2213_32
+ | [ |- wire_digits_is_bounded _ = true ]
+ => apply is_bounded_proj1_wire_digits
+ end.
+
+Lemma add_correct (f g : fe2213_32) : proj1_fe2213_32 (add f g) = carry_add (proj1_fe2213_32 f) (proj1_fe2213_32 g).
+Proof. op_correct_t add addW_correct_and_bounded. Qed.
+Lemma sub_correct (f g : fe2213_32) : proj1_fe2213_32 (sub f g) = carry_sub (proj1_fe2213_32 f) (proj1_fe2213_32 g).
+Proof. op_correct_t sub subW_correct_and_bounded. Qed.
+Lemma mul_correct (f g : fe2213_32) : proj1_fe2213_32 (mul f g) = GF2213_32.mul (proj1_fe2213_32 f) (proj1_fe2213_32 g).
+Proof. op_correct_t mul mulW_correct_and_bounded. Qed.
+Lemma opp_correct (f : fe2213_32) : proj1_fe2213_32 (opp f) = carry_opp (proj1_fe2213_32 f).
+Proof. op_correct_t opp oppW_correct_and_bounded. Qed.
+Lemma freeze_correct (f : fe2213_32) : proj1_fe2213_32 (freeze f) = GF2213_32.freeze (proj1_fe2213_32 f).
+Proof. op_correct_t freeze freezeW_correct_and_bounded. Qed.
+Lemma ge_modulus_correct (f : fe2213_32) : word64ToZ (ge_modulus f) = GF2213_32.ge_modulus (proj1_fe2213_32 f).
+Proof. op_correct_t ge_modulus ge_modulusW_correct. Qed.
+Lemma pack_correct (f : fe2213_32) : proj1_wire_digits (pack f) = GF2213_32.pack (proj1_fe2213_32 f).
+Proof. op_correct_t pack packW_correct_and_bounded. Qed.
+Lemma unpack_correct (f : wire_digits) : proj1_fe2213_32 (unpack f) = GF2213_32.unpack (proj1_wire_digits f).
+Proof. op_correct_t unpack unpackW_correct_and_bounded. Qed.
+Lemma pow_correct (f : fe2213_32) chain : proj1_fe2213_32 (pow f chain) = GF2213_32.pow (proj1_fe2213_32 f) chain.
+Proof. op_correct_t pow (powW_correct_and_bounded chain). Qed.
+Lemma inv_correct (f : fe2213_32) : proj1_fe2213_32 (inv f) = GF2213_32.inv (proj1_fe2213_32 f).
+Proof. op_correct_t inv invW_correct_and_bounded. Qed.
+Lemma sqrt_correct (f : fe2213_32) : proj1_fe2213_32 (sqrt f) = GF2213_32sqrt (proj1_fe2213_32 f).
+Proof. op_correct_t sqrt sqrtW_correct_and_bounded. Qed.
+
+Import Morphisms.
+
+Local Existing Instance prime_modulus.
+
+Lemma field2213_32_and_homomorphisms
+ : @field fe2213_32 eq zero one opp add sub mul inv div
+ /\ @Ring.is_homomorphism (F _) (@Logic.eq _) 1%F F.add F.mul fe2213_32 eq one add mul encode
+ /\ @Ring.is_homomorphism fe2213_32 eq one add mul (F _) (@Logic.eq _) 1%F F.add F.mul decode.
+Proof.
+ eapply @Field.field_and_homomorphism_from_redundant_representation.
+ { exact (F.field_modulo _). }
+ { cbv [decode encode]; intros; rewrite !proj1_fe2213_32_exist_fe2213_32; apply encode_rep. }
+ { reflexivity. }
+ { reflexivity. }
+ { reflexivity. }
+ { cbv [decode encode]; intros; rewrite opp_correct, carry_opp_correct, carry_opp_opt_correct, carry_opp_rep; apply opp_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite add_correct, carry_add_correct, carry_add_opt_correct, carry_add_rep; apply add_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite sub_correct, carry_sub_correct, carry_sub_opt_correct, carry_sub_rep; apply sub_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite mul_correct, GF2213_32.mul_correct, carry_mul_opt_correct by reflexivity; apply carry_mul_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite inv_correct, GF2213_32.inv_correct, inv_opt_correct by reflexivity; apply inv_rep; reflexivity. }
+ { cbv [decode encode div]; intros; rewrite !proj1_fe2213_32_exist_fe2213_32; apply encode_rep. }
+Qed.
+
+Global Instance field2213_32 : @field fe2213_32 eq zero one opp add sub mul inv div := proj1 field2213_32_and_homomorphisms.
+
+Local Opaque proj1_fe2213_32 exist_fe2213_32 proj1_fe2213_32W exist_fe2213_32W.
+Global Instance homomorphism_F2213_32_encode
+ : @Ring.is_homomorphism (F modulus) Logic.eq F.one F.add F.mul fe2213_32 eq one add mul encode.
+Proof. apply field2213_32_and_homomorphisms. Qed.
+
+Global Instance homomorphism_F2213_32_decode
+ : @Ring.is_homomorphism fe2213_32 eq one add mul (F modulus) Logic.eq F.one F.add F.mul decode.
+Proof. apply field2213_32_and_homomorphisms. Qed.
diff --git a/src/SpecificGen/GF2213_32BoundedCommon.v b/src/SpecificGen/GF2213_32BoundedCommon.v
new file mode 100644
index 000000000..df3f4739e
--- /dev/null
+++ b/src/SpecificGen/GF2213_32BoundedCommon.v
@@ -0,0 +1,693 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF2213_32.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.HList.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+(* BEGIN common curve-specific definitions *)
+Definition bit_width : nat := Eval compute in Z.to_nat (GF2213_32.int_width).
+Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
+Definition bounds_exp : tuple Z length_fe2213_32
+ := Eval compute in
+ Tuple.from_list length_fe2213_32 limb_widths eq_refl.
+Definition bounds : tuple (Z * Z) length_fe2213_32
+ := Eval compute in
+ Tuple.map (fun e => b_of e) bounds_exp.
+Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
+ := Eval compute in Tuple.from_list _ wire_widths eq_refl.
+Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
+ := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
+(* 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.
+Arguments bit_width : simpl never.
+Global Opaque word64.
+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},
+ 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},
+ 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
+ := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
+ Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ : 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}
+ : 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 }.
+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 fe2213_32W := Eval cbv (*-[word64]*) in (tuple word64 length_fe2213_32).
+Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
+Definition fe2213_32 :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => word_of e) bounds_exp.
+Definition wire_digits :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
+
+Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
+ := let res := Tuple.map2
+ (fun bounds v =>
+ let '(lower, upper) := bounds in
+ (lower <=? v) && (v <=? upper))%bool%Z
+ bounds x in
+ List.fold_right andb true (Tuple.to_list _ res).
+
+Definition is_bounded (x : SpecificGen.GF2213_32.fe2213_32) : bool
+ := is_bounded_gen (n:=length_fe2213_32) x bounds.
+
+Definition wire_digits_is_bounded (x : SpecificGen.GF2213_32.wire_digits) : bool
+ := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
+
+(* END basic types *)
+
+Section generic_destructuring.
+ Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
+ := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
+ end.
+ Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
+ := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => @app_on' A n'
+ end.
+ Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
+ Qed.
+ Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
+
+ Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
+ end.
+ Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => @app_on_h' A F n'
+ end.
+ Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
+ Qed.
+ Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
+
+ Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
+ : forall (f : tuple A (length wire_widths)), T f
+ := Eval compute in fun f => @app_on A (length wire_widths) T f P.
+ Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
+ := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
+ Definition app_fe2213_32W_dep {A T} (P : forall x : tuple A length_fe2213_32, T x)
+ : forall (f : tuple A length_fe2213_32), T f
+ := Eval compute in fun f => @app_on A length_fe2213_32 T f P.
+ Definition app_fe2213_32W {A T} (f : tuple A length_fe2213_32) (P : tuple A length_fe2213_32 -> T)
+ := Eval compute in @app_fe2213_32W_dep A (fun _ => T) P f.
+ Definition app_fe2213_32_dep {T} (P : forall x : fe2213_32, T x)
+ : forall f : fe2213_32, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe2213_32 bounds_exp T f P.
+ Definition app_fe2213_32 {T} (f : fe2213_32) (P : hlist (fun e => word_of e) bounds_exp -> T)
+ := Eval compute in @app_fe2213_32_dep (fun _ => T) P f.
+ Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
+ : forall f : wire_digits, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
+ Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
+ := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
+
+ Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
+ := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
+ Definition app_fe2213_32W_dep_correct {A T} f P : @app_fe2213_32W_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_fe2213_32W_correct {A T} f P : @app_fe2213_32W A T f P = P f
+ := @app_fe2213_32W_dep_correct A (fun _ => T) f P.
+ Definition app_fe2213_32_dep_correct {T} f P : @app_fe2213_32_dep T P f = P f
+ := app_on_h_correct (fun e => word_of e) bounds_exp f P.
+ Definition app_fe2213_32_correct {T} f P : @app_fe2213_32 T f P = P f
+ := @app_fe2213_32_dep_correct (fun _ => T) f P.
+ Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
+ := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
+ Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
+ := @app_wire_digits_dep_correct (fun _ => T) f P.
+
+ Definition appify2 {T} (op : fe2213_32W -> fe2213_32W -> T) (f g : fe2213_32W) :=
+ app_fe2213_32W f (fun f0 => (app_fe2213_32W g (fun g0 => op f0 g0))).
+
+ Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
+ Proof.
+ intros. cbv [appify2].
+ etransitivity; apply app_fe2213_32W_correct.
+ Qed.
+End generic_destructuring.
+
+Definition eta_fe2213_32W_sig (x : fe2213_32W) : { v : fe2213_32W | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+Defined.
+Definition eta_fe2213_32W (x : fe2213_32W) : fe2213_32W
+ := Eval cbv [proj1_sig eta_fe2213_32W_sig] in proj1_sig (eta_fe2213_32W_sig x).
+Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+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.
+Proof. apply NToWord_wordToN. Qed.
+Local Opaque word64.
+
+Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ rewrite N2Z.id, NToWord_wordToN; reflexivity.
+Qed.
+
+(* BEGIN precomputation. *)
+
+Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
+Definition word_bounded {lower upper} (v : bounded_word lower upper)
+ : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
+ := project_is_boundedT (proj2_sig v).
+Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
+ := exist _ proj_word (Build_is_boundedT word_bounded).
+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)
+ : bounded_word lower upper
+ := Build_bounded_word'
+ proj_word
+ (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
+ | 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.
+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.
+ match goal with
+ | [ |- context[andb ?x ?y] ]
+ => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
+ end;
+ intros; omega.
+Qed.
+Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
+Proof.
+ refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
+ abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
+Defined.
+Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
+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.
+Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe2213_32W app_fe2213_32 HList.mapt HList.mapt' Tuple.map on_tuple List.map List.app length_fe2213_32 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
+Definition fe2213_32WToZ (x : fe2213_32W) : SpecificGen.GF2213_32.fe2213_32
+ := Eval app_tuple_map in
+ app_fe2213_32W x (Tuple.map (fun v : word64 => v : Z)).
+Definition fe2213_32ZToW (x : SpecificGen.GF2213_32.fe2213_32) : fe2213_32W
+ := Eval app_tuple_map in
+ app_fe2213_32W x (Tuple.map (fun v : Z => v : word64)).
+Definition wire_digitsWToZ (x : wire_digitsW) : SpecificGen.GF2213_32.wire_digits
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
+Definition wire_digitsZToW (x : SpecificGen.GF2213_32.wire_digits) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
+Definition fe2213_32W_word64ize (x : fe2213_32W) : fe2213_32W
+ := Eval app_tuple_map in
+ app_fe2213_32W x (Tuple.map word64ize).
+Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map word64ize).
+
+(** TODO: Turn this into a lemma to speed up proofs *)
+Ltac unfold_is_bounded_in H :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe2213_32WToZ, wire_digitsWToZ in H;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe2213_32 List.length wire_widths] in H;
+ rewrite ?Bool.andb_true_iff in H.
+
+Ltac unfold_is_bounded :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe2213_32WToZ, wire_digitsWToZ;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe2213_32 List.length wire_widths];
+ 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.
+Local Opaque bit_width.
+
+Local Ltac prove_lt_bit_width :=
+ rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
+
+Lemma fe2213_32ZToW_WToZ (x : fe2213_32W) : fe2213_32ZToW (fe2213_32WToZ x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe2213_32WToZ fe2213_32ZToW].
+ rewrite !ZToWord64_word64ToZ; reflexivity.
+Qed.
+
+Lemma fe2213_32WToZ_ZToW x : is_bounded x = true -> fe2213_32WToZ (fe2213_32ZToW x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe2213_32WToZ fe2213_32ZToW].
+ intro H.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma fe2213_32W_word64ize_id x : fe2213_32W_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [fe2213_32W_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [wire_digitsW_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+
+Definition uncurry_unop_fe2213_32W {T} (op : fe2213_32W -> T)
+ := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe2213_32) op.
+Definition curry_unop_fe2213_32W {T} op : fe2213_32W -> T
+ := Eval cbv (*-[word64]*) in fun f => app_fe2213_32W f (Tuple.curry (n:=length_fe2213_32) op).
+Definition uncurry_binop_fe2213_32W {T} (op : fe2213_32W -> fe2213_32W -> T)
+ := Eval cbv (*-[word64]*) in uncurry_unop_fe2213_32W (fun f => uncurry_unop_fe2213_32W (op f)).
+Definition curry_binop_fe2213_32W {T} op : fe2213_32W -> fe2213_32W -> T
+ := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe2213_32W (curry_unop_fe2213_32W op f)).
+
+Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
+ := Eval cbv (*-[word64]*) 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).
+
+
+Definition proj1_fe2213_32W (x : fe2213_32) : fe2213_32W
+ := Eval app_tuple_map in
+ app_fe2213_32 x (HList.mapt (fun _ => (@proj_word _ _))).
+Coercion proj1_fe2213_32 (x : fe2213_32) : SpecificGen.GF2213_32.fe2213_32
+ := fe2213_32WToZ (proj1_fe2213_32W x).
+
+Lemma is_bounded_proj1_fe2213_32 (x : fe2213_32) : is_bounded (proj1_fe2213_32 x) = true.
+Proof.
+ revert x; refine (app_fe2213_32_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [is_bounded proj1_fe2213_32 proj1_fe2213_32W fe2213_32WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe2213_32 is_bounded_gen].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
+ := app_wire_digits x (HList.mapt (fun _ => proj_word)).
+Coercion proj1_wire_digits (x : wire_digits) : SpecificGen.GF2213_32.wire_digits
+ := wire_digitsWToZ (proj1_wire_digitsW x).
+
+Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
+Proof.
+ revert x; refine (app_wire_digits_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Local Ltac make_exist_W' x app_W_dep :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x);
+ cbv [tuple tuple' length_fe2213_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v H :=
+ first [ let v' := (eval cbv [snd fst] in (snd v)) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
+ | let v' := (eval cbv [snd fst] in v) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
+ let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
+ let T := type of H' in
+ let T := (eval cbv [id
+ List.fold_right List.map List.length List.app ListUtil.map2 List.rev
+ Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
+ fe2213_32 bounds fe2213_32WToZ length_fe2213_32
+ wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
+ let H' := constr:(H' : T) in
+ let v := (eval unfold x' in x') in
+ do_refine v H'.
+Local Ltac make_exist'' x exist_W ZToW :=
+ let H := fresh in
+ intro H; apply (exist_W (ZToW x));
+ abstract (
+ hnf in x; destruct_head' prod;
+ let H' := fresh in
+ pose proof H as H';
+ unfold_is_bounded_in H;
+ destruct_head' and; simpl in *;
+ Z.ltb_to_lt;
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ assumption
+ ).
+Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x) in *;
+ cbv [tuple tuple' length_fe2213_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v :=
+ first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
+ | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
+ let v := (eval unfold x' in (exist'' x' H)) in
+ do_refine v.
+
+Definition exist_fe2213_32W' (x : fe2213_32W) : is_bounded (fe2213_32WToZ x) = true -> fe2213_32.
+Proof. make_exist_W' x (@app_fe2213_32W_dep). Defined.
+Definition exist_fe2213_32W (x : fe2213_32W) : is_bounded (fe2213_32WToZ x) = true -> fe2213_32
+ := Eval cbv [app_fe2213_32W_dep exist_fe2213_32W' fe2213_32ZToW] in exist_fe2213_32W' x.
+Definition exist_fe2213_32'' (x : SpecificGen.GF2213_32.fe2213_32) : is_bounded x = true -> fe2213_32.
+Proof. make_exist'' x exist_fe2213_32W fe2213_32ZToW. Defined.
+Definition exist_fe2213_32' (x : SpecificGen.GF2213_32.fe2213_32) : is_bounded x = true -> fe2213_32.
+Proof. make_exist' x (@app_fe2213_32W_dep) exist_fe2213_32'' exist_fe2213_32W fe2213_32ZToW. Defined.
+Definition exist_fe2213_32 (x : SpecificGen.GF2213_32.fe2213_32) : is_bounded x = true -> fe2213_32
+ := Eval cbv [exist_fe2213_32' exist_fe2213_32W exist_fe2213_32' app_fe2213_32 app_fe2213_32W_dep] in
+ exist_fe2213_32' x.
+
+Lemma proj1_fe2213_32_exist_fe2213_32W x pf : proj1_fe2213_32 (exist_fe2213_32W x pf) = fe2213_32WToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe2213_32W_exist_fe2213_32 x pf : proj1_fe2213_32W (exist_fe2213_32 x pf) = fe2213_32ZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe2213_32_exist_fe2213_32 x pf : proj1_fe2213_32 (exist_fe2213_32 x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_fe2213_32 exist_fe2213_32 proj1_fe2213_32W fe2213_32WToZ proj_word Build_bounded_word Build_bounded_word'].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition exist_wire_digitsW' (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
+Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
+Definition exist_wire_digitsW (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
+ := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
+Definition exist_wire_digits'' (x : SpecificGen.GF2213_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits' (x : SpecificGen.GF2213_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits (x : SpecificGen.GF2213_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits
+ := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
+ exist_wire_digits' x.
+
+Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 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 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.
+End opt.
+
+Local Transparent bit_width.
+Local Ltac do_change lem :=
+ match lem with
+ | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
+ => 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
+ let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
+ do_change L'
+ end
+ | _ => lem
+ end.
+Definition fe2213_32_word64ize (x : fe2213_32) : fe2213_32.
+Proof.
+ set (x' := x).
+ hnf in x; destruct_head' prod.
+ let lem := constr:(exist_fe2213_32W (fe2213_32W_word64ize (proj1_fe2213_32W x'))) in
+ let lem := (eval cbv [proj1_fe2213_32W x' fe2213_32W_word64ize proj_word exist_fe2213_32W Build_bounded_word' Build_bounded_word] in lem) in
+ let lem := do_change lem in
+ refine (lem _);
+ change (is_bounded (fe2213_32WToZ (fe2213_32W_word64ize (proj1_fe2213_32W x'))) = true);
+ abstract (rewrite fe2213_32W_word64ize_id; apply is_bounded_proj1_fe2213_32).
+Defined.
+Definition wire_digits_word64ize (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 := do_change lem in
+ let lem := (eval cbv [word64ize opt.word64ize 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).
+Defined.
+
+Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
+ : forall n : nat,
+ (n < length limb_widths)%nat
+ -> (0 <= nth_default 0 (Tuple.to_list length_fe2213_32 x) n <=
+ snd (b_of (nth_default (-1) limb_widths n)))%Z.
+Proof.
+ hnf in x; destruct_head' prod.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ unfold nth_default; simpl.
+ intros.
+ repeat match goal with
+ | [ |- context[nth_error _ ?x] ]
+ => is_var x; destruct x; simpl
+ end;
+ omega.
+Qed.
+
+(* END precomputation *)
+
+(* Precompute constants *)
+
+Definition one' := Eval vm_compute in exist_fe2213_32 SpecificGen.GF2213_32.one_ eq_refl.
+Definition one := Eval cbv [one' fe2213_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe2213_32_word64ize one'.
+
+Definition zero' := Eval vm_compute in exist_fe2213_32 SpecificGen.GF2213_32.zero_ eq_refl.
+Definition zero := Eval cbv [zero' fe2213_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe2213_32_word64ize 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)
+ (Hid : id = F id')
+ (Hop_bounded : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> is_bounded (op (F x) (F y)) = true)
+ (Hop : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> op (F x) (F y) = F (op' x y))
+ (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
+ : F (fold_chain_opt id' op' chain ls)
+ = fold_chain_opt id op chain (List.map F ls)
+ /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
+Proof.
+ rewrite !fold_chain_opt_correct.
+ revert dependent ls; induction chain as [|x xs IHxs]; intros.
+ { pose proof (Hls_bounded 0%nat).
+ destruct ls; simpl; split; trivial; congruence. }
+ { destruct x; simpl; unfold Let_In; simpl.
+ rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
+ { do 2 f_equal.
+ rewrite <- Hop, Hid by auto.
+ rewrite !map_nth_default_always.
+ split; try reflexivity.
+ apply (IHxs (_::_)).
+ intros [|?]; autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. }
+ { intros [|?]; simpl;
+ autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. } }
+Qed.
+
+Lemma encode_bounded x : is_bounded (encode x) = true.
+Proof.
+ pose proof (bounded_encode x).
+ generalize dependent (encode x).
+ intro t; compute in t; intros.
+ destruct_head' prod.
+ unfold Pow2Base.bounded in H.
+ cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params2213_32] in H.
+ repeat match type of H with
+ | context[nth_error (cons _ _) _]
+ => let H' := fresh in
+ pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
+ cbv beta iota in H'
+ end.
+ clear H.
+ simpl in *.
+ cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
+ unfold is_bounded.
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [is_bounded proj1_fe2213_32 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe2213_32].
+ repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
+Qed.
+
+Definition encode (x : F modulus) : fe2213_32
+ := exist_fe2213_32 (encode x) (encode_bounded x).
+
+Definition decode (x : fe2213_32) : F modulus
+ := ModularBaseSystem.decode (proj1_fe2213_32 x).
+
+Lemma proj1_fe2213_32_encode x
+ : proj1_fe2213_32 (encode x) = ModularBaseSystem.encode x.
+Proof.
+ cbv [encode].
+ generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
+ intros y pf; intros; hnf in y; destruct_head_hnf' prod.
+ cbv [proj1_fe2213_32 exist_fe2213_32 proj1_fe2213_32W Build_bounded_word Build_bounded_word' fe2213_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma decode_exist_fe2213_32 x pf
+ : decode (exist_fe2213_32 x pf) = ModularBaseSystem.decode x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [decode proj1_fe2213_32 exist_fe2213_32 proj1_fe2213_32W Build_bounded_word Build_bounded_word' fe2213_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition div (f g : fe2213_32) : fe2213_32
+ := exist_fe2213_32 (div (proj1_fe2213_32 f) (proj1_fe2213_32 g)) (encode_bounded _).
+
+Definition eq (f g : fe2213_32) : Prop := eq (proj1_fe2213_32 f) (proj1_fe2213_32 g).
+
+
+Notation ibinop_correct_and_bounded irop op
+ := (forall x y,
+ is_bounded (fe2213_32WToZ x) = true
+ -> is_bounded (fe2213_32WToZ y) = true
+ -> fe2213_32WToZ (irop x y) = op (fe2213_32WToZ x) (fe2213_32WToZ y)
+ /\ is_bounded (fe2213_32WToZ (irop x y)) = true) (only parsing).
+Notation iunop_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe2213_32WToZ x) = true
+ -> fe2213_32WToZ (irop x) = op (fe2213_32WToZ x)
+ /\ is_bounded (fe2213_32WToZ (irop x)) = true) (only parsing).
+Notation iunop_FEToZ_correct irop op
+ := (forall x,
+ is_bounded (fe2213_32WToZ x) = true
+ -> word64ToZ (irop x) = op (fe2213_32WToZ x)) (only parsing).
+Notation iunop_FEToWire_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe2213_32WToZ x) = true
+ -> wire_digitsWToZ (irop x) = op (fe2213_32WToZ x)
+ /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
+Notation iunop_WireToFE_correct_and_bounded irop op
+ := (forall x,
+ wire_digits_is_bounded (wire_digitsWToZ x) = true
+ -> fe2213_32WToZ (irop x) = op (wire_digitsWToZ x)
+ /\ is_bounded (fe2213_32WToZ (irop x)) = true) (only parsing).
+
+Definition prefreeze := GF2213_32.prefreeze.
diff --git a/src/SpecificGen/GF2213_32Reflective.v b/src/SpecificGen/GF2213_32Reflective.v
new file mode 100644
index 000000000..840cfe090
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective.v
@@ -0,0 +1,119 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Export Crypto.SpecificGen.GF2213_32.
+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.Reify.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.SpecificGen.GF2213_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+Definition radd : ExprBinOp := Eval vm_compute in rcarry_addW.
+Definition rsub : ExprBinOp := Eval vm_compute in rcarry_subW.
+Definition rmul : ExprBinOp := Eval vm_compute in rmulW.
+Definition ropp : ExprUnOp := Eval vm_compute in rcarry_oppW.
+Definition rprefreeze : ExprUnOp := Eval vm_compute in rprefreezeW.
+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.
+
+Declare Reduction asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+Ltac asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+
+
+Definition interp_radd : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := Eval asm_interp in interp_bexpr (rword64ize radd).
+(*Print interp_radd.*)
+Definition interp_radd_correct : interp_radd = interp_bexpr radd := eq_refl.
+Definition interp_rsub : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rsub).
+(*Print interp_rsub.*)
+Definition interp_rsub_correct : interp_rsub = interp_bexpr rsub := eq_refl.
+Definition interp_rmul : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rmul).
+(*Print interp_rmul.*)
+Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
+Definition interp_ropp : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := Eval asm_interp in interp_uexpr (rword64ize ropp).
+(*Print interp_ropp.*)
+Definition interp_ropp_correct : interp_ropp = interp_uexpr ropp := eq_refl.
+Definition interp_rprefreeze : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := Eval asm_interp in interp_uexpr (rword64ize rprefreeze).
+(*Print interp_rprefreeze.*)
+Definition interp_rprefreeze_correct : interp_rprefreeze = interp_uexpr rprefreeze := eq_refl.
+
+Definition interp_rge_modulus : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.word64
+ := Eval asm_interp in interp_uexpr_FEToZ (rword64ize rge_modulus).
+Definition interp_rge_modulus_correct : interp_rge_modulus = interp_uexpr_FEToZ rge_modulus := eq_refl.
+
+Definition interp_rpack : SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.wire_digitsW
+ := Eval asm_interp in interp_uexpr_FEToWire (rword64ize rpack).
+Definition interp_rpack_correct : interp_rpack = interp_uexpr_FEToWire rpack := eq_refl.
+
+Definition interp_runpack : SpecificGen.GF2213_32BoundedCommon.wire_digitsW -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := Eval asm_interp in interp_uexpr_WireToFE (rword64ize 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.
+Proof. exact rcarry_addW_correct_and_bounded. Qed.
+Lemma rsub_correct_and_bounded : binop_correct_and_bounded rsub carry_sub.
+Proof. exact rcarry_subW_correct_and_bounded. Qed.
+Lemma rmul_correct_and_bounded : binop_correct_and_bounded rmul mul.
+Proof. exact rmulW_correct_and_bounded. Qed.
+Lemma ropp_correct_and_bounded : unop_correct_and_bounded ropp carry_opp.
+Proof. exact rcarry_oppW_correct_and_bounded. Qed.
+Lemma rprefreeze_correct_and_bounded : unop_correct_and_bounded rprefreeze prefreeze.
+Proof. exact rprefreezeW_correct_and_bounded. Qed.
+Lemma rge_modulus_correct_and_bounded : unop_FEToZ_correct rge_modulus ge_modulus.
+Proof. exact rge_modulusW_correct_and_bounded. Qed.
+Lemma rpack_correct_and_bounded : unop_FEToWire_correct_and_bounded rpack pack.
+Proof. exact rpackW_correct_and_bounded. Qed.
+Lemma runpack_correct_and_bounded : unop_WireToFE_correct_and_bounded runpack unpack.
+Proof. exact runpackW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF2213_32Reflective/Common.v b/src/SpecificGen/GF2213_32Reflective/Common.v
new file mode 100644
index 000000000..c01957df8
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Common.v
@@ -0,0 +1,438 @@
+Require Export Coq.ZArith.ZArith.
+Require Export Coq.Strings.String.
+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.Reify.
+Require Export Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.InterpWfRel.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Reflection.MapInterpWf.
+Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.LetIn.
+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).
+
+Local Ltac make_type_from uncurried_op :=
+ let T := (type of uncurried_op) in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ exact rT.
+
+Definition ExprBinOpT : type base_type.
+Proof. make_type_from (uncurry_binop_fe2213_32 carry_add). Defined.
+Definition ExprUnOpT : type base_type.
+Proof. make_type_from (uncurry_unop_fe2213_32 carry_opp). Defined.
+Definition ExprUnOpFEToZT : type base_type.
+Proof. make_type_from (uncurry_unop_fe2213_32 ge_modulus). Defined.
+Definition ExprUnOpWireToFET : type base_type.
+Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Definition ExprUnOpFEToWireT : type base_type.
+Proof. make_type_from (uncurry_unop_fe2213_32 pack). Defined.
+Definition ExprBinOp : Type := Expr ExprBinOpT.
+Definition ExprUnOp : Type := Expr ExprUnOpT.
+Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
+Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
+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 :: ?xs)%list => let bs := bounds_from_list xs in
+ constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ end.
+
+Local Ltac make_bounds ls :=
+ compute;
+ let v := bounds_from_list (List.rev ls) in
+ let v := (eval compute in v) in
+ exact v.
+
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
+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).
+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).
+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).
+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).
+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).
+
+Notation binop_correct_and_bounded rop op
+ := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+Notation unop_correct_and_bounded rop op
+ := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
+Notation unop_FEToZ_correct rop op
+ := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
+Notation unop_FEToWire_correct_and_bounded rop op
+ := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
+Notation unop_WireToFE_correct_and_bounded rop op
+ := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
+
+Ltac rexpr_cbv :=
+ lazymatch goal with
+ | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ => let operf := head oper in
+ let uncurryf := head uncurry in
+ try cbv delta [T]; try cbv delta [oper];
+ try cbv beta iota delta [uncurryf]
+ end;
+ cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+
+Ltac reify_sig :=
+ rexpr_cbv; eexists; Reify_rhs; reflexivity.
+
+Local Notation rexpr_sig T uncurried_op :=
+ { rexprZ
+ | interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op }
+ (only parsing).
+
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (uncurry_binop_fe2213_32 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT (uncurry_unop_fe2213_32 op)) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT (uncurry_unop_fe2213_32 op)) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT (uncurry_unop_fe2213_32 op)) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET (uncurry_unop_wire_digits op)) (only parsing).
+
+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
+ 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))
+ (only parsing).
+
+Ltac app_tuples x y :=
+ let tx := type of x in
+ lazymatch (eval hnf in tx) with
+ | prod _ _ => let xs := app_tuples (snd x) y in
+ constr:((fst x, xs))
+ | _ => constr:((x, y))
+ end.
+
+Local Arguments Tuple.map2 : simpl never.
+Local Arguments Tuple.map : simpl never.
+
+Fixpoint args_to_bounded_helperT {n}
+ (v : Tuple.tuple' Word64.word64 n)
+ (bounds : Tuple.tuple' (Z * Z) n)
+ (pf : List.fold_right
+ andb true
+ (Tuple.to_list
+ _
+ (Tuple.map2
+ (n:=S 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)
+ (res : Type)
+ {struct n}
+ : Type.
+Proof.
+ refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
+ List.fold_right
+ _ _ (Tuple.to_list
+ _
+ (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
+ | 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
+ end v bounds pf).
+ { clear -pf0.
+ abstract (
+ destruct v', bounds'; simpl @fst;
+ rewrite Tuple.map_S in pf0;
+ simpl in pf0;
+ rewrite Tuple.map2_S in pf0;
+ simpl @List.fold_right in *;
+ rewrite Bool.andb_true_iff in pf0; tauto
+ ). }
+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.
+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 |}
+ | 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) |}))
+ end.
+ { clear -pf pf'.
+ unfold Tuple.map2, Tuple.map in pf; simpl in *.
+ abstract (
+ destruct bounds;
+ simpl in *;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+ { simpl in *.
+ clear -pf pf'.
+ abstract (
+ destruct bounds as [? [? ?] ], v; simpl in *;
+ rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
+ simpl in pf;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+Defined.
+
+Definition assoc_right''
+ := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
+
+Definition args_to_bounded {n} v bounds pf
+ := Eval cbv [args_to_bounded_helper assoc_right''] in
+ @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
+
+Local Ltac get_len T :=
+ match (eval hnf in T) with
+ | prod ?A ?B
+ => let a := get_len A in
+ let b := get_len B in
+ (eval compute in (a + b)%nat)
+ | _ => constr:(1%nat)
+ end.
+
+Local Ltac args_to_bounded x H :=
+ let x' := fresh in
+ set (x' := x);
+ compute in x;
+ let len := (let T := type of x in get_len T) in
+ destruct_head' prod;
+ let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
+ let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
+ let c := (eval cbv [all_binders_for ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
+ apply c; compute; clear;
+ try abstract (
+ repeat split;
+ solve [ reflexivity
+ | refine (fun v => match v with eq_refl => I end) ]
+ ).
+
+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).
+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).
+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).
+Proof.
+ let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
+ exact v.
+Defined.
+
+Ltac assoc_right_tuple x so_far :=
+ let t := type of x in
+ lazymatch (eval hnf in t) with
+ | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
+ assoc_right_tuple (fst x) so_far
+ | _ => lazymatch so_far with
+ | @None => x
+ | _ => constr:((x, so_far))
+ end
+ end.
+
+Local Ltac make_bounds_prop bounds orig_bounds :=
+ let bounds' := fresh "bounds'" in
+ let bounds_bad := fresh "bounds_bad" in
+ rename bounds into bounds_bad;
+ let boundsv := assoc_right_tuple bounds_bad (@None) in
+ pose boundsv as bounds;
+ pose orig_bounds as bounds';
+ repeat (refine (match fst bounds' with
+ | Some bounds' => let (l, u) := fst bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end && _)%bool;
+ destruct bounds' as [_ bounds'], bounds as [_ bounds]);
+ try exact (match bounds' with
+ | Some bounds' => let (l, u) := bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end).
+
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprBinOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToWireT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpWireToFET)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToZT)) : bool.
+Proof.
+ refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
+Defined.
+
+(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
+ various kinds of correct and boundedness, and abstract in Gallina
+ rather than Ltac *)
+
+Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
+ let Heq := fresh "Heq" in
+ let Hbounds0 := fresh "Hbounds0" in
+ let Hbounds1 := fresh "Hbounds1" in
+ let Hbounds2 := fresh "Hbounds2" in
+ pose proof (proj2_sig ropZ_sig) as Heq;
+ cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
+ curry_binop_fe2213_32W curry_unop_fe2213_32W curry_unop_wire_digitsW
+ curry_binop_fe2213_32 curry_unop_fe2213_32 curry_unop_wire_digits
+ uncurry_binop_fe2213_32W uncurry_unop_fe2213_32W uncurry_unop_wire_digitsW
+ uncurry_binop_fe2213_32 uncurry_unop_fe2213_32 uncurry_unop_wire_digits
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET
+ interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ cbv zeta in *;
+ simpl @fe2213_32WToZ; simpl @wire_digitsWToZ;
+ rewrite <- Heq; clear Heq;
+ destruct Hbounds as [Heq Hbounds];
+ change interp_op with (@Z.interp_op) in *;
+ 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;
+ specialize_by repeat first [ progress intros
+ | reflexivity
+ | assumption
+ | progress destruct_head' base_type
+ | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' and
+ | progress repeat apply conj ];
+ specialize (Hbounds_left args H0);
+ specialize (Hbounds_right args H0);
+ cbv beta in *;
+ lazymatch type of Hbounds_right with
+ | match ?e with _ => _ end
+ => lazymatch type of H1 with
+ | match ?e' with _ => _ end
+ => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
+ end
+ end;
+ 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;
+ match goal with
+ | [ |- fe2213_32WToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- wire_digitsWToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- _ = _ ]
+ => exact Hbounds_left
+ end;
+ change word64ToZ with Word64.word64ToZ in *;
+ (split; [ exact Hbounds_left | ]);
+ cbv [interp_flat_type] in *;
+ cbv [fst snd
+ binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good
+ ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds] in H1;
+ destruct_head' ZBounds.bounds;
+ unfold_is_bounded_in H1;
+ simpl @fe2213_32WToZ; simpl @wire_digitsWToZ;
+ unfold_is_bounded;
+ destruct_head' and;
+ Z.ltb_to_lt;
+ change Word64.word64ToZ with word64ToZ in *;
+ repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
+
+
+Ltac rexpr_correct :=
+ let ropW' := fresh in
+ let ropZ_sig := fresh in
+ intros ropW' ropZ_sig;
+ let wf_ropW := fresh "wf_ropW" in
+ assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
+ cbv zeta; repeat apply conj;
+ [ vm_compute; reflexivity
+ | apply @InterpRelWf;
+ [ | 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 compute_bounds opW bounds
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ (only parsing).
+
+
+Module Export PrettyPrinting.
+ (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
+ 8.5/8.6. Because [Set] is special and things break if
+ [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
+ to debug. *)
+ Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
+
+ Inductive result := yes | no | borked.
+
+ Definition ZBounds_to_bounds_on
+ := 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 |}
+ => in_range l u
+ | None
+ => overflow
+ end
+ end.
+
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result
+ := match t return interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result with
+ | Tbase TZ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
+ | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
+ | no, no => no
+ | yes, no | no, yes | yes, yes => yes
+ | borked, _ | _, borked => borked
+ end
+ end.
+
+ (** This gives a slightly easier to read version of the bounds *)
+ Notation compute_bounds_for_display opW bounds
+ := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
+ Notation sanity_compute opW bounds
+ := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
+ Notation sanity_check opW bounds
+ := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
+End PrettyPrinting.
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v b/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
new file mode 100644
index 000000000..be248d708
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
@@ -0,0 +1,50 @@
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprBinOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_binop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall xy
+ (xy := (eta_fe2213_32W (fst xy), eta_fe2213_32W (snd xy)))
+ (Hxy : is_bounded (fe2213_32WToZ (fst xy)) = true
+ /\ is_bounded (fe2213_32WToZ (snd xy)) = true),
+ let Hx := let (Hx, Hy) := Hxy in Hx in
+ 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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall xy
+ (xy := (eta_fe2213_32W (fst xy), eta_fe2213_32W (snd xy)))
+ (Hxy : is_bounded (fe2213_32WToZ (fst xy)) = true
+ /\ is_bounded (fe2213_32WToZ (snd xy)) = true),
+ 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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => binop_bounds_good bounds = true
+ | None => False
+ end)
+ : binop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x y Hx Hy.
+ pose x as x'; pose y as y'.
+ hnf in x, y; destruct_head' prod.
+ specialize (H0 (x', y') (conj Hx Hy)).
+ specialize (H1 (x', y') (conj Hx Hy)).
+ let args := constr:(binop_args_to_bounded (x', y') Hx Hy) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
new file mode 100644
index 000000000..3d52bcf78
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe2213_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unop_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
new file mode 100644
index 000000000..a08efcfd8
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToWire_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe2213_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToWire_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToWire_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
new file mode 100644
index 000000000..1f3a5f8a0
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToZ_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe2213_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToZ_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToZ_correct (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
new file mode 100644
index 000000000..727936a13
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpWireToFE_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_wire_digitsW x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopWireToFE_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_WireToFE_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified.v b/src/SpecificGen/GF2213_32Reflective/Reified.v
new file mode 100644
index 000000000..445c5d28d
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified.v
@@ -0,0 +1,13 @@
+(** We split the reification up into separate files, one operation per
+ file, so that it can run in parallel. *)
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.Add.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.CarryAdd.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.Sub.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.CarrySub.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.Mul.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.Opp.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.CarryOpp.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.PreFreeze.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.GeModulus.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.Pack.
+Require Export Crypto.SpecificGen.GF2213_32Reflective.Reified.Unpack.
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/Add.v b/src/SpecificGen/GF2213_32Reflective/Reified/Add.v
new file mode 100644
index 000000000..d278a8f26
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/Add.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonBinOp.
+
+Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
+Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
+Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
+Proof. rexpr_correct. Qed.
+Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
+Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
+Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/CarryAdd.v b/src/SpecificGen/GF2213_32Reflective/Reified/CarryAdd.v
new file mode 100644
index 000000000..516a36bf5
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/CarryAdd.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonBinOp.
+
+Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
+Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
+Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_addW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/CarryOpp.v b/src/SpecificGen/GF2213_32Reflective/Reified/CarryOpp.v
new file mode 100644
index 000000000..a2ab83ce2
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/CarryOpp.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOp.
+
+Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
+Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
+Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_oppW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/CarrySub.v b/src/SpecificGen/GF2213_32Reflective/Reified/CarrySub.v
new file mode 100644
index 000000000..04a097dd7
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/CarrySub.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonBinOp.
+
+Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
+Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
+Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_subW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/GeModulus.v b/src/SpecificGen/GF2213_32Reflective/Reified/GeModulus.v
new file mode 100644
index 000000000..659e712d9
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/GeModulus.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOpFEToZ.
+
+Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
+Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
+Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rge_modulusW_correct_and_bounded
+ := ExprUnOpFEToZ_correct_and_bounded
+ rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/Mul.v b/src/SpecificGen/GF2213_32Reflective/Reified/Mul.v
new file mode 100644
index 000000000..08c235a9c
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/Mul.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonBinOp.
+
+Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
+Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
+Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rmulW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
+Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
+Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/Opp.v b/src/SpecificGen/GF2213_32Reflective/Reified/Opp.v
new file mode 100644
index 000000000..091cb96de
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/Opp.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOp.
+
+Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
+Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
+Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
+Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
+Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/Pack.v b/src/SpecificGen/GF2213_32Reflective/Reified/Pack.v
new file mode 100644
index 000000000..f96f81a53
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/Pack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOpFEToWire.
+
+Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
+Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
+Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rpackW_correct_and_bounded
+ := ExprUnOpFEToWire_correct_and_bounded
+ rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v
new file mode 100644
index 000000000..60ab908b0
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOp.
+
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF2213_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
+Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rprefreezeW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/Sub.v b/src/SpecificGen/GF2213_32Reflective/Reified/Sub.v
new file mode 100644
index 000000000..cc7ba69dd
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/Sub.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonBinOp.
+
+Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
+Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
+Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
+Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
+Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/Unpack.v b/src/SpecificGen/GF2213_32Reflective/Reified/Unpack.v
new file mode 100644
index 000000000..5ef4ff3d8
--- /dev/null
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/Unpack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOpWireToFE.
+
+Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
+Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
+Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition runpackW_correct_and_bounded
+ := ExprUnOpWireToFE_correct_and_bounded
+ runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/SpecificGen/GF2519_32Bounded.v b/src/SpecificGen/GF2519_32Bounded.v
new file mode 100644
index 000000000..292ee9adb
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Bounded.v
@@ -0,0 +1,462 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF2519_32.
+Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
+Require Import Crypto.SpecificGen.GF2519_32Reflective.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+
+Local Ltac bounded_t opW blem :=
+ apply blem; apply is_bounded_proj1_fe2519_32.
+Local Ltac bounded_wire_digits_t opW blem :=
+ apply blem; apply is_bounded_proj1_wire_digits.
+
+Local Ltac define_binop f g opW blem :=
+ refine (exist_fe2519_32W (opW (proj1_fe2519_32W f) (proj1_fe2519_32W g)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop f opW blem :=
+ refine (exist_fe2519_32W (opW (proj1_fe2519_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_FEToZ f opW :=
+ refine (opW (proj1_fe2519_32W f)).
+Local Ltac define_unop_FEToWire f opW blem :=
+ refine (exist_wire_digitsW (opW (proj1_fe2519_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_WireToFE f opW blem :=
+ refine (exist_fe2519_32W (opW (proj1_wire_digitsW f)) _);
+ 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 Arguments interp_radd / _ _.
+Local Arguments interp_rsub / _ _.
+Local Arguments interp_rmul / _ _.
+Local Arguments interp_ropp / _.
+Local Arguments interp_rprefreeze / _.
+Local Arguments interp_rge_modulus / _.
+Local Arguments interp_rpack / _.
+Local Arguments interp_runpack / _.
+Definition addW (f g : fe2519_32W) : fe2519_32W := Eval simpl in interp_radd f g.
+Definition subW (f g : fe2519_32W) : fe2519_32W := Eval simpl in interp_rsub f g.
+Definition mulW (f g : fe2519_32W) : fe2519_32W := Eval simpl in interp_rmul f g.
+Definition oppW (f : fe2519_32W) : fe2519_32W := Eval simpl in interp_ropp f.
+Definition prefreezeW (f : fe2519_32W) : fe2519_32W := Eval simpl in interp_rprefreeze f.
+Definition ge_modulusW (f : fe2519_32W) : word64 := Eval simpl in interp_rge_modulus f.
+Definition packW (f : fe2519_32W) : wire_digitsW := Eval simpl in interp_rpack f.
+Definition unpackW (f : wire_digitsW) : fe2519_32W := Eval simpl in interp_runpack f.
+
+Require Import ModularBaseSystemWord.
+Definition modulusW :=
+ Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z 10 GF2519_32.modulus_digits_)).
+
+Definition postfreeze : GF2519_32.fe2519_32 -> GF2519_32.fe2519_32 :=
+ GF2519_32.postfreeze.
+
+Lemma freeze_prepost_freeze : forall x, postfreeze (prefreeze x) = GF2519_32.freeze x. Proof. reflexivity. Qed.
+
+Definition postfreezeW : fe2519_32W -> fe2519_32W :=
+ (conditional_subtract_modulusW
+ (num_limbs := 10)
+ modulusW
+ ge_modulusW
+ (Interpretations.Word64.neg GF2519_32.int_width)
+ ).
+
+Definition freezeW (f : fe2519_32W) : fe2519_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
+
+Local Transparent Let_In.
+Definition powW (f : fe2519_32W) chain := fold_chain_opt (proj1_fe2519_32W one) mulW chain [f].
+Definition invW (f : fe2519_32W) : fe2519_32W
+ := Eval cbv -[Let_In fe2519_32W mulW] in powW f (chain inv_ec).
+
+Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
+ change opW with (interp_rop);
+ rewrite pre_rewrite;
+ intros; apply rop_cb; assumption.
+
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Proof. port_correct_and_bounded interp_radd_correct addW interp_radd radd_correct_and_bounded. Qed.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW carry_sub.
+Proof. port_correct_and_bounded interp_rsub_correct subW interp_rsub rsub_correct_and_bounded. Qed.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW mul.
+Proof. port_correct_and_bounded interp_rmul_correct mulW interp_rmul rmul_correct_and_bounded. Qed.
+Lemma oppW_correct_and_bounded : iunop_correct_and_bounded oppW carry_opp.
+Proof. port_correct_and_bounded interp_ropp_correct oppW interp_ropp ropp_correct_and_bounded. Qed.
+Lemma prefreezeW_correct_and_bounded : iunop_correct_and_bounded prefreezeW prefreeze.
+Proof. port_correct_and_bounded interp_rprefreeze_correct prefreezeW interp_rprefreeze rprefreeze_correct_and_bounded. Qed.
+Lemma ge_modulusW_correct : iunop_FEToZ_correct ge_modulusW ge_modulus.
+Proof. port_correct_and_bounded interp_rge_modulus_correct ge_modulusW interp_rge_modulus rge_modulus_correct_and_bounded. Qed.
+Lemma packW_correct_and_bounded : iunop_FEToWire_correct_and_bounded packW pack.
+Proof. port_correct_and_bounded interp_rpack_correct packW interp_rpack rpack_correct_and_bounded. Qed.
+Lemma unpackW_correct_and_bounded : iunop_WireToFE_correct_and_bounded unpackW unpack.
+Proof. port_correct_and_bounded interp_runpack_correct unpackW interp_runpack runpack_correct_and_bounded. Qed.
+
+(* TODO : move *)
+Lemma neg_range : forall x y, 0 <= x ->
+ 0 <= ModularBaseSystemListZOperations.neg x y < 2 ^ x.
+Proof.
+ intros.
+ split; auto using ModularBaseSystemListZOperationsProofs.neg_nonneg.
+ eapply Z.le_lt_trans; eauto using ModularBaseSystemListZOperationsProofs.neg_upperbound.
+ rewrite Z.ones_equiv.
+ omega.
+Qed.
+
+Ltac lower_bound_minus_ge_modulus :=
+ apply Z.le_0_sub;
+ 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;
+ etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
+ apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
+
+Ltac upper_bound_minus_ge_modulus :=
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | ];
+ eapply Z.le_lt_trans; [ eassumption | ];
+ instantiate; vm_compute; reflexivity.
+
+Lemma postfreezeW_correct_and_bounded : iunop_correct_and_bounded postfreezeW postfreeze.
+Proof.
+ intros x H.
+ pose proof (ge_modulusW_correct x H) as Hgm.
+ destruct_head_hnf' prod.
+ unfold_is_bounded_in H.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ cbv [postfreezeW].
+ cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
+ change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ rewrite Hgm.
+
+ cbv [modulusW Tuple.map].
+ cbv [on_tuple List.map to_list to_list' from_list from_list'
+ Tuple.map2 on_tuple2 ListUtil.map2 fe2519_32WToZ].
+ cbv [postfreeze GF2519_32.postfreeze].
+ cbv [Let_In].
+
+ split.
+ { match goal with
+ |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?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;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ | |- (_,_) = (_,_) => reflexivity
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ end. }
+
+
+ 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;
+ repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ try match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ | |- _ - _ <= _ => etransitivity; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | instantiate; assumption ]
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- ModularBaseSystemListZOperations.neg _ _ < _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- _ => vm_compute; (discriminate || reflexivity)
+ end.
+Qed.
+
+Lemma freezeW_correct_and_bounded : iunop_correct_and_bounded freezeW freeze.
+Proof.
+ intros f H; rewrite <- freeze_prepost_freeze.
+ change (freezeW f) with (postfreezeW (prefreezeW f)).
+ destruct (prefreezeW_correct_and_bounded f H) as [H0 H1].
+ destruct (postfreezeW_correct_and_bounded _ H1) as [H0' H1'].
+ rewrite H1', H0', H0; split; reflexivity.
+Qed.
+
+Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
+Proof.
+ cbv [powW pow].
+ intro x; intros; apply (fold_chain_opt_gen fe2519_32WToZ is_bounded [x]).
+ { reflexivity. }
+ { reflexivity. }
+ { intros; progress rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption.
+ apply mulW_correct_and_bounded; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption; reflexivity. }
+ { intros [|?]; autorewrite with simpl_nth_default;
+ (assumption || reflexivity). }
+Qed.
+
+Lemma invW_correct_and_bounded : iunop_correct_and_bounded invW inv.
+Proof.
+ intro f.
+ assert (H : forall f, invW f = powW f (chain inv_ec))
+ by abstract (cbv -[Let_In fe2519_32W mulW]; reflexivity).
+ rewrite H.
+ rewrite inv_correct.
+ cbv [inv_opt].
+ rewrite <- pow_correct.
+ apply powW_correct_and_bounded.
+Qed.
+
+Definition fieldwisebW_sig (f g : fe2519_32W)
+ : { b | b = GF2519_32.fieldwiseb (fe2519_32WToZ f) (fe2519_32WToZ g) }.
+Proof.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ cbv [GF2519_32.fieldwiseb fe2519_32WToZ].
+ rewrite ?word64eqb_Zeqb.
+ reflexivity.
+Defined.
+
+Definition fieldwisebW (f g : fe2519_32W) : bool :=
+ Eval cbv [proj1_sig fieldwisebW_sig appify2 app_fe2519_32W] in
+ appify2 (fun f g => proj1_sig (fieldwisebW_sig f g)) f g.
+
+Lemma fieldwisebW_correct f g
+ : fieldwisebW f g = GF2519_32.fieldwiseb (fe2519_32WToZ f) (fe2519_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (fieldwisebW_sig f' g')).
+Qed.
+
+Local Arguments freezeW : simpl never.
+Local Arguments fe2519_32WToZ !_ / .
+Local Opaque freezeW.
+
+Definition eqbW_sig (f g : fe2519_32W)
+ : { b | is_bounded (fe2519_32WToZ f) = true
+ -> is_bounded (fe2519_32WToZ g) = true
+ -> b = GF2519_32.eqb (fe2519_32WToZ f) (fe2519_32WToZ g) }.
+Proof.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded f pf)) as frf.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded g pf)) as frg.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ unfold GF2519_32.eqb.
+ simpl @fe2519_32WToZ in *; cbv beta iota.
+ intros.
+ rewrite <- frf, <- frg by assumption.
+ rewrite <- fieldwisebW_correct.
+ reflexivity.
+Defined.
+
+Definition eqbW (f g : fe2519_32W) : bool :=
+ Eval cbv [proj1_sig eqbW_sig appify2 app_fe2519_32W] in
+ appify2 (fun f g => proj1_sig (eqbW_sig f g)) f g.
+
+Lemma eqbW_correct f g
+ : is_bounded (fe2519_32WToZ f) = true
+ -> is_bounded (fe2519_32WToZ g) = true
+ -> eqbW f g = GF2519_32.eqb (fe2519_32WToZ f) (fe2519_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (eqbW_sig f' g')).
+Qed.
+
+(* TODO(jgross): use NToWord or such for this constant too *)
+Definition sqrt_m1W : fe2519_32W :=
+ Eval vm_compute in fe2519_32ZToW sqrt_m1.
+
+Definition GF2519_32sqrt x : GF2519_32.fe2519_32 :=
+ dlet powx := powW (fe2519_32ZToW x) (chain GF2519_32.sqrt_ec) in
+ GF2519_32.sqrt (fe2519_32WToZ powx) (fe2519_32WToZ (mulW powx powx)) x.
+
+Definition sqrtW_sig
+ : { sqrtW | iunop_correct_and_bounded sqrtW GF2519_32sqrt }.
+Proof.
+ eexists.
+ unfold GF2519_32sqrt, GF2519_32.sqrt.
+ intros.
+ rewrite !fe2519_32ZToW_WToZ.
+ split.
+ { etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody_eq; intros.
+ set_evars.
+ match goal with (* unfold the first dlet ... in, but only if it's binding a var *)
+ | [ |- ?x = dlet y := fe2519_32WToZ ?z in ?f ]
+ => is_var z; change (x = match fe2519_32WToZ z with y => f end)
+ end.
+ change sqrt_m1 with (fe2519_32WToZ sqrt_m1W).
+ rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded sqrt_m1W a X Y)), <- eqbW_correct, (pull_bool_if fe2519_32WToZ)
+ by repeat match goal with
+ | _ => progress subst
+ | [ |- is_bounded (fe2519_32WToZ ?op) = true ]
+ => lazymatch op with
+ | mulW _ _ => apply mulW_correct_and_bounded
+ | powW _ _ => apply powW_correct_and_bounded
+ | sqrt_m1W => vm_compute; reflexivity
+ | _ => assumption
+ end
+ end.
+ subst_evars; reflexivity.
+ } Unfocus.
+ lazymatch goal with
+ | [ |- context G[dlet x := ?v in fe2519_32WToZ (@?f x)] ]
+ => let G' := context G[fe2519_32WToZ (dlet x := v in f x)] in
+ cut G'; cbv beta;
+ [ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
+ end.
+ reflexivity. }
+ { cbv [Let_In].
+ break_if.
+ { apply powW_correct_and_bounded; assumption. }
+ { apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ].
+ apply powW_correct_and_bounded; assumption. } }
+Defined.
+
+Definition sqrtW (f : fe2519_32W) : fe2519_32W :=
+ Eval cbv [proj1_sig sqrtW_sig app_fe2519_32W] in
+ app_fe2519_32W f (proj1_sig sqrtW_sig).
+
+Lemma sqrtW_correct_and_bounded : iunop_correct_and_bounded sqrtW GF2519_32sqrt.
+Proof.
+ intro f.
+ set (f' := f).
+ hnf in f; destruct_head' prod.
+ assert (H : sqrtW f' = proj1_sig sqrtW_sig f')
+ by (subst f'; cbv beta iota delta [proj1_sig sqrtW_sig sqrtW]; reflexivity).
+ rewrite H.
+ exact (proj2_sig sqrtW_sig f').
+Qed.
+
+
+
+Definition add (f g : fe2519_32) : fe2519_32.
+Proof. define_binop f g addW addW_correct_and_bounded. Defined.
+Definition sub (f g : fe2519_32) : fe2519_32.
+Proof. define_binop f g subW subW_correct_and_bounded. Defined.
+Definition mul (f g : fe2519_32) : fe2519_32.
+Proof. define_binop f g mulW mulW_correct_and_bounded. Defined.
+Definition opp (f : fe2519_32) : fe2519_32.
+Proof. define_unop f oppW oppW_correct_and_bounded. Defined.
+Definition freeze (f : fe2519_32) : fe2519_32.
+Proof. define_unop f freezeW freezeW_correct_and_bounded. Defined.
+Definition ge_modulus (f : fe2519_32) : word64.
+Proof. define_unop_FEToZ f ge_modulusW. Defined.
+Definition pack (f : fe2519_32) : wire_digits.
+Proof. define_unop_FEToWire f packW packW_correct_and_bounded. Defined.
+Definition unpack (f : wire_digits) : fe2519_32.
+Proof. define_unop_WireToFE f unpackW unpackW_correct_and_bounded. Defined.
+
+Definition pow (f : fe2519_32) (chain : list (nat * nat)) : fe2519_32.
+Proof. define_unop f (fun x => powW x chain) powW_correct_and_bounded. Defined.
+Definition inv (f : fe2519_32) : fe2519_32.
+Proof. define_unop f invW invW_correct_and_bounded. Defined.
+Definition sqrt (f : fe2519_32) : fe2519_32.
+Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
+
+Local Ltac op_correct_t op opW_correct_and_bounded :=
+ cbv [op];
+ lazymatch goal with
+ | [ |- context[proj1_fe2519_32 (exist_fe2519_32W _ _)] ]
+ => rewrite proj1_fe2519_32_exist_fe2519_32W
+ | [ |- context[proj1_wire_digits (exist_wire_digitsW _ _)] ]
+ => rewrite proj1_wire_digits_exist_wire_digitsW
+ | _ => idtac
+ end;
+ apply opW_correct_and_bounded;
+ lazymatch goal with
+ | [ |- is_bounded _ = true ]
+ => apply is_bounded_proj1_fe2519_32
+ | [ |- wire_digits_is_bounded _ = true ]
+ => apply is_bounded_proj1_wire_digits
+ end.
+
+Lemma add_correct (f g : fe2519_32) : proj1_fe2519_32 (add f g) = carry_add (proj1_fe2519_32 f) (proj1_fe2519_32 g).
+Proof. op_correct_t add addW_correct_and_bounded. Qed.
+Lemma sub_correct (f g : fe2519_32) : proj1_fe2519_32 (sub f g) = carry_sub (proj1_fe2519_32 f) (proj1_fe2519_32 g).
+Proof. op_correct_t sub subW_correct_and_bounded. Qed.
+Lemma mul_correct (f g : fe2519_32) : proj1_fe2519_32 (mul f g) = GF2519_32.mul (proj1_fe2519_32 f) (proj1_fe2519_32 g).
+Proof. op_correct_t mul mulW_correct_and_bounded. Qed.
+Lemma opp_correct (f : fe2519_32) : proj1_fe2519_32 (opp f) = carry_opp (proj1_fe2519_32 f).
+Proof. op_correct_t opp oppW_correct_and_bounded. Qed.
+Lemma freeze_correct (f : fe2519_32) : proj1_fe2519_32 (freeze f) = GF2519_32.freeze (proj1_fe2519_32 f).
+Proof. op_correct_t freeze freezeW_correct_and_bounded. Qed.
+Lemma ge_modulus_correct (f : fe2519_32) : word64ToZ (ge_modulus f) = GF2519_32.ge_modulus (proj1_fe2519_32 f).
+Proof. op_correct_t ge_modulus ge_modulusW_correct. Qed.
+Lemma pack_correct (f : fe2519_32) : proj1_wire_digits (pack f) = GF2519_32.pack (proj1_fe2519_32 f).
+Proof. op_correct_t pack packW_correct_and_bounded. Qed.
+Lemma unpack_correct (f : wire_digits) : proj1_fe2519_32 (unpack f) = GF2519_32.unpack (proj1_wire_digits f).
+Proof. op_correct_t unpack unpackW_correct_and_bounded. Qed.
+Lemma pow_correct (f : fe2519_32) chain : proj1_fe2519_32 (pow f chain) = GF2519_32.pow (proj1_fe2519_32 f) chain.
+Proof. op_correct_t pow (powW_correct_and_bounded chain). Qed.
+Lemma inv_correct (f : fe2519_32) : proj1_fe2519_32 (inv f) = GF2519_32.inv (proj1_fe2519_32 f).
+Proof. op_correct_t inv invW_correct_and_bounded. Qed.
+Lemma sqrt_correct (f : fe2519_32) : proj1_fe2519_32 (sqrt f) = GF2519_32sqrt (proj1_fe2519_32 f).
+Proof. op_correct_t sqrt sqrtW_correct_and_bounded. Qed.
+
+Import Morphisms.
+
+Local Existing Instance prime_modulus.
+
+Lemma field2519_32_and_homomorphisms
+ : @field fe2519_32 eq zero one opp add sub mul inv div
+ /\ @Ring.is_homomorphism (F _) (@Logic.eq _) 1%F F.add F.mul fe2519_32 eq one add mul encode
+ /\ @Ring.is_homomorphism fe2519_32 eq one add mul (F _) (@Logic.eq _) 1%F F.add F.mul decode.
+Proof.
+ eapply @Field.field_and_homomorphism_from_redundant_representation.
+ { exact (F.field_modulo _). }
+ { cbv [decode encode]; intros; rewrite !proj1_fe2519_32_exist_fe2519_32; apply encode_rep. }
+ { reflexivity. }
+ { reflexivity. }
+ { reflexivity. }
+ { cbv [decode encode]; intros; rewrite opp_correct, carry_opp_correct, carry_opp_opt_correct, carry_opp_rep; apply opp_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite add_correct, carry_add_correct, carry_add_opt_correct, carry_add_rep; apply add_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite sub_correct, carry_sub_correct, carry_sub_opt_correct, carry_sub_rep; apply sub_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite mul_correct, GF2519_32.mul_correct, carry_mul_opt_correct by reflexivity; apply carry_mul_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite inv_correct, GF2519_32.inv_correct, inv_opt_correct by reflexivity; apply inv_rep; reflexivity. }
+ { cbv [decode encode div]; intros; rewrite !proj1_fe2519_32_exist_fe2519_32; apply encode_rep. }
+Qed.
+
+Global Instance field2519_32 : @field fe2519_32 eq zero one opp add sub mul inv div := proj1 field2519_32_and_homomorphisms.
+
+Local Opaque proj1_fe2519_32 exist_fe2519_32 proj1_fe2519_32W exist_fe2519_32W.
+Global Instance homomorphism_F2519_32_encode
+ : @Ring.is_homomorphism (F modulus) Logic.eq F.one F.add F.mul fe2519_32 eq one add mul encode.
+Proof. apply field2519_32_and_homomorphisms. Qed.
+
+Global Instance homomorphism_F2519_32_decode
+ : @Ring.is_homomorphism fe2519_32 eq one add mul (F modulus) Logic.eq F.one F.add F.mul decode.
+Proof. apply field2519_32_and_homomorphisms. Qed.
diff --git a/src/SpecificGen/GF2519_32BoundedCommon.v b/src/SpecificGen/GF2519_32BoundedCommon.v
new file mode 100644
index 000000000..cea3e4f10
--- /dev/null
+++ b/src/SpecificGen/GF2519_32BoundedCommon.v
@@ -0,0 +1,693 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF2519_32.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.HList.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+(* BEGIN common curve-specific definitions *)
+Definition bit_width : nat := Eval compute in Z.to_nat (GF2519_32.int_width).
+Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
+Definition bounds_exp : tuple Z length_fe2519_32
+ := Eval compute in
+ Tuple.from_list length_fe2519_32 limb_widths eq_refl.
+Definition bounds : tuple (Z * Z) length_fe2519_32
+ := Eval compute in
+ Tuple.map (fun e => b_of e) bounds_exp.
+Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
+ := Eval compute in Tuple.from_list _ wire_widths eq_refl.
+Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
+ := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
+(* 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.
+Arguments bit_width : simpl never.
+Global Opaque word64.
+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},
+ 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},
+ 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
+ := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
+ Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ : 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}
+ : 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 }.
+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 fe2519_32W := Eval cbv (*-[word64]*) in (tuple word64 length_fe2519_32).
+Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
+Definition fe2519_32 :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => word_of e) bounds_exp.
+Definition wire_digits :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
+
+Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
+ := let res := Tuple.map2
+ (fun bounds v =>
+ let '(lower, upper) := bounds in
+ (lower <=? v) && (v <=? upper))%bool%Z
+ bounds x in
+ List.fold_right andb true (Tuple.to_list _ res).
+
+Definition is_bounded (x : SpecificGen.GF2519_32.fe2519_32) : bool
+ := is_bounded_gen (n:=length_fe2519_32) x bounds.
+
+Definition wire_digits_is_bounded (x : SpecificGen.GF2519_32.wire_digits) : bool
+ := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
+
+(* END basic types *)
+
+Section generic_destructuring.
+ Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
+ := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
+ end.
+ Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
+ := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => @app_on' A n'
+ end.
+ Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
+ Qed.
+ Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
+
+ Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
+ end.
+ Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => @app_on_h' A F n'
+ end.
+ Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
+ Qed.
+ Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
+
+ Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
+ : forall (f : tuple A (length wire_widths)), T f
+ := Eval compute in fun f => @app_on A (length wire_widths) T f P.
+ Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
+ := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
+ Definition app_fe2519_32W_dep {A T} (P : forall x : tuple A length_fe2519_32, T x)
+ : forall (f : tuple A length_fe2519_32), T f
+ := Eval compute in fun f => @app_on A length_fe2519_32 T f P.
+ Definition app_fe2519_32W {A T} (f : tuple A length_fe2519_32) (P : tuple A length_fe2519_32 -> T)
+ := Eval compute in @app_fe2519_32W_dep A (fun _ => T) P f.
+ Definition app_fe2519_32_dep {T} (P : forall x : fe2519_32, T x)
+ : forall f : fe2519_32, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe2519_32 bounds_exp T f P.
+ Definition app_fe2519_32 {T} (f : fe2519_32) (P : hlist (fun e => word_of e) bounds_exp -> T)
+ := Eval compute in @app_fe2519_32_dep (fun _ => T) P f.
+ Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
+ : forall f : wire_digits, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
+ Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
+ := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
+
+ Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
+ := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
+ Definition app_fe2519_32W_dep_correct {A T} f P : @app_fe2519_32W_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_fe2519_32W_correct {A T} f P : @app_fe2519_32W A T f P = P f
+ := @app_fe2519_32W_dep_correct A (fun _ => T) f P.
+ Definition app_fe2519_32_dep_correct {T} f P : @app_fe2519_32_dep T P f = P f
+ := app_on_h_correct (fun e => word_of e) bounds_exp f P.
+ Definition app_fe2519_32_correct {T} f P : @app_fe2519_32 T f P = P f
+ := @app_fe2519_32_dep_correct (fun _ => T) f P.
+ Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
+ := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
+ Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
+ := @app_wire_digits_dep_correct (fun _ => T) f P.
+
+ Definition appify2 {T} (op : fe2519_32W -> fe2519_32W -> T) (f g : fe2519_32W) :=
+ app_fe2519_32W f (fun f0 => (app_fe2519_32W g (fun g0 => op f0 g0))).
+
+ Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
+ Proof.
+ intros. cbv [appify2].
+ etransitivity; apply app_fe2519_32W_correct.
+ Qed.
+End generic_destructuring.
+
+Definition eta_fe2519_32W_sig (x : fe2519_32W) : { v : fe2519_32W | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+Defined.
+Definition eta_fe2519_32W (x : fe2519_32W) : fe2519_32W
+ := Eval cbv [proj1_sig eta_fe2519_32W_sig] in proj1_sig (eta_fe2519_32W_sig x).
+Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+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.
+Proof. apply NToWord_wordToN. Qed.
+Local Opaque word64.
+
+Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ rewrite N2Z.id, NToWord_wordToN; reflexivity.
+Qed.
+
+(* BEGIN precomputation. *)
+
+Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
+Definition word_bounded {lower upper} (v : bounded_word lower upper)
+ : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
+ := project_is_boundedT (proj2_sig v).
+Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
+ := exist _ proj_word (Build_is_boundedT word_bounded).
+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)
+ : bounded_word lower upper
+ := Build_bounded_word'
+ proj_word
+ (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
+ | 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.
+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.
+ match goal with
+ | [ |- context[andb ?x ?y] ]
+ => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
+ end;
+ intros; omega.
+Qed.
+Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
+Proof.
+ refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
+ abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
+Defined.
+Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
+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.
+Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe2519_32W app_fe2519_32 HList.mapt HList.mapt' Tuple.map on_tuple List.map List.app length_fe2519_32 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
+Definition fe2519_32WToZ (x : fe2519_32W) : SpecificGen.GF2519_32.fe2519_32
+ := Eval app_tuple_map in
+ app_fe2519_32W x (Tuple.map (fun v : word64 => v : Z)).
+Definition fe2519_32ZToW (x : SpecificGen.GF2519_32.fe2519_32) : fe2519_32W
+ := Eval app_tuple_map in
+ app_fe2519_32W x (Tuple.map (fun v : Z => v : word64)).
+Definition wire_digitsWToZ (x : wire_digitsW) : SpecificGen.GF2519_32.wire_digits
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
+Definition wire_digitsZToW (x : SpecificGen.GF2519_32.wire_digits) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
+Definition fe2519_32W_word64ize (x : fe2519_32W) : fe2519_32W
+ := Eval app_tuple_map in
+ app_fe2519_32W x (Tuple.map word64ize).
+Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map word64ize).
+
+(** TODO: Turn this into a lemma to speed up proofs *)
+Ltac unfold_is_bounded_in H :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe2519_32WToZ, wire_digitsWToZ in H;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe2519_32 List.length wire_widths] in H;
+ rewrite ?Bool.andb_true_iff in H.
+
+Ltac unfold_is_bounded :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe2519_32WToZ, wire_digitsWToZ;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe2519_32 List.length wire_widths];
+ 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.
+Local Opaque bit_width.
+
+Local Ltac prove_lt_bit_width :=
+ rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
+
+Lemma fe2519_32ZToW_WToZ (x : fe2519_32W) : fe2519_32ZToW (fe2519_32WToZ x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe2519_32WToZ fe2519_32ZToW].
+ rewrite !ZToWord64_word64ToZ; reflexivity.
+Qed.
+
+Lemma fe2519_32WToZ_ZToW x : is_bounded x = true -> fe2519_32WToZ (fe2519_32ZToW x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe2519_32WToZ fe2519_32ZToW].
+ intro H.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma fe2519_32W_word64ize_id x : fe2519_32W_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [fe2519_32W_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [wire_digitsW_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+
+Definition uncurry_unop_fe2519_32W {T} (op : fe2519_32W -> T)
+ := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe2519_32) op.
+Definition curry_unop_fe2519_32W {T} op : fe2519_32W -> T
+ := Eval cbv (*-[word64]*) in fun f => app_fe2519_32W f (Tuple.curry (n:=length_fe2519_32) op).
+Definition uncurry_binop_fe2519_32W {T} (op : fe2519_32W -> fe2519_32W -> T)
+ := Eval cbv (*-[word64]*) in uncurry_unop_fe2519_32W (fun f => uncurry_unop_fe2519_32W (op f)).
+Definition curry_binop_fe2519_32W {T} op : fe2519_32W -> fe2519_32W -> T
+ := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe2519_32W (curry_unop_fe2519_32W op f)).
+
+Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
+ := Eval cbv (*-[word64]*) 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).
+
+
+Definition proj1_fe2519_32W (x : fe2519_32) : fe2519_32W
+ := Eval app_tuple_map in
+ app_fe2519_32 x (HList.mapt (fun _ => (@proj_word _ _))).
+Coercion proj1_fe2519_32 (x : fe2519_32) : SpecificGen.GF2519_32.fe2519_32
+ := fe2519_32WToZ (proj1_fe2519_32W x).
+
+Lemma is_bounded_proj1_fe2519_32 (x : fe2519_32) : is_bounded (proj1_fe2519_32 x) = true.
+Proof.
+ revert x; refine (app_fe2519_32_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [is_bounded proj1_fe2519_32 proj1_fe2519_32W fe2519_32WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe2519_32 is_bounded_gen].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
+ := app_wire_digits x (HList.mapt (fun _ => proj_word)).
+Coercion proj1_wire_digits (x : wire_digits) : SpecificGen.GF2519_32.wire_digits
+ := wire_digitsWToZ (proj1_wire_digitsW x).
+
+Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
+Proof.
+ revert x; refine (app_wire_digits_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Local Ltac make_exist_W' x app_W_dep :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x);
+ cbv [tuple tuple' length_fe2519_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v H :=
+ first [ let v' := (eval cbv [snd fst] in (snd v)) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
+ | let v' := (eval cbv [snd fst] in v) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
+ let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
+ let T := type of H' in
+ let T := (eval cbv [id
+ List.fold_right List.map List.length List.app ListUtil.map2 List.rev
+ Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
+ fe2519_32 bounds fe2519_32WToZ length_fe2519_32
+ wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
+ let H' := constr:(H' : T) in
+ let v := (eval unfold x' in x') in
+ do_refine v H'.
+Local Ltac make_exist'' x exist_W ZToW :=
+ let H := fresh in
+ intro H; apply (exist_W (ZToW x));
+ abstract (
+ hnf in x; destruct_head' prod;
+ let H' := fresh in
+ pose proof H as H';
+ unfold_is_bounded_in H;
+ destruct_head' and; simpl in *;
+ Z.ltb_to_lt;
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ assumption
+ ).
+Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x) in *;
+ cbv [tuple tuple' length_fe2519_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v :=
+ first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
+ | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
+ let v := (eval unfold x' in (exist'' x' H)) in
+ do_refine v.
+
+Definition exist_fe2519_32W' (x : fe2519_32W) : is_bounded (fe2519_32WToZ x) = true -> fe2519_32.
+Proof. make_exist_W' x (@app_fe2519_32W_dep). Defined.
+Definition exist_fe2519_32W (x : fe2519_32W) : is_bounded (fe2519_32WToZ x) = true -> fe2519_32
+ := Eval cbv [app_fe2519_32W_dep exist_fe2519_32W' fe2519_32ZToW] in exist_fe2519_32W' x.
+Definition exist_fe2519_32'' (x : SpecificGen.GF2519_32.fe2519_32) : is_bounded x = true -> fe2519_32.
+Proof. make_exist'' x exist_fe2519_32W fe2519_32ZToW. Defined.
+Definition exist_fe2519_32' (x : SpecificGen.GF2519_32.fe2519_32) : is_bounded x = true -> fe2519_32.
+Proof. make_exist' x (@app_fe2519_32W_dep) exist_fe2519_32'' exist_fe2519_32W fe2519_32ZToW. Defined.
+Definition exist_fe2519_32 (x : SpecificGen.GF2519_32.fe2519_32) : is_bounded x = true -> fe2519_32
+ := Eval cbv [exist_fe2519_32' exist_fe2519_32W exist_fe2519_32' app_fe2519_32 app_fe2519_32W_dep] in
+ exist_fe2519_32' x.
+
+Lemma proj1_fe2519_32_exist_fe2519_32W x pf : proj1_fe2519_32 (exist_fe2519_32W x pf) = fe2519_32WToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe2519_32W_exist_fe2519_32 x pf : proj1_fe2519_32W (exist_fe2519_32 x pf) = fe2519_32ZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe2519_32_exist_fe2519_32 x pf : proj1_fe2519_32 (exist_fe2519_32 x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_fe2519_32 exist_fe2519_32 proj1_fe2519_32W fe2519_32WToZ proj_word Build_bounded_word Build_bounded_word'].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition exist_wire_digitsW' (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
+Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
+Definition exist_wire_digitsW (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
+ := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
+Definition exist_wire_digits'' (x : SpecificGen.GF2519_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits' (x : SpecificGen.GF2519_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits (x : SpecificGen.GF2519_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits
+ := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
+ exist_wire_digits' x.
+
+Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 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 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.
+End opt.
+
+Local Transparent bit_width.
+Local Ltac do_change lem :=
+ match lem with
+ | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
+ => 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
+ let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
+ do_change L'
+ end
+ | _ => lem
+ end.
+Definition fe2519_32_word64ize (x : fe2519_32) : fe2519_32.
+Proof.
+ set (x' := x).
+ hnf in x; destruct_head' prod.
+ let lem := constr:(exist_fe2519_32W (fe2519_32W_word64ize (proj1_fe2519_32W x'))) in
+ let lem := (eval cbv [proj1_fe2519_32W x' fe2519_32W_word64ize proj_word exist_fe2519_32W Build_bounded_word' Build_bounded_word] in lem) in
+ let lem := do_change lem in
+ refine (lem _);
+ change (is_bounded (fe2519_32WToZ (fe2519_32W_word64ize (proj1_fe2519_32W x'))) = true);
+ abstract (rewrite fe2519_32W_word64ize_id; apply is_bounded_proj1_fe2519_32).
+Defined.
+Definition wire_digits_word64ize (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 := do_change lem in
+ let lem := (eval cbv [word64ize opt.word64ize 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).
+Defined.
+
+Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
+ : forall n : nat,
+ (n < length limb_widths)%nat
+ -> (0 <= nth_default 0 (Tuple.to_list length_fe2519_32 x) n <=
+ snd (b_of (nth_default (-1) limb_widths n)))%Z.
+Proof.
+ hnf in x; destruct_head' prod.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ unfold nth_default; simpl.
+ intros.
+ repeat match goal with
+ | [ |- context[nth_error _ ?x] ]
+ => is_var x; destruct x; simpl
+ end;
+ omega.
+Qed.
+
+(* END precomputation *)
+
+(* Precompute constants *)
+
+Definition one' := Eval vm_compute in exist_fe2519_32 SpecificGen.GF2519_32.one_ eq_refl.
+Definition one := Eval cbv [one' fe2519_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe2519_32_word64ize one'.
+
+Definition zero' := Eval vm_compute in exist_fe2519_32 SpecificGen.GF2519_32.zero_ eq_refl.
+Definition zero := Eval cbv [zero' fe2519_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe2519_32_word64ize 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)
+ (Hid : id = F id')
+ (Hop_bounded : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> is_bounded (op (F x) (F y)) = true)
+ (Hop : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> op (F x) (F y) = F (op' x y))
+ (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
+ : F (fold_chain_opt id' op' chain ls)
+ = fold_chain_opt id op chain (List.map F ls)
+ /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
+Proof.
+ rewrite !fold_chain_opt_correct.
+ revert dependent ls; induction chain as [|x xs IHxs]; intros.
+ { pose proof (Hls_bounded 0%nat).
+ destruct ls; simpl; split; trivial; congruence. }
+ { destruct x; simpl; unfold Let_In; simpl.
+ rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
+ { do 2 f_equal.
+ rewrite <- Hop, Hid by auto.
+ rewrite !map_nth_default_always.
+ split; try reflexivity.
+ apply (IHxs (_::_)).
+ intros [|?]; autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. }
+ { intros [|?]; simpl;
+ autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. } }
+Qed.
+
+Lemma encode_bounded x : is_bounded (encode x) = true.
+Proof.
+ pose proof (bounded_encode x).
+ generalize dependent (encode x).
+ intro t; compute in t; intros.
+ destruct_head' prod.
+ unfold Pow2Base.bounded in H.
+ cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params2519_32] in H.
+ repeat match type of H with
+ | context[nth_error (cons _ _) _]
+ => let H' := fresh in
+ pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
+ cbv beta iota in H'
+ end.
+ clear H.
+ simpl in *.
+ cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
+ unfold is_bounded.
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [is_bounded proj1_fe2519_32 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe2519_32].
+ repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
+Qed.
+
+Definition encode (x : F modulus) : fe2519_32
+ := exist_fe2519_32 (encode x) (encode_bounded x).
+
+Definition decode (x : fe2519_32) : F modulus
+ := ModularBaseSystem.decode (proj1_fe2519_32 x).
+
+Lemma proj1_fe2519_32_encode x
+ : proj1_fe2519_32 (encode x) = ModularBaseSystem.encode x.
+Proof.
+ cbv [encode].
+ generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
+ intros y pf; intros; hnf in y; destruct_head_hnf' prod.
+ cbv [proj1_fe2519_32 exist_fe2519_32 proj1_fe2519_32W Build_bounded_word Build_bounded_word' fe2519_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma decode_exist_fe2519_32 x pf
+ : decode (exist_fe2519_32 x pf) = ModularBaseSystem.decode x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [decode proj1_fe2519_32 exist_fe2519_32 proj1_fe2519_32W Build_bounded_word Build_bounded_word' fe2519_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition div (f g : fe2519_32) : fe2519_32
+ := exist_fe2519_32 (div (proj1_fe2519_32 f) (proj1_fe2519_32 g)) (encode_bounded _).
+
+Definition eq (f g : fe2519_32) : Prop := eq (proj1_fe2519_32 f) (proj1_fe2519_32 g).
+
+
+Notation ibinop_correct_and_bounded irop op
+ := (forall x y,
+ is_bounded (fe2519_32WToZ x) = true
+ -> is_bounded (fe2519_32WToZ y) = true
+ -> fe2519_32WToZ (irop x y) = op (fe2519_32WToZ x) (fe2519_32WToZ y)
+ /\ is_bounded (fe2519_32WToZ (irop x y)) = true) (only parsing).
+Notation iunop_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe2519_32WToZ x) = true
+ -> fe2519_32WToZ (irop x) = op (fe2519_32WToZ x)
+ /\ is_bounded (fe2519_32WToZ (irop x)) = true) (only parsing).
+Notation iunop_FEToZ_correct irop op
+ := (forall x,
+ is_bounded (fe2519_32WToZ x) = true
+ -> word64ToZ (irop x) = op (fe2519_32WToZ x)) (only parsing).
+Notation iunop_FEToWire_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe2519_32WToZ x) = true
+ -> wire_digitsWToZ (irop x) = op (fe2519_32WToZ x)
+ /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
+Notation iunop_WireToFE_correct_and_bounded irop op
+ := (forall x,
+ wire_digits_is_bounded (wire_digitsWToZ x) = true
+ -> fe2519_32WToZ (irop x) = op (wire_digitsWToZ x)
+ /\ is_bounded (fe2519_32WToZ (irop x)) = true) (only parsing).
+
+Definition prefreeze := GF2519_32.prefreeze.
diff --git a/src/SpecificGen/GF2519_32Reflective.v b/src/SpecificGen/GF2519_32Reflective.v
new file mode 100644
index 000000000..964d0f608
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective.v
@@ -0,0 +1,119 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Export Crypto.SpecificGen.GF2519_32.
+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.Reify.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.SpecificGen.GF2519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+Definition radd : ExprBinOp := Eval vm_compute in rcarry_addW.
+Definition rsub : ExprBinOp := Eval vm_compute in rcarry_subW.
+Definition rmul : ExprBinOp := Eval vm_compute in rmulW.
+Definition ropp : ExprUnOp := Eval vm_compute in rcarry_oppW.
+Definition rprefreeze : ExprUnOp := Eval vm_compute in rprefreezeW.
+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.
+
+Declare Reduction asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+Ltac asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+
+
+Definition interp_radd : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := Eval asm_interp in interp_bexpr (rword64ize radd).
+(*Print interp_radd.*)
+Definition interp_radd_correct : interp_radd = interp_bexpr radd := eq_refl.
+Definition interp_rsub : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rsub).
+(*Print interp_rsub.*)
+Definition interp_rsub_correct : interp_rsub = interp_bexpr rsub := eq_refl.
+Definition interp_rmul : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rmul).
+(*Print interp_rmul.*)
+Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
+Definition interp_ropp : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := Eval asm_interp in interp_uexpr (rword64ize ropp).
+(*Print interp_ropp.*)
+Definition interp_ropp_correct : interp_ropp = interp_uexpr ropp := eq_refl.
+Definition interp_rprefreeze : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := Eval asm_interp in interp_uexpr (rword64ize rprefreeze).
+(*Print interp_rprefreeze.*)
+Definition interp_rprefreeze_correct : interp_rprefreeze = interp_uexpr rprefreeze := eq_refl.
+
+Definition interp_rge_modulus : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.word64
+ := Eval asm_interp in interp_uexpr_FEToZ (rword64ize rge_modulus).
+Definition interp_rge_modulus_correct : interp_rge_modulus = interp_uexpr_FEToZ rge_modulus := eq_refl.
+
+Definition interp_rpack : SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.wire_digitsW
+ := Eval asm_interp in interp_uexpr_FEToWire (rword64ize rpack).
+Definition interp_rpack_correct : interp_rpack = interp_uexpr_FEToWire rpack := eq_refl.
+
+Definition interp_runpack : SpecificGen.GF2519_32BoundedCommon.wire_digitsW -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := Eval asm_interp in interp_uexpr_WireToFE (rword64ize 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.
+Proof. exact rcarry_addW_correct_and_bounded. Qed.
+Lemma rsub_correct_and_bounded : binop_correct_and_bounded rsub carry_sub.
+Proof. exact rcarry_subW_correct_and_bounded. Qed.
+Lemma rmul_correct_and_bounded : binop_correct_and_bounded rmul mul.
+Proof. exact rmulW_correct_and_bounded. Qed.
+Lemma ropp_correct_and_bounded : unop_correct_and_bounded ropp carry_opp.
+Proof. exact rcarry_oppW_correct_and_bounded. Qed.
+Lemma rprefreeze_correct_and_bounded : unop_correct_and_bounded rprefreeze prefreeze.
+Proof. exact rprefreezeW_correct_and_bounded. Qed.
+Lemma rge_modulus_correct_and_bounded : unop_FEToZ_correct rge_modulus ge_modulus.
+Proof. exact rge_modulusW_correct_and_bounded. Qed.
+Lemma rpack_correct_and_bounded : unop_FEToWire_correct_and_bounded rpack pack.
+Proof. exact rpackW_correct_and_bounded. Qed.
+Lemma runpack_correct_and_bounded : unop_WireToFE_correct_and_bounded runpack unpack.
+Proof. exact runpackW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF2519_32Reflective/Common.v b/src/SpecificGen/GF2519_32Reflective/Common.v
new file mode 100644
index 000000000..16a7dd72c
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Common.v
@@ -0,0 +1,438 @@
+Require Export Coq.ZArith.ZArith.
+Require Export Coq.Strings.String.
+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.Reify.
+Require Export Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.InterpWfRel.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Reflection.MapInterpWf.
+Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.LetIn.
+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).
+
+Local Ltac make_type_from uncurried_op :=
+ let T := (type of uncurried_op) in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ exact rT.
+
+Definition ExprBinOpT : type base_type.
+Proof. make_type_from (uncurry_binop_fe2519_32 carry_add). Defined.
+Definition ExprUnOpT : type base_type.
+Proof. make_type_from (uncurry_unop_fe2519_32 carry_opp). Defined.
+Definition ExprUnOpFEToZT : type base_type.
+Proof. make_type_from (uncurry_unop_fe2519_32 ge_modulus). Defined.
+Definition ExprUnOpWireToFET : type base_type.
+Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Definition ExprUnOpFEToWireT : type base_type.
+Proof. make_type_from (uncurry_unop_fe2519_32 pack). Defined.
+Definition ExprBinOp : Type := Expr ExprBinOpT.
+Definition ExprUnOp : Type := Expr ExprUnOpT.
+Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
+Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
+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 :: ?xs)%list => let bs := bounds_from_list xs in
+ constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ end.
+
+Local Ltac make_bounds ls :=
+ compute;
+ let v := bounds_from_list (List.rev ls) in
+ let v := (eval compute in v) in
+ exact v.
+
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
+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).
+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).
+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).
+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).
+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).
+
+Notation binop_correct_and_bounded rop op
+ := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+Notation unop_correct_and_bounded rop op
+ := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
+Notation unop_FEToZ_correct rop op
+ := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
+Notation unop_FEToWire_correct_and_bounded rop op
+ := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
+Notation unop_WireToFE_correct_and_bounded rop op
+ := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
+
+Ltac rexpr_cbv :=
+ lazymatch goal with
+ | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ => let operf := head oper in
+ let uncurryf := head uncurry in
+ try cbv delta [T]; try cbv delta [oper];
+ try cbv beta iota delta [uncurryf]
+ end;
+ cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+
+Ltac reify_sig :=
+ rexpr_cbv; eexists; Reify_rhs; reflexivity.
+
+Local Notation rexpr_sig T uncurried_op :=
+ { rexprZ
+ | interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op }
+ (only parsing).
+
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (uncurry_binop_fe2519_32 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT (uncurry_unop_fe2519_32 op)) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT (uncurry_unop_fe2519_32 op)) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT (uncurry_unop_fe2519_32 op)) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET (uncurry_unop_wire_digits op)) (only parsing).
+
+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
+ 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))
+ (only parsing).
+
+Ltac app_tuples x y :=
+ let tx := type of x in
+ lazymatch (eval hnf in tx) with
+ | prod _ _ => let xs := app_tuples (snd x) y in
+ constr:((fst x, xs))
+ | _ => constr:((x, y))
+ end.
+
+Local Arguments Tuple.map2 : simpl never.
+Local Arguments Tuple.map : simpl never.
+
+Fixpoint args_to_bounded_helperT {n}
+ (v : Tuple.tuple' Word64.word64 n)
+ (bounds : Tuple.tuple' (Z * Z) n)
+ (pf : List.fold_right
+ andb true
+ (Tuple.to_list
+ _
+ (Tuple.map2
+ (n:=S 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)
+ (res : Type)
+ {struct n}
+ : Type.
+Proof.
+ refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
+ List.fold_right
+ _ _ (Tuple.to_list
+ _
+ (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
+ | 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
+ end v bounds pf).
+ { clear -pf0.
+ abstract (
+ destruct v', bounds'; simpl @fst;
+ rewrite Tuple.map_S in pf0;
+ simpl in pf0;
+ rewrite Tuple.map2_S in pf0;
+ simpl @List.fold_right in *;
+ rewrite Bool.andb_true_iff in pf0; tauto
+ ). }
+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.
+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 |}
+ | 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) |}))
+ end.
+ { clear -pf pf'.
+ unfold Tuple.map2, Tuple.map in pf; simpl in *.
+ abstract (
+ destruct bounds;
+ simpl in *;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+ { simpl in *.
+ clear -pf pf'.
+ abstract (
+ destruct bounds as [? [? ?] ], v; simpl in *;
+ rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
+ simpl in pf;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+Defined.
+
+Definition assoc_right''
+ := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
+
+Definition args_to_bounded {n} v bounds pf
+ := Eval cbv [args_to_bounded_helper assoc_right''] in
+ @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
+
+Local Ltac get_len T :=
+ match (eval hnf in T) with
+ | prod ?A ?B
+ => let a := get_len A in
+ let b := get_len B in
+ (eval compute in (a + b)%nat)
+ | _ => constr:(1%nat)
+ end.
+
+Local Ltac args_to_bounded x H :=
+ let x' := fresh in
+ set (x' := x);
+ compute in x;
+ let len := (let T := type of x in get_len T) in
+ destruct_head' prod;
+ let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
+ let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
+ let c := (eval cbv [all_binders_for ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
+ apply c; compute; clear;
+ try abstract (
+ repeat split;
+ solve [ reflexivity
+ | refine (fun v => match v with eq_refl => I end) ]
+ ).
+
+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).
+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).
+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).
+Proof.
+ let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
+ exact v.
+Defined.
+
+Ltac assoc_right_tuple x so_far :=
+ let t := type of x in
+ lazymatch (eval hnf in t) with
+ | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
+ assoc_right_tuple (fst x) so_far
+ | _ => lazymatch so_far with
+ | @None => x
+ | _ => constr:((x, so_far))
+ end
+ end.
+
+Local Ltac make_bounds_prop bounds orig_bounds :=
+ let bounds' := fresh "bounds'" in
+ let bounds_bad := fresh "bounds_bad" in
+ rename bounds into bounds_bad;
+ let boundsv := assoc_right_tuple bounds_bad (@None) in
+ pose boundsv as bounds;
+ pose orig_bounds as bounds';
+ repeat (refine (match fst bounds' with
+ | Some bounds' => let (l, u) := fst bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end && _)%bool;
+ destruct bounds' as [_ bounds'], bounds as [_ bounds]);
+ try exact (match bounds' with
+ | Some bounds' => let (l, u) := bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end).
+
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprBinOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToWireT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpWireToFET)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToZT)) : bool.
+Proof.
+ refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
+Defined.
+
+(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
+ various kinds of correct and boundedness, and abstract in Gallina
+ rather than Ltac *)
+
+Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
+ let Heq := fresh "Heq" in
+ let Hbounds0 := fresh "Hbounds0" in
+ let Hbounds1 := fresh "Hbounds1" in
+ let Hbounds2 := fresh "Hbounds2" in
+ pose proof (proj2_sig ropZ_sig) as Heq;
+ cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
+ curry_binop_fe2519_32W curry_unop_fe2519_32W curry_unop_wire_digitsW
+ curry_binop_fe2519_32 curry_unop_fe2519_32 curry_unop_wire_digits
+ uncurry_binop_fe2519_32W uncurry_unop_fe2519_32W uncurry_unop_wire_digitsW
+ uncurry_binop_fe2519_32 uncurry_unop_fe2519_32 uncurry_unop_wire_digits
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET
+ interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ cbv zeta in *;
+ simpl @fe2519_32WToZ; simpl @wire_digitsWToZ;
+ rewrite <- Heq; clear Heq;
+ destruct Hbounds as [Heq Hbounds];
+ change interp_op with (@Z.interp_op) in *;
+ 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;
+ specialize_by repeat first [ progress intros
+ | reflexivity
+ | assumption
+ | progress destruct_head' base_type
+ | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' and
+ | progress repeat apply conj ];
+ specialize (Hbounds_left args H0);
+ specialize (Hbounds_right args H0);
+ cbv beta in *;
+ lazymatch type of Hbounds_right with
+ | match ?e with _ => _ end
+ => lazymatch type of H1 with
+ | match ?e' with _ => _ end
+ => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
+ end
+ end;
+ 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;
+ match goal with
+ | [ |- fe2519_32WToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- wire_digitsWToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- _ = _ ]
+ => exact Hbounds_left
+ end;
+ change word64ToZ with Word64.word64ToZ in *;
+ (split; [ exact Hbounds_left | ]);
+ cbv [interp_flat_type] in *;
+ cbv [fst snd
+ binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good
+ ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds] in H1;
+ destruct_head' ZBounds.bounds;
+ unfold_is_bounded_in H1;
+ simpl @fe2519_32WToZ; simpl @wire_digitsWToZ;
+ unfold_is_bounded;
+ destruct_head' and;
+ Z.ltb_to_lt;
+ change Word64.word64ToZ with word64ToZ in *;
+ repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
+
+
+Ltac rexpr_correct :=
+ let ropW' := fresh in
+ let ropZ_sig := fresh in
+ intros ropW' ropZ_sig;
+ let wf_ropW := fresh "wf_ropW" in
+ assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
+ cbv zeta; repeat apply conj;
+ [ vm_compute; reflexivity
+ | apply @InterpRelWf;
+ [ | 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 compute_bounds opW bounds
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ (only parsing).
+
+
+Module Export PrettyPrinting.
+ (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
+ 8.5/8.6. Because [Set] is special and things break if
+ [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
+ to debug. *)
+ Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
+
+ Inductive result := yes | no | borked.
+
+ Definition ZBounds_to_bounds_on
+ := 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 |}
+ => in_range l u
+ | None
+ => overflow
+ end
+ end.
+
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result
+ := match t return interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result with
+ | Tbase TZ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
+ | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
+ | no, no => no
+ | yes, no | no, yes | yes, yes => yes
+ | borked, _ | _, borked => borked
+ end
+ end.
+
+ (** This gives a slightly easier to read version of the bounds *)
+ Notation compute_bounds_for_display opW bounds
+ := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
+ Notation sanity_compute opW bounds
+ := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
+ Notation sanity_check opW bounds
+ := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
+End PrettyPrinting.
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v b/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
new file mode 100644
index 000000000..d2bce7c3f
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
@@ -0,0 +1,50 @@
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprBinOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_binop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall xy
+ (xy := (eta_fe2519_32W (fst xy), eta_fe2519_32W (snd xy)))
+ (Hxy : is_bounded (fe2519_32WToZ (fst xy)) = true
+ /\ is_bounded (fe2519_32WToZ (snd xy)) = true),
+ let Hx := let (Hx, Hy) := Hxy in Hx in
+ 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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall xy
+ (xy := (eta_fe2519_32W (fst xy), eta_fe2519_32W (snd xy)))
+ (Hxy : is_bounded (fe2519_32WToZ (fst xy)) = true
+ /\ is_bounded (fe2519_32WToZ (snd xy)) = true),
+ 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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => binop_bounds_good bounds = true
+ | None => False
+ end)
+ : binop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x y Hx Hy.
+ pose x as x'; pose y as y'.
+ hnf in x, y; destruct_head' prod.
+ specialize (H0 (x', y') (conj Hx Hy)).
+ specialize (H1 (x', y') (conj Hx Hy)).
+ let args := constr:(binop_args_to_bounded (x', y') Hx Hy) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
new file mode 100644
index 000000000..d651e6118
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe2519_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unop_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
new file mode 100644
index 000000000..10dc0866e
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToWire_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe2519_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToWire_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToWire_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
new file mode 100644
index 000000000..ec00c6233
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToZ_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe2519_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToZ_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToZ_correct (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
new file mode 100644
index 000000000..6ca61afe9
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpWireToFE_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_wire_digitsW x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopWireToFE_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_WireToFE_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified.v b/src/SpecificGen/GF2519_32Reflective/Reified.v
new file mode 100644
index 000000000..4630e532f
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified.v
@@ -0,0 +1,13 @@
+(** We split the reification up into separate files, one operation per
+ file, so that it can run in parallel. *)
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.Add.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.CarryAdd.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.Sub.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.CarrySub.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.Mul.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.Opp.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.CarryOpp.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.PreFreeze.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.GeModulus.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.Pack.
+Require Export Crypto.SpecificGen.GF2519_32Reflective.Reified.Unpack.
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/Add.v b/src/SpecificGen/GF2519_32Reflective/Reified/Add.v
new file mode 100644
index 000000000..c379c398a
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/Add.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonBinOp.
+
+Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
+Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
+Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
+Proof. rexpr_correct. Qed.
+Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
+Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
+Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/CarryAdd.v b/src/SpecificGen/GF2519_32Reflective/Reified/CarryAdd.v
new file mode 100644
index 000000000..91571ac77
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/CarryAdd.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonBinOp.
+
+Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
+Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
+Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_addW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/CarryOpp.v b/src/SpecificGen/GF2519_32Reflective/Reified/CarryOpp.v
new file mode 100644
index 000000000..708dd54bf
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/CarryOpp.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOp.
+
+Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
+Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
+Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_oppW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/CarrySub.v b/src/SpecificGen/GF2519_32Reflective/Reified/CarrySub.v
new file mode 100644
index 000000000..f69b61455
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/CarrySub.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonBinOp.
+
+Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
+Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
+Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_subW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/GeModulus.v b/src/SpecificGen/GF2519_32Reflective/Reified/GeModulus.v
new file mode 100644
index 000000000..e5dfa9e43
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/GeModulus.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOpFEToZ.
+
+Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
+Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
+Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rge_modulusW_correct_and_bounded
+ := ExprUnOpFEToZ_correct_and_bounded
+ rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/Mul.v b/src/SpecificGen/GF2519_32Reflective/Reified/Mul.v
new file mode 100644
index 000000000..dc1e0bf49
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/Mul.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonBinOp.
+
+Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
+Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
+Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rmulW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
+Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
+Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/Opp.v b/src/SpecificGen/GF2519_32Reflective/Reified/Opp.v
new file mode 100644
index 000000000..4d38eb057
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/Opp.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOp.
+
+Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
+Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
+Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
+Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
+Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/Pack.v b/src/SpecificGen/GF2519_32Reflective/Reified/Pack.v
new file mode 100644
index 000000000..f82b07842
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/Pack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOpFEToWire.
+
+Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
+Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
+Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rpackW_correct_and_bounded
+ := ExprUnOpFEToWire_correct_and_bounded
+ rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v
new file mode 100644
index 000000000..dc9e2dcef
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOp.
+
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF2519_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
+Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rprefreezeW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/Sub.v b/src/SpecificGen/GF2519_32Reflective/Reified/Sub.v
new file mode 100644
index 000000000..1875487fb
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/Sub.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonBinOp.
+
+Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
+Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
+Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
+Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
+Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/Unpack.v b/src/SpecificGen/GF2519_32Reflective/Reified/Unpack.v
new file mode 100644
index 000000000..f839f4f0c
--- /dev/null
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/Unpack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOpWireToFE.
+
+Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
+Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
+Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition runpackW_correct_and_bounded
+ := ExprUnOpWireToFE_correct_and_bounded
+ runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/SpecificGen/GF25519_32Bounded.v b/src/SpecificGen/GF25519_32Bounded.v
new file mode 100644
index 000000000..311001606
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Bounded.v
@@ -0,0 +1,462 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF25519_32.
+Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
+Require Import Crypto.SpecificGen.GF25519_32Reflective.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+
+Local Ltac bounded_t opW blem :=
+ apply blem; apply is_bounded_proj1_fe25519_32.
+Local Ltac bounded_wire_digits_t opW blem :=
+ apply blem; apply is_bounded_proj1_wire_digits.
+
+Local Ltac define_binop f g opW blem :=
+ refine (exist_fe25519_32W (opW (proj1_fe25519_32W f) (proj1_fe25519_32W g)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop f opW blem :=
+ refine (exist_fe25519_32W (opW (proj1_fe25519_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_FEToZ f opW :=
+ refine (opW (proj1_fe25519_32W f)).
+Local Ltac define_unop_FEToWire f opW blem :=
+ refine (exist_wire_digitsW (opW (proj1_fe25519_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_WireToFE f opW blem :=
+ refine (exist_fe25519_32W (opW (proj1_wire_digitsW f)) _);
+ 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 Arguments interp_radd / _ _.
+Local Arguments interp_rsub / _ _.
+Local Arguments interp_rmul / _ _.
+Local Arguments interp_ropp / _.
+Local Arguments interp_rprefreeze / _.
+Local Arguments interp_rge_modulus / _.
+Local Arguments interp_rpack / _.
+Local Arguments interp_runpack / _.
+Definition addW (f g : fe25519_32W) : fe25519_32W := Eval simpl in interp_radd f g.
+Definition subW (f g : fe25519_32W) : fe25519_32W := Eval simpl in interp_rsub f g.
+Definition mulW (f g : fe25519_32W) : fe25519_32W := Eval simpl in interp_rmul f g.
+Definition oppW (f : fe25519_32W) : fe25519_32W := Eval simpl in interp_ropp f.
+Definition prefreezeW (f : fe25519_32W) : fe25519_32W := Eval simpl in interp_rprefreeze f.
+Definition ge_modulusW (f : fe25519_32W) : word64 := Eval simpl in interp_rge_modulus f.
+Definition packW (f : fe25519_32W) : wire_digitsW := Eval simpl in interp_rpack f.
+Definition unpackW (f : wire_digitsW) : fe25519_32W := Eval simpl in interp_runpack f.
+
+Require Import ModularBaseSystemWord.
+Definition modulusW :=
+ Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z 10 GF25519_32.modulus_digits_)).
+
+Definition postfreeze : GF25519_32.fe25519_32 -> GF25519_32.fe25519_32 :=
+ GF25519_32.postfreeze.
+
+Lemma freeze_prepost_freeze : forall x, postfreeze (prefreeze x) = GF25519_32.freeze x. Proof. reflexivity. Qed.
+
+Definition postfreezeW : fe25519_32W -> fe25519_32W :=
+ (conditional_subtract_modulusW
+ (num_limbs := 10)
+ modulusW
+ ge_modulusW
+ (Interpretations.Word64.neg GF25519_32.int_width)
+ ).
+
+Definition freezeW (f : fe25519_32W) : fe25519_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
+
+Local Transparent Let_In.
+Definition powW (f : fe25519_32W) chain := fold_chain_opt (proj1_fe25519_32W one) mulW chain [f].
+Definition invW (f : fe25519_32W) : fe25519_32W
+ := Eval cbv -[Let_In fe25519_32W mulW] in powW f (chain inv_ec).
+
+Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
+ change opW with (interp_rop);
+ rewrite pre_rewrite;
+ intros; apply rop_cb; assumption.
+
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Proof. port_correct_and_bounded interp_radd_correct addW interp_radd radd_correct_and_bounded. Qed.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW carry_sub.
+Proof. port_correct_and_bounded interp_rsub_correct subW interp_rsub rsub_correct_and_bounded. Qed.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW mul.
+Proof. port_correct_and_bounded interp_rmul_correct mulW interp_rmul rmul_correct_and_bounded. Qed.
+Lemma oppW_correct_and_bounded : iunop_correct_and_bounded oppW carry_opp.
+Proof. port_correct_and_bounded interp_ropp_correct oppW interp_ropp ropp_correct_and_bounded. Qed.
+Lemma prefreezeW_correct_and_bounded : iunop_correct_and_bounded prefreezeW prefreeze.
+Proof. port_correct_and_bounded interp_rprefreeze_correct prefreezeW interp_rprefreeze rprefreeze_correct_and_bounded. Qed.
+Lemma ge_modulusW_correct : iunop_FEToZ_correct ge_modulusW ge_modulus.
+Proof. port_correct_and_bounded interp_rge_modulus_correct ge_modulusW interp_rge_modulus rge_modulus_correct_and_bounded. Qed.
+Lemma packW_correct_and_bounded : iunop_FEToWire_correct_and_bounded packW pack.
+Proof. port_correct_and_bounded interp_rpack_correct packW interp_rpack rpack_correct_and_bounded. Qed.
+Lemma unpackW_correct_and_bounded : iunop_WireToFE_correct_and_bounded unpackW unpack.
+Proof. port_correct_and_bounded interp_runpack_correct unpackW interp_runpack runpack_correct_and_bounded. Qed.
+
+(* TODO : move *)
+Lemma neg_range : forall x y, 0 <= x ->
+ 0 <= ModularBaseSystemListZOperations.neg x y < 2 ^ x.
+Proof.
+ intros.
+ split; auto using ModularBaseSystemListZOperationsProofs.neg_nonneg.
+ eapply Z.le_lt_trans; eauto using ModularBaseSystemListZOperationsProofs.neg_upperbound.
+ rewrite Z.ones_equiv.
+ omega.
+Qed.
+
+Ltac lower_bound_minus_ge_modulus :=
+ apply Z.le_0_sub;
+ 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;
+ etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
+ apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
+
+Ltac upper_bound_minus_ge_modulus :=
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | ];
+ eapply Z.le_lt_trans; [ eassumption | ];
+ instantiate; vm_compute; reflexivity.
+
+Lemma postfreezeW_correct_and_bounded : iunop_correct_and_bounded postfreezeW postfreeze.
+Proof.
+ intros x H.
+ pose proof (ge_modulusW_correct x H) as Hgm.
+ destruct_head_hnf' prod.
+ unfold_is_bounded_in H.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ cbv [postfreezeW].
+ cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
+ change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ rewrite Hgm.
+
+ cbv [modulusW Tuple.map].
+ cbv [on_tuple List.map to_list to_list' from_list from_list'
+ Tuple.map2 on_tuple2 ListUtil.map2 fe25519_32WToZ].
+ cbv [postfreeze GF25519_32.postfreeze].
+ cbv [Let_In].
+
+ split.
+ { match goal with
+ |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?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;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ | |- (_,_) = (_,_) => reflexivity
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ end. }
+
+
+ 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;
+ repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ try match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ | |- _ - _ <= _ => etransitivity; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | instantiate; assumption ]
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- ModularBaseSystemListZOperations.neg _ _ < _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- _ => vm_compute; (discriminate || reflexivity)
+ end.
+Qed.
+
+Lemma freezeW_correct_and_bounded : iunop_correct_and_bounded freezeW freeze.
+Proof.
+ intros f H; rewrite <- freeze_prepost_freeze.
+ change (freezeW f) with (postfreezeW (prefreezeW f)).
+ destruct (prefreezeW_correct_and_bounded f H) as [H0 H1].
+ destruct (postfreezeW_correct_and_bounded _ H1) as [H0' H1'].
+ rewrite H1', H0', H0; split; reflexivity.
+Qed.
+
+Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
+Proof.
+ cbv [powW pow].
+ intro x; intros; apply (fold_chain_opt_gen fe25519_32WToZ is_bounded [x]).
+ { reflexivity. }
+ { reflexivity. }
+ { intros; progress rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption.
+ apply mulW_correct_and_bounded; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption; reflexivity. }
+ { intros [|?]; autorewrite with simpl_nth_default;
+ (assumption || reflexivity). }
+Qed.
+
+Lemma invW_correct_and_bounded : iunop_correct_and_bounded invW inv.
+Proof.
+ intro f.
+ assert (H : forall f, invW f = powW f (chain inv_ec))
+ by abstract (cbv -[Let_In fe25519_32W mulW]; reflexivity).
+ rewrite H.
+ rewrite inv_correct.
+ cbv [inv_opt].
+ rewrite <- pow_correct.
+ apply powW_correct_and_bounded.
+Qed.
+
+Definition fieldwisebW_sig (f g : fe25519_32W)
+ : { b | b = GF25519_32.fieldwiseb (fe25519_32WToZ f) (fe25519_32WToZ g) }.
+Proof.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ cbv [GF25519_32.fieldwiseb fe25519_32WToZ].
+ rewrite ?word64eqb_Zeqb.
+ reflexivity.
+Defined.
+
+Definition fieldwisebW (f g : fe25519_32W) : bool :=
+ Eval cbv [proj1_sig fieldwisebW_sig appify2 app_fe25519_32W] in
+ appify2 (fun f g => proj1_sig (fieldwisebW_sig f g)) f g.
+
+Lemma fieldwisebW_correct f g
+ : fieldwisebW f g = GF25519_32.fieldwiseb (fe25519_32WToZ f) (fe25519_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (fieldwisebW_sig f' g')).
+Qed.
+
+Local Arguments freezeW : simpl never.
+Local Arguments fe25519_32WToZ !_ / .
+Local Opaque freezeW.
+
+Definition eqbW_sig (f g : fe25519_32W)
+ : { b | is_bounded (fe25519_32WToZ f) = true
+ -> is_bounded (fe25519_32WToZ g) = true
+ -> b = GF25519_32.eqb (fe25519_32WToZ f) (fe25519_32WToZ g) }.
+Proof.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded f pf)) as frf.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded g pf)) as frg.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ unfold GF25519_32.eqb.
+ simpl @fe25519_32WToZ in *; cbv beta iota.
+ intros.
+ rewrite <- frf, <- frg by assumption.
+ rewrite <- fieldwisebW_correct.
+ reflexivity.
+Defined.
+
+Definition eqbW (f g : fe25519_32W) : bool :=
+ Eval cbv [proj1_sig eqbW_sig appify2 app_fe25519_32W] in
+ appify2 (fun f g => proj1_sig (eqbW_sig f g)) f g.
+
+Lemma eqbW_correct f g
+ : is_bounded (fe25519_32WToZ f) = true
+ -> is_bounded (fe25519_32WToZ g) = true
+ -> eqbW f g = GF25519_32.eqb (fe25519_32WToZ f) (fe25519_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (eqbW_sig f' g')).
+Qed.
+
+(* TODO(jgross): use NToWord or such for this constant too *)
+Definition sqrt_m1W : fe25519_32W :=
+ Eval vm_compute in fe25519_32ZToW sqrt_m1.
+
+Definition GF25519_32sqrt x : GF25519_32.fe25519_32 :=
+ dlet powx := powW (fe25519_32ZToW x) (chain GF25519_32.sqrt_ec) in
+ GF25519_32.sqrt (fe25519_32WToZ powx) (fe25519_32WToZ (mulW powx powx)) x.
+
+Definition sqrtW_sig
+ : { sqrtW | iunop_correct_and_bounded sqrtW GF25519_32sqrt }.
+Proof.
+ eexists.
+ unfold GF25519_32sqrt, GF25519_32.sqrt.
+ intros.
+ rewrite !fe25519_32ZToW_WToZ.
+ split.
+ { etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody_eq; intros.
+ set_evars.
+ match goal with (* unfold the first dlet ... in, but only if it's binding a var *)
+ | [ |- ?x = dlet y := fe25519_32WToZ ?z in ?f ]
+ => is_var z; change (x = match fe25519_32WToZ z with y => f end)
+ end.
+ change sqrt_m1 with (fe25519_32WToZ sqrt_m1W).
+ rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded sqrt_m1W a X Y)), <- eqbW_correct, (pull_bool_if fe25519_32WToZ)
+ by repeat match goal with
+ | _ => progress subst
+ | [ |- is_bounded (fe25519_32WToZ ?op) = true ]
+ => lazymatch op with
+ | mulW _ _ => apply mulW_correct_and_bounded
+ | powW _ _ => apply powW_correct_and_bounded
+ | sqrt_m1W => vm_compute; reflexivity
+ | _ => assumption
+ end
+ end.
+ subst_evars; reflexivity.
+ } Unfocus.
+ lazymatch goal with
+ | [ |- context G[dlet x := ?v in fe25519_32WToZ (@?f x)] ]
+ => let G' := context G[fe25519_32WToZ (dlet x := v in f x)] in
+ cut G'; cbv beta;
+ [ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
+ end.
+ reflexivity. }
+ { cbv [Let_In].
+ break_if.
+ { apply powW_correct_and_bounded; assumption. }
+ { apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ].
+ apply powW_correct_and_bounded; assumption. } }
+Defined.
+
+Definition sqrtW (f : fe25519_32W) : fe25519_32W :=
+ Eval cbv [proj1_sig sqrtW_sig app_fe25519_32W] in
+ app_fe25519_32W f (proj1_sig sqrtW_sig).
+
+Lemma sqrtW_correct_and_bounded : iunop_correct_and_bounded sqrtW GF25519_32sqrt.
+Proof.
+ intro f.
+ set (f' := f).
+ hnf in f; destruct_head' prod.
+ assert (H : sqrtW f' = proj1_sig sqrtW_sig f')
+ by (subst f'; cbv beta iota delta [proj1_sig sqrtW_sig sqrtW]; reflexivity).
+ rewrite H.
+ exact (proj2_sig sqrtW_sig f').
+Qed.
+
+
+
+Definition add (f g : fe25519_32) : fe25519_32.
+Proof. define_binop f g addW addW_correct_and_bounded. Defined.
+Definition sub (f g : fe25519_32) : fe25519_32.
+Proof. define_binop f g subW subW_correct_and_bounded. Defined.
+Definition mul (f g : fe25519_32) : fe25519_32.
+Proof. define_binop f g mulW mulW_correct_and_bounded. Defined.
+Definition opp (f : fe25519_32) : fe25519_32.
+Proof. define_unop f oppW oppW_correct_and_bounded. Defined.
+Definition freeze (f : fe25519_32) : fe25519_32.
+Proof. define_unop f freezeW freezeW_correct_and_bounded. Defined.
+Definition ge_modulus (f : fe25519_32) : word64.
+Proof. define_unop_FEToZ f ge_modulusW. Defined.
+Definition pack (f : fe25519_32) : wire_digits.
+Proof. define_unop_FEToWire f packW packW_correct_and_bounded. Defined.
+Definition unpack (f : wire_digits) : fe25519_32.
+Proof. define_unop_WireToFE f unpackW unpackW_correct_and_bounded. Defined.
+
+Definition pow (f : fe25519_32) (chain : list (nat * nat)) : fe25519_32.
+Proof. define_unop f (fun x => powW x chain) powW_correct_and_bounded. Defined.
+Definition inv (f : fe25519_32) : fe25519_32.
+Proof. define_unop f invW invW_correct_and_bounded. Defined.
+Definition sqrt (f : fe25519_32) : fe25519_32.
+Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
+
+Local Ltac op_correct_t op opW_correct_and_bounded :=
+ cbv [op];
+ lazymatch goal with
+ | [ |- context[proj1_fe25519_32 (exist_fe25519_32W _ _)] ]
+ => rewrite proj1_fe25519_32_exist_fe25519_32W
+ | [ |- context[proj1_wire_digits (exist_wire_digitsW _ _)] ]
+ => rewrite proj1_wire_digits_exist_wire_digitsW
+ | _ => idtac
+ end;
+ apply opW_correct_and_bounded;
+ lazymatch goal with
+ | [ |- is_bounded _ = true ]
+ => apply is_bounded_proj1_fe25519_32
+ | [ |- wire_digits_is_bounded _ = true ]
+ => apply is_bounded_proj1_wire_digits
+ end.
+
+Lemma add_correct (f g : fe25519_32) : proj1_fe25519_32 (add f g) = carry_add (proj1_fe25519_32 f) (proj1_fe25519_32 g).
+Proof. op_correct_t add addW_correct_and_bounded. Qed.
+Lemma sub_correct (f g : fe25519_32) : proj1_fe25519_32 (sub f g) = carry_sub (proj1_fe25519_32 f) (proj1_fe25519_32 g).
+Proof. op_correct_t sub subW_correct_and_bounded. Qed.
+Lemma mul_correct (f g : fe25519_32) : proj1_fe25519_32 (mul f g) = GF25519_32.mul (proj1_fe25519_32 f) (proj1_fe25519_32 g).
+Proof. op_correct_t mul mulW_correct_and_bounded. Qed.
+Lemma opp_correct (f : fe25519_32) : proj1_fe25519_32 (opp f) = carry_opp (proj1_fe25519_32 f).
+Proof. op_correct_t opp oppW_correct_and_bounded. Qed.
+Lemma freeze_correct (f : fe25519_32) : proj1_fe25519_32 (freeze f) = GF25519_32.freeze (proj1_fe25519_32 f).
+Proof. op_correct_t freeze freezeW_correct_and_bounded. Qed.
+Lemma ge_modulus_correct (f : fe25519_32) : word64ToZ (ge_modulus f) = GF25519_32.ge_modulus (proj1_fe25519_32 f).
+Proof. op_correct_t ge_modulus ge_modulusW_correct. Qed.
+Lemma pack_correct (f : fe25519_32) : proj1_wire_digits (pack f) = GF25519_32.pack (proj1_fe25519_32 f).
+Proof. op_correct_t pack packW_correct_and_bounded. Qed.
+Lemma unpack_correct (f : wire_digits) : proj1_fe25519_32 (unpack f) = GF25519_32.unpack (proj1_wire_digits f).
+Proof. op_correct_t unpack unpackW_correct_and_bounded. Qed.
+Lemma pow_correct (f : fe25519_32) chain : proj1_fe25519_32 (pow f chain) = GF25519_32.pow (proj1_fe25519_32 f) chain.
+Proof. op_correct_t pow (powW_correct_and_bounded chain). Qed.
+Lemma inv_correct (f : fe25519_32) : proj1_fe25519_32 (inv f) = GF25519_32.inv (proj1_fe25519_32 f).
+Proof. op_correct_t inv invW_correct_and_bounded. Qed.
+Lemma sqrt_correct (f : fe25519_32) : proj1_fe25519_32 (sqrt f) = GF25519_32sqrt (proj1_fe25519_32 f).
+Proof. op_correct_t sqrt sqrtW_correct_and_bounded. Qed.
+
+Import Morphisms.
+
+Local Existing Instance prime_modulus.
+
+Lemma field25519_32_and_homomorphisms
+ : @field fe25519_32 eq zero one opp add sub mul inv div
+ /\ @Ring.is_homomorphism (F _) (@Logic.eq _) 1%F F.add F.mul fe25519_32 eq one add mul encode
+ /\ @Ring.is_homomorphism fe25519_32 eq one add mul (F _) (@Logic.eq _) 1%F F.add F.mul decode.
+Proof.
+ eapply @Field.field_and_homomorphism_from_redundant_representation.
+ { exact (F.field_modulo _). }
+ { cbv [decode encode]; intros; rewrite !proj1_fe25519_32_exist_fe25519_32; apply encode_rep. }
+ { reflexivity. }
+ { reflexivity. }
+ { reflexivity. }
+ { cbv [decode encode]; intros; rewrite opp_correct, carry_opp_correct, carry_opp_opt_correct, carry_opp_rep; apply opp_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite add_correct, carry_add_correct, carry_add_opt_correct, carry_add_rep; apply add_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite sub_correct, carry_sub_correct, carry_sub_opt_correct, carry_sub_rep; apply sub_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite mul_correct, GF25519_32.mul_correct, carry_mul_opt_correct by reflexivity; apply carry_mul_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite inv_correct, GF25519_32.inv_correct, inv_opt_correct by reflexivity; apply inv_rep; reflexivity. }
+ { cbv [decode encode div]; intros; rewrite !proj1_fe25519_32_exist_fe25519_32; apply encode_rep. }
+Qed.
+
+Global Instance field25519_32 : @field fe25519_32 eq zero one opp add sub mul inv div := proj1 field25519_32_and_homomorphisms.
+
+Local Opaque proj1_fe25519_32 exist_fe25519_32 proj1_fe25519_32W exist_fe25519_32W.
+Global Instance homomorphism_F25519_32_encode
+ : @Ring.is_homomorphism (F modulus) Logic.eq F.one F.add F.mul fe25519_32 eq one add mul encode.
+Proof. apply field25519_32_and_homomorphisms. Qed.
+
+Global Instance homomorphism_F25519_32_decode
+ : @Ring.is_homomorphism fe25519_32 eq one add mul (F modulus) Logic.eq F.one F.add F.mul decode.
+Proof. apply field25519_32_and_homomorphisms. Qed.
diff --git a/src/SpecificGen/GF25519_32BoundedCommon.v b/src/SpecificGen/GF25519_32BoundedCommon.v
new file mode 100644
index 000000000..f88b8a152
--- /dev/null
+++ b/src/SpecificGen/GF25519_32BoundedCommon.v
@@ -0,0 +1,693 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF25519_32.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.HList.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+(* BEGIN common curve-specific definitions *)
+Definition bit_width : nat := Eval compute in Z.to_nat (GF25519_32.int_width).
+Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
+Definition bounds_exp : tuple Z length_fe25519_32
+ := Eval compute in
+ Tuple.from_list length_fe25519_32 limb_widths eq_refl.
+Definition bounds : tuple (Z * Z) length_fe25519_32
+ := Eval compute in
+ Tuple.map (fun e => b_of e) bounds_exp.
+Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
+ := Eval compute in Tuple.from_list _ wire_widths eq_refl.
+Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
+ := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
+(* 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.
+Arguments bit_width : simpl never.
+Global Opaque word64.
+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},
+ 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},
+ 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
+ := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
+ Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ : 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}
+ : 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 }.
+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_32W := Eval cbv (*-[word64]*) in (tuple word64 length_fe25519_32).
+Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
+Definition fe25519_32 :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => word_of e) bounds_exp.
+Definition wire_digits :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
+
+Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
+ := let res := Tuple.map2
+ (fun bounds v =>
+ let '(lower, upper) := bounds in
+ (lower <=? v) && (v <=? upper))%bool%Z
+ bounds x in
+ List.fold_right andb true (Tuple.to_list _ res).
+
+Definition is_bounded (x : SpecificGen.GF25519_32.fe25519_32) : bool
+ := is_bounded_gen (n:=length_fe25519_32) x bounds.
+
+Definition wire_digits_is_bounded (x : SpecificGen.GF25519_32.wire_digits) : bool
+ := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
+
+(* END basic types *)
+
+Section generic_destructuring.
+ Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
+ := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
+ end.
+ Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
+ := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => @app_on' A n'
+ end.
+ Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
+ Qed.
+ Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
+
+ Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
+ end.
+ Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => @app_on_h' A F n'
+ end.
+ Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
+ Qed.
+ Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
+
+ Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
+ : forall (f : tuple A (length wire_widths)), T f
+ := Eval compute in fun f => @app_on A (length wire_widths) T f P.
+ Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
+ := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
+ Definition app_fe25519_32W_dep {A T} (P : forall x : tuple A length_fe25519_32, T x)
+ : forall (f : tuple A length_fe25519_32), T f
+ := Eval compute in fun f => @app_on A length_fe25519_32 T f P.
+ Definition app_fe25519_32W {A T} (f : tuple A length_fe25519_32) (P : tuple A length_fe25519_32 -> T)
+ := Eval compute in @app_fe25519_32W_dep A (fun _ => T) P f.
+ Definition app_fe25519_32_dep {T} (P : forall x : fe25519_32, T x)
+ : forall f : fe25519_32, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe25519_32 bounds_exp T f P.
+ Definition app_fe25519_32 {T} (f : fe25519_32) (P : hlist (fun e => word_of e) bounds_exp -> T)
+ := Eval compute in @app_fe25519_32_dep (fun _ => T) P f.
+ Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
+ : forall f : wire_digits, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
+ Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
+ := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
+
+ Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
+ := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
+ Definition app_fe25519_32W_dep_correct {A T} f P : @app_fe25519_32W_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_fe25519_32W_correct {A T} f P : @app_fe25519_32W A T f P = P f
+ := @app_fe25519_32W_dep_correct A (fun _ => T) f P.
+ Definition app_fe25519_32_dep_correct {T} f P : @app_fe25519_32_dep T P f = P f
+ := app_on_h_correct (fun e => word_of e) bounds_exp f P.
+ Definition app_fe25519_32_correct {T} f P : @app_fe25519_32 T f P = P f
+ := @app_fe25519_32_dep_correct (fun _ => T) f P.
+ Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
+ := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
+ Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
+ := @app_wire_digits_dep_correct (fun _ => T) f P.
+
+ Definition appify2 {T} (op : fe25519_32W -> fe25519_32W -> T) (f g : fe25519_32W) :=
+ app_fe25519_32W f (fun f0 => (app_fe25519_32W g (fun g0 => op f0 g0))).
+
+ Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
+ Proof.
+ intros. cbv [appify2].
+ etransitivity; apply app_fe25519_32W_correct.
+ Qed.
+End generic_destructuring.
+
+Definition eta_fe25519_32W_sig (x : fe25519_32W) : { v : fe25519_32W | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+Defined.
+Definition eta_fe25519_32W (x : fe25519_32W) : fe25519_32W
+ := Eval cbv [proj1_sig eta_fe25519_32W_sig] in proj1_sig (eta_fe25519_32W_sig x).
+Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+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.
+Proof. apply NToWord_wordToN. Qed.
+Local Opaque word64.
+
+Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ rewrite N2Z.id, NToWord_wordToN; reflexivity.
+Qed.
+
+(* BEGIN precomputation. *)
+
+Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
+Definition word_bounded {lower upper} (v : bounded_word lower upper)
+ : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
+ := project_is_boundedT (proj2_sig v).
+Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
+ := exist _ proj_word (Build_is_boundedT word_bounded).
+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)
+ : bounded_word lower upper
+ := Build_bounded_word'
+ proj_word
+ (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
+ | 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.
+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.
+ match goal with
+ | [ |- context[andb ?x ?y] ]
+ => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
+ end;
+ intros; omega.
+Qed.
+Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
+Proof.
+ refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
+ abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
+Defined.
+Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
+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.
+Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe25519_32W app_fe25519_32 HList.mapt HList.mapt' Tuple.map on_tuple List.map List.app length_fe25519_32 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
+Definition fe25519_32WToZ (x : fe25519_32W) : SpecificGen.GF25519_32.fe25519_32
+ := Eval app_tuple_map in
+ app_fe25519_32W x (Tuple.map (fun v : word64 => v : Z)).
+Definition fe25519_32ZToW (x : SpecificGen.GF25519_32.fe25519_32) : fe25519_32W
+ := Eval app_tuple_map in
+ app_fe25519_32W x (Tuple.map (fun v : Z => v : word64)).
+Definition wire_digitsWToZ (x : wire_digitsW) : SpecificGen.GF25519_32.wire_digits
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
+Definition wire_digitsZToW (x : SpecificGen.GF25519_32.wire_digits) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
+Definition fe25519_32W_word64ize (x : fe25519_32W) : fe25519_32W
+ := Eval app_tuple_map in
+ app_fe25519_32W x (Tuple.map word64ize).
+Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map word64ize).
+
+(** TODO: Turn this into a lemma to speed up proofs *)
+Ltac unfold_is_bounded_in H :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe25519_32WToZ, wire_digitsWToZ in H;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe25519_32 List.length wire_widths] in H;
+ rewrite ?Bool.andb_true_iff in H.
+
+Ltac unfold_is_bounded :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe25519_32WToZ, wire_digitsWToZ;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe25519_32 List.length wire_widths];
+ 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.
+Local Opaque bit_width.
+
+Local Ltac prove_lt_bit_width :=
+ rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
+
+Lemma fe25519_32ZToW_WToZ (x : fe25519_32W) : fe25519_32ZToW (fe25519_32WToZ x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe25519_32WToZ fe25519_32ZToW].
+ rewrite !ZToWord64_word64ToZ; reflexivity.
+Qed.
+
+Lemma fe25519_32WToZ_ZToW x : is_bounded x = true -> fe25519_32WToZ (fe25519_32ZToW x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe25519_32WToZ fe25519_32ZToW].
+ intro H.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma fe25519_32W_word64ize_id x : fe25519_32W_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [fe25519_32W_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [wire_digitsW_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+
+Definition uncurry_unop_fe25519_32W {T} (op : fe25519_32W -> T)
+ := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe25519_32) op.
+Definition curry_unop_fe25519_32W {T} op : fe25519_32W -> T
+ := Eval cbv (*-[word64]*) in fun f => app_fe25519_32W f (Tuple.curry (n:=length_fe25519_32) op).
+Definition uncurry_binop_fe25519_32W {T} (op : fe25519_32W -> fe25519_32W -> T)
+ := Eval cbv (*-[word64]*) in uncurry_unop_fe25519_32W (fun f => uncurry_unop_fe25519_32W (op f)).
+Definition curry_binop_fe25519_32W {T} op : fe25519_32W -> fe25519_32W -> T
+ := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe25519_32W (curry_unop_fe25519_32W op f)).
+
+Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
+ := Eval cbv (*-[word64]*) 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).
+
+
+Definition proj1_fe25519_32W (x : fe25519_32) : fe25519_32W
+ := Eval app_tuple_map in
+ app_fe25519_32 x (HList.mapt (fun _ => (@proj_word _ _))).
+Coercion proj1_fe25519_32 (x : fe25519_32) : SpecificGen.GF25519_32.fe25519_32
+ := fe25519_32WToZ (proj1_fe25519_32W x).
+
+Lemma is_bounded_proj1_fe25519_32 (x : fe25519_32) : is_bounded (proj1_fe25519_32 x) = true.
+Proof.
+ revert x; refine (app_fe25519_32_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [is_bounded proj1_fe25519_32 proj1_fe25519_32W fe25519_32WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe25519_32 is_bounded_gen].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
+ := app_wire_digits x (HList.mapt (fun _ => proj_word)).
+Coercion proj1_wire_digits (x : wire_digits) : SpecificGen.GF25519_32.wire_digits
+ := wire_digitsWToZ (proj1_wire_digitsW x).
+
+Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
+Proof.
+ revert x; refine (app_wire_digits_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Local Ltac make_exist_W' x app_W_dep :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x);
+ cbv [tuple tuple' length_fe25519_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v H :=
+ first [ let v' := (eval cbv [snd fst] in (snd v)) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
+ | let v' := (eval cbv [snd fst] in v) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
+ let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
+ let T := type of H' in
+ let T := (eval cbv [id
+ List.fold_right List.map List.length List.app ListUtil.map2 List.rev
+ Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
+ fe25519_32 bounds fe25519_32WToZ length_fe25519_32
+ wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
+ let H' := constr:(H' : T) in
+ let v := (eval unfold x' in x') in
+ do_refine v H'.
+Local Ltac make_exist'' x exist_W ZToW :=
+ let H := fresh in
+ intro H; apply (exist_W (ZToW x));
+ abstract (
+ hnf in x; destruct_head' prod;
+ let H' := fresh in
+ pose proof H as H';
+ unfold_is_bounded_in H;
+ destruct_head' and; simpl in *;
+ Z.ltb_to_lt;
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ assumption
+ ).
+Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x) in *;
+ cbv [tuple tuple' length_fe25519_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v :=
+ first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
+ | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
+ let v := (eval unfold x' in (exist'' x' H)) in
+ do_refine v.
+
+Definition exist_fe25519_32W' (x : fe25519_32W) : is_bounded (fe25519_32WToZ x) = true -> fe25519_32.
+Proof. make_exist_W' x (@app_fe25519_32W_dep). Defined.
+Definition exist_fe25519_32W (x : fe25519_32W) : is_bounded (fe25519_32WToZ x) = true -> fe25519_32
+ := Eval cbv [app_fe25519_32W_dep exist_fe25519_32W' fe25519_32ZToW] in exist_fe25519_32W' x.
+Definition exist_fe25519_32'' (x : SpecificGen.GF25519_32.fe25519_32) : is_bounded x = true -> fe25519_32.
+Proof. make_exist'' x exist_fe25519_32W fe25519_32ZToW. Defined.
+Definition exist_fe25519_32' (x : SpecificGen.GF25519_32.fe25519_32) : is_bounded x = true -> fe25519_32.
+Proof. make_exist' x (@app_fe25519_32W_dep) exist_fe25519_32'' exist_fe25519_32W fe25519_32ZToW. Defined.
+Definition exist_fe25519_32 (x : SpecificGen.GF25519_32.fe25519_32) : is_bounded x = true -> fe25519_32
+ := Eval cbv [exist_fe25519_32' exist_fe25519_32W exist_fe25519_32' app_fe25519_32 app_fe25519_32W_dep] in
+ exist_fe25519_32' x.
+
+Lemma proj1_fe25519_32_exist_fe25519_32W x pf : proj1_fe25519_32 (exist_fe25519_32W x pf) = fe25519_32WToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe25519_32W_exist_fe25519_32 x pf : proj1_fe25519_32W (exist_fe25519_32 x pf) = fe25519_32ZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe25519_32_exist_fe25519_32 x pf : proj1_fe25519_32 (exist_fe25519_32 x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_fe25519_32 exist_fe25519_32 proj1_fe25519_32W fe25519_32WToZ proj_word Build_bounded_word Build_bounded_word'].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition exist_wire_digitsW' (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
+Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
+Definition exist_wire_digitsW (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
+ := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
+Definition exist_wire_digits'' (x : SpecificGen.GF25519_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits' (x : SpecificGen.GF25519_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits (x : SpecificGen.GF25519_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits
+ := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
+ exist_wire_digits' x.
+
+Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 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 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.
+End opt.
+
+Local Transparent bit_width.
+Local Ltac do_change lem :=
+ match lem with
+ | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
+ => 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
+ let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
+ do_change L'
+ end
+ | _ => lem
+ end.
+Definition fe25519_32_word64ize (x : fe25519_32) : fe25519_32.
+Proof.
+ set (x' := x).
+ hnf in x; destruct_head' prod.
+ let lem := constr:(exist_fe25519_32W (fe25519_32W_word64ize (proj1_fe25519_32W x'))) in
+ let lem := (eval cbv [proj1_fe25519_32W x' fe25519_32W_word64ize proj_word exist_fe25519_32W Build_bounded_word' Build_bounded_word] in lem) in
+ let lem := do_change lem in
+ refine (lem _);
+ change (is_bounded (fe25519_32WToZ (fe25519_32W_word64ize (proj1_fe25519_32W x'))) = true);
+ abstract (rewrite fe25519_32W_word64ize_id; apply is_bounded_proj1_fe25519_32).
+Defined.
+Definition wire_digits_word64ize (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 := do_change lem in
+ let lem := (eval cbv [word64ize opt.word64ize 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).
+Defined.
+
+Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
+ : forall n : nat,
+ (n < length limb_widths)%nat
+ -> (0 <= nth_default 0 (Tuple.to_list length_fe25519_32 x) n <=
+ snd (b_of (nth_default (-1) limb_widths n)))%Z.
+Proof.
+ hnf in x; destruct_head' prod.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ unfold nth_default; simpl.
+ intros.
+ repeat match goal with
+ | [ |- context[nth_error _ ?x] ]
+ => is_var x; destruct x; simpl
+ end;
+ omega.
+Qed.
+
+(* END precomputation *)
+
+(* Precompute constants *)
+
+Definition one' := Eval vm_compute in exist_fe25519_32 SpecificGen.GF25519_32.one_ eq_refl.
+Definition one := Eval cbv [one' fe25519_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_32_word64ize one'.
+
+Definition zero' := Eval vm_compute in exist_fe25519_32 SpecificGen.GF25519_32.zero_ eq_refl.
+Definition zero := Eval cbv [zero' fe25519_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_32_word64ize 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)
+ (Hid : id = F id')
+ (Hop_bounded : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> is_bounded (op (F x) (F y)) = true)
+ (Hop : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> op (F x) (F y) = F (op' x y))
+ (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
+ : F (fold_chain_opt id' op' chain ls)
+ = fold_chain_opt id op chain (List.map F ls)
+ /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
+Proof.
+ rewrite !fold_chain_opt_correct.
+ revert dependent ls; induction chain as [|x xs IHxs]; intros.
+ { pose proof (Hls_bounded 0%nat).
+ destruct ls; simpl; split; trivial; congruence. }
+ { destruct x; simpl; unfold Let_In; simpl.
+ rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
+ { do 2 f_equal.
+ rewrite <- Hop, Hid by auto.
+ rewrite !map_nth_default_always.
+ split; try reflexivity.
+ apply (IHxs (_::_)).
+ intros [|?]; autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. }
+ { intros [|?]; simpl;
+ autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. } }
+Qed.
+
+Lemma encode_bounded x : is_bounded (encode x) = true.
+Proof.
+ pose proof (bounded_encode x).
+ generalize dependent (encode x).
+ intro t; compute in t; intros.
+ destruct_head' prod.
+ unfold Pow2Base.bounded in H.
+ cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params25519_32] in H.
+ repeat match type of H with
+ | context[nth_error (cons _ _) _]
+ => let H' := fresh in
+ pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
+ cbv beta iota in H'
+ end.
+ clear H.
+ simpl in *.
+ cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
+ unfold is_bounded.
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [is_bounded proj1_fe25519_32 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe25519_32].
+ repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
+Qed.
+
+Definition encode (x : F modulus) : fe25519_32
+ := exist_fe25519_32 (encode x) (encode_bounded x).
+
+Definition decode (x : fe25519_32) : F modulus
+ := ModularBaseSystem.decode (proj1_fe25519_32 x).
+
+Lemma proj1_fe25519_32_encode x
+ : proj1_fe25519_32 (encode x) = ModularBaseSystem.encode x.
+Proof.
+ cbv [encode].
+ generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
+ intros y pf; intros; hnf in y; destruct_head_hnf' prod.
+ cbv [proj1_fe25519_32 exist_fe25519_32 proj1_fe25519_32W Build_bounded_word Build_bounded_word' fe25519_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma decode_exist_fe25519_32 x pf
+ : decode (exist_fe25519_32 x pf) = ModularBaseSystem.decode x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [decode proj1_fe25519_32 exist_fe25519_32 proj1_fe25519_32W Build_bounded_word Build_bounded_word' fe25519_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition div (f g : fe25519_32) : fe25519_32
+ := exist_fe25519_32 (div (proj1_fe25519_32 f) (proj1_fe25519_32 g)) (encode_bounded _).
+
+Definition eq (f g : fe25519_32) : Prop := eq (proj1_fe25519_32 f) (proj1_fe25519_32 g).
+
+
+Notation ibinop_correct_and_bounded irop op
+ := (forall x y,
+ is_bounded (fe25519_32WToZ x) = true
+ -> is_bounded (fe25519_32WToZ y) = true
+ -> fe25519_32WToZ (irop x y) = op (fe25519_32WToZ x) (fe25519_32WToZ y)
+ /\ is_bounded (fe25519_32WToZ (irop x y)) = true) (only parsing).
+Notation iunop_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe25519_32WToZ x) = true
+ -> fe25519_32WToZ (irop x) = op (fe25519_32WToZ x)
+ /\ is_bounded (fe25519_32WToZ (irop x)) = true) (only parsing).
+Notation iunop_FEToZ_correct irop op
+ := (forall x,
+ is_bounded (fe25519_32WToZ x) = true
+ -> word64ToZ (irop x) = op (fe25519_32WToZ x)) (only parsing).
+Notation iunop_FEToWire_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe25519_32WToZ x) = true
+ -> wire_digitsWToZ (irop x) = op (fe25519_32WToZ x)
+ /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
+Notation iunop_WireToFE_correct_and_bounded irop op
+ := (forall x,
+ wire_digits_is_bounded (wire_digitsWToZ x) = true
+ -> fe25519_32WToZ (irop x) = op (wire_digitsWToZ x)
+ /\ is_bounded (fe25519_32WToZ (irop x)) = true) (only parsing).
+
+Definition prefreeze := GF25519_32.prefreeze.
diff --git a/src/SpecificGen/GF25519_32Reflective.v b/src/SpecificGen/GF25519_32Reflective.v
new file mode 100644
index 000000000..0459f10ab
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective.v
@@ -0,0 +1,119 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Export Crypto.SpecificGen.GF25519_32.
+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.Reify.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.SpecificGen.GF25519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+Definition radd : ExprBinOp := Eval vm_compute in rcarry_addW.
+Definition rsub : ExprBinOp := Eval vm_compute in rcarry_subW.
+Definition rmul : ExprBinOp := Eval vm_compute in rmulW.
+Definition ropp : ExprUnOp := Eval vm_compute in rcarry_oppW.
+Definition rprefreeze : ExprUnOp := Eval vm_compute in rprefreezeW.
+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.
+
+Declare Reduction asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+Ltac asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+
+
+Definition interp_radd : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := Eval asm_interp in interp_bexpr (rword64ize radd).
+(*Print interp_radd.*)
+Definition interp_radd_correct : interp_radd = interp_bexpr radd := eq_refl.
+Definition interp_rsub : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rsub).
+(*Print interp_rsub.*)
+Definition interp_rsub_correct : interp_rsub = interp_bexpr rsub := eq_refl.
+Definition interp_rmul : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rmul).
+(*Print interp_rmul.*)
+Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
+Definition interp_ropp : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := Eval asm_interp in interp_uexpr (rword64ize ropp).
+(*Print interp_ropp.*)
+Definition interp_ropp_correct : interp_ropp = interp_uexpr ropp := eq_refl.
+Definition interp_rprefreeze : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := Eval asm_interp in interp_uexpr (rword64ize rprefreeze).
+(*Print interp_rprefreeze.*)
+Definition interp_rprefreeze_correct : interp_rprefreeze = interp_uexpr rprefreeze := eq_refl.
+
+Definition interp_rge_modulus : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.word64
+ := Eval asm_interp in interp_uexpr_FEToZ (rword64ize rge_modulus).
+Definition interp_rge_modulus_correct : interp_rge_modulus = interp_uexpr_FEToZ rge_modulus := eq_refl.
+
+Definition interp_rpack : SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.wire_digitsW
+ := Eval asm_interp in interp_uexpr_FEToWire (rword64ize rpack).
+Definition interp_rpack_correct : interp_rpack = interp_uexpr_FEToWire rpack := eq_refl.
+
+Definition interp_runpack : SpecificGen.GF25519_32BoundedCommon.wire_digitsW -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := Eval asm_interp in interp_uexpr_WireToFE (rword64ize 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.
+Proof. exact rcarry_addW_correct_and_bounded. Qed.
+Lemma rsub_correct_and_bounded : binop_correct_and_bounded rsub carry_sub.
+Proof. exact rcarry_subW_correct_and_bounded. Qed.
+Lemma rmul_correct_and_bounded : binop_correct_and_bounded rmul mul.
+Proof. exact rmulW_correct_and_bounded. Qed.
+Lemma ropp_correct_and_bounded : unop_correct_and_bounded ropp carry_opp.
+Proof. exact rcarry_oppW_correct_and_bounded. Qed.
+Lemma rprefreeze_correct_and_bounded : unop_correct_and_bounded rprefreeze prefreeze.
+Proof. exact rprefreezeW_correct_and_bounded. Qed.
+Lemma rge_modulus_correct_and_bounded : unop_FEToZ_correct rge_modulus ge_modulus.
+Proof. exact rge_modulusW_correct_and_bounded. Qed.
+Lemma rpack_correct_and_bounded : unop_FEToWire_correct_and_bounded rpack pack.
+Proof. exact rpackW_correct_and_bounded. Qed.
+Lemma runpack_correct_and_bounded : unop_WireToFE_correct_and_bounded runpack unpack.
+Proof. exact runpackW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF25519_32Reflective/Common.v b/src/SpecificGen/GF25519_32Reflective/Common.v
new file mode 100644
index 000000000..4105b4173
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Common.v
@@ -0,0 +1,438 @@
+Require Export Coq.ZArith.ZArith.
+Require Export Coq.Strings.String.
+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.Reify.
+Require Export Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.InterpWfRel.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Reflection.MapInterpWf.
+Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.LetIn.
+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).
+
+Local Ltac make_type_from uncurried_op :=
+ let T := (type of uncurried_op) in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ exact rT.
+
+Definition ExprBinOpT : type base_type.
+Proof. make_type_from (uncurry_binop_fe25519_32 carry_add). Defined.
+Definition ExprUnOpT : type base_type.
+Proof. make_type_from (uncurry_unop_fe25519_32 carry_opp). Defined.
+Definition ExprUnOpFEToZT : type base_type.
+Proof. make_type_from (uncurry_unop_fe25519_32 ge_modulus). Defined.
+Definition ExprUnOpWireToFET : type base_type.
+Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Definition ExprUnOpFEToWireT : type base_type.
+Proof. make_type_from (uncurry_unop_fe25519_32 pack). Defined.
+Definition ExprBinOp : Type := Expr ExprBinOpT.
+Definition ExprUnOp : Type := Expr ExprUnOpT.
+Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
+Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
+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 :: ?xs)%list => let bs := bounds_from_list xs in
+ constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ end.
+
+Local Ltac make_bounds ls :=
+ compute;
+ let v := bounds_from_list (List.rev ls) in
+ let v := (eval compute in v) in
+ exact v.
+
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
+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).
+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).
+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).
+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).
+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).
+
+Notation binop_correct_and_bounded rop op
+ := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+Notation unop_correct_and_bounded rop op
+ := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
+Notation unop_FEToZ_correct rop op
+ := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
+Notation unop_FEToWire_correct_and_bounded rop op
+ := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
+Notation unop_WireToFE_correct_and_bounded rop op
+ := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
+
+Ltac rexpr_cbv :=
+ lazymatch goal with
+ | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ => let operf := head oper in
+ let uncurryf := head uncurry in
+ try cbv delta [T]; try cbv delta [oper];
+ try cbv beta iota delta [uncurryf]
+ end;
+ cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+
+Ltac reify_sig :=
+ rexpr_cbv; eexists; Reify_rhs; reflexivity.
+
+Local Notation rexpr_sig T uncurried_op :=
+ { rexprZ
+ | interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op }
+ (only parsing).
+
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (uncurry_binop_fe25519_32 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT (uncurry_unop_fe25519_32 op)) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT (uncurry_unop_fe25519_32 op)) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT (uncurry_unop_fe25519_32 op)) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET (uncurry_unop_wire_digits op)) (only parsing).
+
+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
+ 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))
+ (only parsing).
+
+Ltac app_tuples x y :=
+ let tx := type of x in
+ lazymatch (eval hnf in tx) with
+ | prod _ _ => let xs := app_tuples (snd x) y in
+ constr:((fst x, xs))
+ | _ => constr:((x, y))
+ end.
+
+Local Arguments Tuple.map2 : simpl never.
+Local Arguments Tuple.map : simpl never.
+
+Fixpoint args_to_bounded_helperT {n}
+ (v : Tuple.tuple' Word64.word64 n)
+ (bounds : Tuple.tuple' (Z * Z) n)
+ (pf : List.fold_right
+ andb true
+ (Tuple.to_list
+ _
+ (Tuple.map2
+ (n:=S 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)
+ (res : Type)
+ {struct n}
+ : Type.
+Proof.
+ refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
+ List.fold_right
+ _ _ (Tuple.to_list
+ _
+ (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
+ | 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
+ end v bounds pf).
+ { clear -pf0.
+ abstract (
+ destruct v', bounds'; simpl @fst;
+ rewrite Tuple.map_S in pf0;
+ simpl in pf0;
+ rewrite Tuple.map2_S in pf0;
+ simpl @List.fold_right in *;
+ rewrite Bool.andb_true_iff in pf0; tauto
+ ). }
+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.
+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 |}
+ | 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) |}))
+ end.
+ { clear -pf pf'.
+ unfold Tuple.map2, Tuple.map in pf; simpl in *.
+ abstract (
+ destruct bounds;
+ simpl in *;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+ { simpl in *.
+ clear -pf pf'.
+ abstract (
+ destruct bounds as [? [? ?] ], v; simpl in *;
+ rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
+ simpl in pf;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+Defined.
+
+Definition assoc_right''
+ := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
+
+Definition args_to_bounded {n} v bounds pf
+ := Eval cbv [args_to_bounded_helper assoc_right''] in
+ @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
+
+Local Ltac get_len T :=
+ match (eval hnf in T) with
+ | prod ?A ?B
+ => let a := get_len A in
+ let b := get_len B in
+ (eval compute in (a + b)%nat)
+ | _ => constr:(1%nat)
+ end.
+
+Local Ltac args_to_bounded x H :=
+ let x' := fresh in
+ set (x' := x);
+ compute in x;
+ let len := (let T := type of x in get_len T) in
+ destruct_head' prod;
+ let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
+ let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
+ let c := (eval cbv [all_binders_for ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
+ apply c; compute; clear;
+ try abstract (
+ repeat split;
+ solve [ reflexivity
+ | refine (fun v => match v with eq_refl => I end) ]
+ ).
+
+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).
+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).
+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).
+Proof.
+ let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
+ exact v.
+Defined.
+
+Ltac assoc_right_tuple x so_far :=
+ let t := type of x in
+ lazymatch (eval hnf in t) with
+ | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
+ assoc_right_tuple (fst x) so_far
+ | _ => lazymatch so_far with
+ | @None => x
+ | _ => constr:((x, so_far))
+ end
+ end.
+
+Local Ltac make_bounds_prop bounds orig_bounds :=
+ let bounds' := fresh "bounds'" in
+ let bounds_bad := fresh "bounds_bad" in
+ rename bounds into bounds_bad;
+ let boundsv := assoc_right_tuple bounds_bad (@None) in
+ pose boundsv as bounds;
+ pose orig_bounds as bounds';
+ repeat (refine (match fst bounds' with
+ | Some bounds' => let (l, u) := fst bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end && _)%bool;
+ destruct bounds' as [_ bounds'], bounds as [_ bounds]);
+ try exact (match bounds' with
+ | Some bounds' => let (l, u) := bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end).
+
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprBinOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToWireT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpWireToFET)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToZT)) : bool.
+Proof.
+ refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
+Defined.
+
+(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
+ various kinds of correct and boundedness, and abstract in Gallina
+ rather than Ltac *)
+
+Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
+ let Heq := fresh "Heq" in
+ let Hbounds0 := fresh "Hbounds0" in
+ let Hbounds1 := fresh "Hbounds1" in
+ let Hbounds2 := fresh "Hbounds2" in
+ pose proof (proj2_sig ropZ_sig) as Heq;
+ cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
+ curry_binop_fe25519_32W curry_unop_fe25519_32W curry_unop_wire_digitsW
+ curry_binop_fe25519_32 curry_unop_fe25519_32 curry_unop_wire_digits
+ uncurry_binop_fe25519_32W uncurry_unop_fe25519_32W uncurry_unop_wire_digitsW
+ uncurry_binop_fe25519_32 uncurry_unop_fe25519_32 uncurry_unop_wire_digits
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET
+ interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ cbv zeta in *;
+ simpl @fe25519_32WToZ; simpl @wire_digitsWToZ;
+ rewrite <- Heq; clear Heq;
+ destruct Hbounds as [Heq Hbounds];
+ change interp_op with (@Z.interp_op) in *;
+ 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;
+ specialize_by repeat first [ progress intros
+ | reflexivity
+ | assumption
+ | progress destruct_head' base_type
+ | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' and
+ | progress repeat apply conj ];
+ specialize (Hbounds_left args H0);
+ specialize (Hbounds_right args H0);
+ cbv beta in *;
+ lazymatch type of Hbounds_right with
+ | match ?e with _ => _ end
+ => lazymatch type of H1 with
+ | match ?e' with _ => _ end
+ => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
+ end
+ end;
+ 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;
+ match goal with
+ | [ |- fe25519_32WToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- wire_digitsWToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- _ = _ ]
+ => exact Hbounds_left
+ end;
+ change word64ToZ with Word64.word64ToZ in *;
+ (split; [ exact Hbounds_left | ]);
+ cbv [interp_flat_type] in *;
+ cbv [fst snd
+ binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good
+ ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds] in H1;
+ destruct_head' ZBounds.bounds;
+ unfold_is_bounded_in H1;
+ simpl @fe25519_32WToZ; simpl @wire_digitsWToZ;
+ unfold_is_bounded;
+ destruct_head' and;
+ Z.ltb_to_lt;
+ change Word64.word64ToZ with word64ToZ in *;
+ repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
+
+
+Ltac rexpr_correct :=
+ let ropW' := fresh in
+ let ropZ_sig := fresh in
+ intros ropW' ropZ_sig;
+ let wf_ropW := fresh "wf_ropW" in
+ assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
+ cbv zeta; repeat apply conj;
+ [ vm_compute; reflexivity
+ | apply @InterpRelWf;
+ [ | 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 compute_bounds opW bounds
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ (only parsing).
+
+
+Module Export PrettyPrinting.
+ (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
+ 8.5/8.6. Because [Set] is special and things break if
+ [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
+ to debug. *)
+ Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
+
+ Inductive result := yes | no | borked.
+
+ Definition ZBounds_to_bounds_on
+ := 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 |}
+ => in_range l u
+ | None
+ => overflow
+ end
+ end.
+
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result
+ := match t return interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result with
+ | Tbase TZ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
+ | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
+ | no, no => no
+ | yes, no | no, yes | yes, yes => yes
+ | borked, _ | _, borked => borked
+ end
+ end.
+
+ (** This gives a slightly easier to read version of the bounds *)
+ Notation compute_bounds_for_display opW bounds
+ := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
+ Notation sanity_compute opW bounds
+ := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
+ Notation sanity_check opW bounds
+ := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
+End PrettyPrinting.
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v b/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
new file mode 100644
index 000000000..837229900
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
@@ -0,0 +1,50 @@
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprBinOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_binop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall xy
+ (xy := (eta_fe25519_32W (fst xy), eta_fe25519_32W (snd xy)))
+ (Hxy : is_bounded (fe25519_32WToZ (fst xy)) = true
+ /\ is_bounded (fe25519_32WToZ (snd xy)) = true),
+ let Hx := let (Hx, Hy) := Hxy in Hx in
+ 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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall xy
+ (xy := (eta_fe25519_32W (fst xy), eta_fe25519_32W (snd xy)))
+ (Hxy : is_bounded (fe25519_32WToZ (fst xy)) = true
+ /\ is_bounded (fe25519_32WToZ (snd xy)) = true),
+ 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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => binop_bounds_good bounds = true
+ | None => False
+ end)
+ : binop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x y Hx Hy.
+ pose x as x'; pose y as y'.
+ hnf in x, y; destruct_head' prod.
+ specialize (H0 (x', y') (conj Hx Hy)).
+ specialize (H1 (x', y') (conj Hx Hy)).
+ let args := constr:(binop_args_to_bounded (x', y') Hx Hy) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
new file mode 100644
index 000000000..864efb92b
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe25519_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unop_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
new file mode 100644
index 000000000..68efe5b86
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToWire_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe25519_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToWire_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToWire_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
new file mode 100644
index 000000000..6b9864d6f
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToZ_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe25519_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToZ_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToZ_correct (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
new file mode 100644
index 000000000..09ac72cdc
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpWireToFE_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_wire_digitsW x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopWireToFE_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_WireToFE_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified.v b/src/SpecificGen/GF25519_32Reflective/Reified.v
new file mode 100644
index 000000000..106a6beb8
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified.v
@@ -0,0 +1,13 @@
+(** We split the reification up into separate files, one operation per
+ file, so that it can run in parallel. *)
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.Add.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.CarryAdd.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.Sub.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.CarrySub.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.Mul.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.Opp.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.CarryOpp.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.PreFreeze.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.GeModulus.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.Pack.
+Require Export Crypto.SpecificGen.GF25519_32Reflective.Reified.Unpack.
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/Add.v b/src/SpecificGen/GF25519_32Reflective/Reified/Add.v
new file mode 100644
index 000000000..c35de3c76
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/Add.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonBinOp.
+
+Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
+Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
+Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
+Proof. rexpr_correct. Qed.
+Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
+Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
+Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/CarryAdd.v b/src/SpecificGen/GF25519_32Reflective/Reified/CarryAdd.v
new file mode 100644
index 000000000..d7a9923fb
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/CarryAdd.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonBinOp.
+
+Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
+Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
+Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_addW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/CarryOpp.v b/src/SpecificGen/GF25519_32Reflective/Reified/CarryOpp.v
new file mode 100644
index 000000000..9a3eb5cb4
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/CarryOpp.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOp.
+
+Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
+Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
+Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_oppW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/CarrySub.v b/src/SpecificGen/GF25519_32Reflective/Reified/CarrySub.v
new file mode 100644
index 000000000..a632e6ae7
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/CarrySub.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonBinOp.
+
+Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
+Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
+Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_subW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/GeModulus.v b/src/SpecificGen/GF25519_32Reflective/Reified/GeModulus.v
new file mode 100644
index 000000000..961940d22
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/GeModulus.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOpFEToZ.
+
+Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
+Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
+Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rge_modulusW_correct_and_bounded
+ := ExprUnOpFEToZ_correct_and_bounded
+ rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/Mul.v b/src/SpecificGen/GF25519_32Reflective/Reified/Mul.v
new file mode 100644
index 000000000..4cfa15372
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/Mul.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonBinOp.
+
+Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
+Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
+Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rmulW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
+Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
+Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/Opp.v b/src/SpecificGen/GF25519_32Reflective/Reified/Opp.v
new file mode 100644
index 000000000..c105233ba
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/Opp.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOp.
+
+Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
+Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
+Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
+Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
+Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/Pack.v b/src/SpecificGen/GF25519_32Reflective/Reified/Pack.v
new file mode 100644
index 000000000..9f32e104a
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/Pack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOpFEToWire.
+
+Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
+Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
+Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rpackW_correct_and_bounded
+ := ExprUnOpFEToWire_correct_and_bounded
+ rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v
new file mode 100644
index 000000000..1ba921294
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOp.
+
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF25519_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
+Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rprefreezeW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/Sub.v b/src/SpecificGen/GF25519_32Reflective/Reified/Sub.v
new file mode 100644
index 000000000..12881bf22
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/Sub.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonBinOp.
+
+Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
+Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
+Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
+Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
+Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/Unpack.v b/src/SpecificGen/GF25519_32Reflective/Reified/Unpack.v
new file mode 100644
index 000000000..04433583d
--- /dev/null
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/Unpack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOpWireToFE.
+
+Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
+Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
+Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition runpackW_correct_and_bounded
+ := ExprUnOpWireToFE_correct_and_bounded
+ runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/SpecificGen/GF25519_64Bounded.v b/src/SpecificGen/GF25519_64Bounded.v
new file mode 100644
index 000000000..e474eb2bc
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Bounded.v
@@ -0,0 +1,462 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF25519_64.
+Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
+Require Import Crypto.SpecificGen.GF25519_64Reflective.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+
+Local Ltac bounded_t opW blem :=
+ apply blem; apply is_bounded_proj1_fe25519_64.
+Local Ltac bounded_wire_digits_t opW blem :=
+ apply blem; apply is_bounded_proj1_wire_digits.
+
+Local Ltac define_binop f g opW blem :=
+ refine (exist_fe25519_64W (opW (proj1_fe25519_64W f) (proj1_fe25519_64W g)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop f opW blem :=
+ refine (exist_fe25519_64W (opW (proj1_fe25519_64W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_FEToZ f opW :=
+ refine (opW (proj1_fe25519_64W f)).
+Local Ltac define_unop_FEToWire f opW blem :=
+ refine (exist_wire_digitsW (opW (proj1_fe25519_64W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_WireToFE f opW blem :=
+ refine (exist_fe25519_64W (opW (proj1_wire_digitsW f)) _);
+ 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 Arguments interp_radd / _ _.
+Local Arguments interp_rsub / _ _.
+Local Arguments interp_rmul / _ _.
+Local Arguments interp_ropp / _.
+Local Arguments interp_rprefreeze / _.
+Local Arguments interp_rge_modulus / _.
+Local Arguments interp_rpack / _.
+Local Arguments interp_runpack / _.
+Definition addW (f g : fe25519_64W) : fe25519_64W := Eval simpl in interp_radd f g.
+Definition subW (f g : fe25519_64W) : fe25519_64W := Eval simpl in interp_rsub f g.
+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 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.
+
+Require Import ModularBaseSystemWord.
+Definition modulusW :=
+ Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z 10 GF25519_64.modulus_digits_)).
+
+Definition postfreeze : GF25519_64.fe25519_64 -> GF25519_64.fe25519_64 :=
+ GF25519_64.postfreeze.
+
+Lemma freeze_prepost_freeze : forall x, postfreeze (prefreeze x) = GF25519_64.freeze x. Proof. reflexivity. Qed.
+
+Definition postfreezeW : fe25519_64W -> fe25519_64W :=
+ (conditional_subtract_modulusW
+ (num_limbs := 10)
+ modulusW
+ ge_modulusW
+ (Interpretations.Word64.neg GF25519_64.int_width)
+ ).
+
+Definition freezeW (f : fe25519_64W) : fe25519_64W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
+
+Local Transparent Let_In.
+Definition powW (f : fe25519_64W) chain := fold_chain_opt (proj1_fe25519_64W one) mulW chain [f].
+Definition invW (f : fe25519_64W) : fe25519_64W
+ := Eval cbv -[Let_In fe25519_64W mulW] in powW f (chain inv_ec).
+
+Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
+ change opW with (interp_rop);
+ rewrite pre_rewrite;
+ intros; apply rop_cb; assumption.
+
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Proof. port_correct_and_bounded interp_radd_correct addW interp_radd radd_correct_and_bounded. Qed.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW carry_sub.
+Proof. port_correct_and_bounded interp_rsub_correct subW interp_rsub rsub_correct_and_bounded. Qed.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW mul.
+Proof. port_correct_and_bounded interp_rmul_correct mulW interp_rmul rmul_correct_and_bounded. Qed.
+Lemma oppW_correct_and_bounded : iunop_correct_and_bounded oppW carry_opp.
+Proof. port_correct_and_bounded interp_ropp_correct oppW interp_ropp ropp_correct_and_bounded. Qed.
+Lemma prefreezeW_correct_and_bounded : iunop_correct_and_bounded prefreezeW prefreeze.
+Proof. port_correct_and_bounded interp_rprefreeze_correct prefreezeW interp_rprefreeze rprefreeze_correct_and_bounded. Qed.
+Lemma ge_modulusW_correct : iunop_FEToZ_correct ge_modulusW ge_modulus.
+Proof. port_correct_and_bounded interp_rge_modulus_correct ge_modulusW interp_rge_modulus rge_modulus_correct_and_bounded. Qed.
+Lemma packW_correct_and_bounded : iunop_FEToWire_correct_and_bounded packW pack.
+Proof. port_correct_and_bounded interp_rpack_correct packW interp_rpack rpack_correct_and_bounded. Qed.
+Lemma unpackW_correct_and_bounded : iunop_WireToFE_correct_and_bounded unpackW unpack.
+Proof. port_correct_and_bounded interp_runpack_correct unpackW interp_runpack runpack_correct_and_bounded. Qed.
+
+(* TODO : move *)
+Lemma neg_range : forall x y, 0 <= x ->
+ 0 <= ModularBaseSystemListZOperations.neg x y < 2 ^ x.
+Proof.
+ intros.
+ split; auto using ModularBaseSystemListZOperationsProofs.neg_nonneg.
+ eapply Z.le_lt_trans; eauto using ModularBaseSystemListZOperationsProofs.neg_upperbound.
+ rewrite Z.ones_equiv.
+ omega.
+Qed.
+
+Ltac lower_bound_minus_ge_modulus :=
+ apply Z.le_0_sub;
+ 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;
+ etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
+ apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
+
+Ltac upper_bound_minus_ge_modulus :=
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | ];
+ eapply Z.le_lt_trans; [ eassumption | ];
+ instantiate; vm_compute; reflexivity.
+
+Lemma postfreezeW_correct_and_bounded : iunop_correct_and_bounded postfreezeW postfreeze.
+Proof.
+ intros x H.
+ pose proof (ge_modulusW_correct x H) as Hgm.
+ destruct_head_hnf' prod.
+ unfold_is_bounded_in H.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ cbv [postfreezeW].
+ cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
+ change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ rewrite Hgm.
+
+ cbv [modulusW Tuple.map].
+ cbv [on_tuple List.map to_list to_list' from_list from_list'
+ Tuple.map2 on_tuple2 ListUtil.map2 fe25519_64WToZ].
+ cbv [postfreeze GF25519_64.postfreeze].
+ cbv [Let_In].
+
+ split.
+ { match goal with
+ |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?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;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ | |- (_,_) = (_,_) => reflexivity
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ end. }
+
+
+ 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;
+ repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ try match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ | |- _ - _ <= _ => etransitivity; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | instantiate; assumption ]
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- ModularBaseSystemListZOperations.neg _ _ < _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- _ => vm_compute; (discriminate || reflexivity)
+ end.
+Qed.
+
+Lemma freezeW_correct_and_bounded : iunop_correct_and_bounded freezeW freeze.
+Proof.
+ intros f H; rewrite <- freeze_prepost_freeze.
+ change (freezeW f) with (postfreezeW (prefreezeW f)).
+ destruct (prefreezeW_correct_and_bounded f H) as [H0 H1].
+ destruct (postfreezeW_correct_and_bounded _ H1) as [H0' H1'].
+ rewrite H1', H0', H0; split; reflexivity.
+Qed.
+
+Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
+Proof.
+ cbv [powW pow].
+ intro x; intros; apply (fold_chain_opt_gen fe25519_64WToZ is_bounded [x]).
+ { reflexivity. }
+ { reflexivity. }
+ { intros; progress rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption.
+ apply mulW_correct_and_bounded; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption; reflexivity. }
+ { intros [|?]; autorewrite with simpl_nth_default;
+ (assumption || reflexivity). }
+Qed.
+
+Lemma invW_correct_and_bounded : iunop_correct_and_bounded invW inv.
+Proof.
+ intro f.
+ assert (H : forall f, invW f = powW f (chain inv_ec))
+ by abstract (cbv -[Let_In fe25519_64W mulW]; reflexivity).
+ rewrite H.
+ rewrite inv_correct.
+ cbv [inv_opt].
+ rewrite <- pow_correct.
+ apply powW_correct_and_bounded.
+Qed.
+
+Definition fieldwisebW_sig (f g : fe25519_64W)
+ : { b | b = GF25519_64.fieldwiseb (fe25519_64WToZ f) (fe25519_64WToZ g) }.
+Proof.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ cbv [GF25519_64.fieldwiseb fe25519_64WToZ].
+ rewrite ?word64eqb_Zeqb.
+ reflexivity.
+Defined.
+
+Definition fieldwisebW (f g : fe25519_64W) : bool :=
+ Eval cbv [proj1_sig fieldwisebW_sig appify2 app_fe25519_64W] in
+ appify2 (fun f g => proj1_sig (fieldwisebW_sig f g)) f g.
+
+Lemma fieldwisebW_correct f g
+ : fieldwisebW f g = GF25519_64.fieldwiseb (fe25519_64WToZ f) (fe25519_64WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (fieldwisebW_sig f' g')).
+Qed.
+
+Local Arguments freezeW : simpl never.
+Local Arguments fe25519_64WToZ !_ / .
+Local Opaque freezeW.
+
+Definition eqbW_sig (f g : fe25519_64W)
+ : { b | is_bounded (fe25519_64WToZ f) = true
+ -> is_bounded (fe25519_64WToZ g) = true
+ -> b = GF25519_64.eqb (fe25519_64WToZ f) (fe25519_64WToZ g) }.
+Proof.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded f pf)) as frf.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded g pf)) as frg.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ unfold GF25519_64.eqb.
+ simpl @fe25519_64WToZ in *; cbv beta iota.
+ intros.
+ rewrite <- frf, <- frg by assumption.
+ rewrite <- fieldwisebW_correct.
+ reflexivity.
+Defined.
+
+Definition eqbW (f g : fe25519_64W) : bool :=
+ Eval cbv [proj1_sig eqbW_sig appify2 app_fe25519_64W] in
+ appify2 (fun f g => proj1_sig (eqbW_sig f g)) f g.
+
+Lemma eqbW_correct f g
+ : is_bounded (fe25519_64WToZ f) = true
+ -> is_bounded (fe25519_64WToZ g) = true
+ -> eqbW f g = GF25519_64.eqb (fe25519_64WToZ f) (fe25519_64WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (eqbW_sig f' g')).
+Qed.
+
+(* TODO(jgross): use NToWord or such for this constant too *)
+Definition sqrt_m1W : fe25519_64W :=
+ Eval vm_compute in fe25519_64ZToW sqrt_m1.
+
+Definition GF25519_64sqrt x : GF25519_64.fe25519_64 :=
+ dlet powx := powW (fe25519_64ZToW x) (chain GF25519_64.sqrt_ec) in
+ GF25519_64.sqrt (fe25519_64WToZ powx) (fe25519_64WToZ (mulW powx powx)) x.
+
+Definition sqrtW_sig
+ : { sqrtW | iunop_correct_and_bounded sqrtW GF25519_64sqrt }.
+Proof.
+ eexists.
+ unfold GF25519_64sqrt, GF25519_64.sqrt.
+ intros.
+ rewrite !fe25519_64ZToW_WToZ.
+ split.
+ { etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody_eq; intros.
+ set_evars.
+ match goal with (* unfold the first dlet ... in, but only if it's binding a var *)
+ | [ |- ?x = dlet y := fe25519_64WToZ ?z in ?f ]
+ => is_var z; change (x = match fe25519_64WToZ z with y => f end)
+ end.
+ change sqrt_m1 with (fe25519_64WToZ sqrt_m1W).
+ rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded sqrt_m1W a X Y)), <- eqbW_correct, (pull_bool_if fe25519_64WToZ)
+ by repeat match goal with
+ | _ => progress subst
+ | [ |- is_bounded (fe25519_64WToZ ?op) = true ]
+ => lazymatch op with
+ | mulW _ _ => apply mulW_correct_and_bounded
+ | powW _ _ => apply powW_correct_and_bounded
+ | sqrt_m1W => vm_compute; reflexivity
+ | _ => assumption
+ end
+ end.
+ subst_evars; reflexivity.
+ } Unfocus.
+ lazymatch goal with
+ | [ |- context G[dlet x := ?v in fe25519_64WToZ (@?f x)] ]
+ => let G' := context G[fe25519_64WToZ (dlet x := v in f x)] in
+ cut G'; cbv beta;
+ [ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
+ end.
+ reflexivity. }
+ { cbv [Let_In].
+ break_if.
+ { apply powW_correct_and_bounded; assumption. }
+ { apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ].
+ apply powW_correct_and_bounded; assumption. } }
+Defined.
+
+Definition sqrtW (f : fe25519_64W) : fe25519_64W :=
+ Eval cbv [proj1_sig sqrtW_sig app_fe25519_64W] in
+ app_fe25519_64W f (proj1_sig sqrtW_sig).
+
+Lemma sqrtW_correct_and_bounded : iunop_correct_and_bounded sqrtW GF25519_64sqrt.
+Proof.
+ intro f.
+ set (f' := f).
+ hnf in f; destruct_head' prod.
+ assert (H : sqrtW f' = proj1_sig sqrtW_sig f')
+ by (subst f'; cbv beta iota delta [proj1_sig sqrtW_sig sqrtW]; reflexivity).
+ rewrite H.
+ exact (proj2_sig sqrtW_sig f').
+Qed.
+
+
+
+Definition add (f g : fe25519_64) : fe25519_64.
+Proof. define_binop f g addW addW_correct_and_bounded. Defined.
+Definition sub (f g : fe25519_64) : fe25519_64.
+Proof. define_binop f g subW subW_correct_and_bounded. Defined.
+Definition mul (f g : fe25519_64) : fe25519_64.
+Proof. define_binop f g mulW mulW_correct_and_bounded. Defined.
+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.
+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.
+Definition unpack (f : wire_digits) : fe25519_64.
+Proof. define_unop_WireToFE f unpackW unpackW_correct_and_bounded. Defined.
+
+Definition pow (f : fe25519_64) (chain : list (nat * nat)) : fe25519_64.
+Proof. define_unop f (fun x => powW x chain) powW_correct_and_bounded. Defined.
+Definition inv (f : fe25519_64) : fe25519_64.
+Proof. define_unop f invW invW_correct_and_bounded. Defined.
+Definition sqrt (f : fe25519_64) : fe25519_64.
+Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
+
+Local Ltac op_correct_t op opW_correct_and_bounded :=
+ cbv [op];
+ lazymatch goal with
+ | [ |- context[proj1_fe25519_64 (exist_fe25519_64W _ _)] ]
+ => rewrite proj1_fe25519_64_exist_fe25519_64W
+ | [ |- context[proj1_wire_digits (exist_wire_digitsW _ _)] ]
+ => rewrite proj1_wire_digits_exist_wire_digitsW
+ | _ => idtac
+ end;
+ apply opW_correct_and_bounded;
+ lazymatch goal with
+ | [ |- is_bounded _ = true ]
+ => apply is_bounded_proj1_fe25519_64
+ | [ |- wire_digits_is_bounded _ = true ]
+ => apply is_bounded_proj1_wire_digits
+ end.
+
+Lemma add_correct (f g : fe25519_64) : proj1_fe25519_64 (add f g) = carry_add (proj1_fe25519_64 f) (proj1_fe25519_64 g).
+Proof. op_correct_t add addW_correct_and_bounded. Qed.
+Lemma sub_correct (f g : fe25519_64) : proj1_fe25519_64 (sub f g) = carry_sub (proj1_fe25519_64 f) (proj1_fe25519_64 g).
+Proof. op_correct_t sub subW_correct_and_bounded. Qed.
+Lemma mul_correct (f g : fe25519_64) : proj1_fe25519_64 (mul f g) = GF25519_64.mul (proj1_fe25519_64 f) (proj1_fe25519_64 g).
+Proof. op_correct_t mul mulW_correct_and_bounded. Qed.
+Lemma opp_correct (f : fe25519_64) : proj1_fe25519_64 (opp f) = carry_opp (proj1_fe25519_64 f).
+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).
+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.
+Lemma unpack_correct (f : wire_digits) : proj1_fe25519_64 (unpack f) = GF25519_64.unpack (proj1_wire_digits f).
+Proof. op_correct_t unpack unpackW_correct_and_bounded. Qed.
+Lemma pow_correct (f : fe25519_64) chain : proj1_fe25519_64 (pow f chain) = GF25519_64.pow (proj1_fe25519_64 f) chain.
+Proof. op_correct_t pow (powW_correct_and_bounded chain). Qed.
+Lemma inv_correct (f : fe25519_64) : proj1_fe25519_64 (inv f) = GF25519_64.inv (proj1_fe25519_64 f).
+Proof. op_correct_t inv invW_correct_and_bounded. Qed.
+Lemma sqrt_correct (f : fe25519_64) : proj1_fe25519_64 (sqrt f) = GF25519_64sqrt (proj1_fe25519_64 f).
+Proof. op_correct_t sqrt sqrtW_correct_and_bounded. Qed.
+
+Import Morphisms.
+
+Local Existing Instance prime_modulus.
+
+Lemma field25519_64_and_homomorphisms
+ : @field fe25519_64 eq zero one opp add sub mul inv div
+ /\ @Ring.is_homomorphism (F _) (@Logic.eq _) 1%F F.add F.mul fe25519_64 eq one add mul encode
+ /\ @Ring.is_homomorphism fe25519_64 eq one add mul (F _) (@Logic.eq _) 1%F F.add F.mul decode.
+Proof.
+ eapply @Field.field_and_homomorphism_from_redundant_representation.
+ { exact (F.field_modulo _). }
+ { cbv [decode encode]; intros; rewrite !proj1_fe25519_64_exist_fe25519_64; apply encode_rep. }
+ { reflexivity. }
+ { reflexivity. }
+ { reflexivity. }
+ { cbv [decode encode]; intros; rewrite opp_correct, carry_opp_correct, carry_opp_opt_correct, carry_opp_rep; apply opp_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite add_correct, carry_add_correct, carry_add_opt_correct, carry_add_rep; apply add_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite sub_correct, carry_sub_correct, carry_sub_opt_correct, carry_sub_rep; apply sub_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite mul_correct, GF25519_64.mul_correct, carry_mul_opt_correct by reflexivity; apply carry_mul_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite inv_correct, GF25519_64.inv_correct, inv_opt_correct by reflexivity; apply inv_rep; reflexivity. }
+ { cbv [decode encode div]; intros; rewrite !proj1_fe25519_64_exist_fe25519_64; apply encode_rep. }
+Qed.
+
+Global Instance field25519_64 : @field fe25519_64 eq zero one opp add sub mul inv div := proj1 field25519_64_and_homomorphisms.
+
+Local Opaque proj1_fe25519_64 exist_fe25519_64 proj1_fe25519_64W exist_fe25519_64W.
+Global Instance homomorphism_F25519_64_encode
+ : @Ring.is_homomorphism (F modulus) Logic.eq F.one F.add F.mul fe25519_64 eq one add mul encode.
+Proof. apply field25519_64_and_homomorphisms. Qed.
+
+Global Instance homomorphism_F25519_64_decode
+ : @Ring.is_homomorphism fe25519_64 eq one add mul (F modulus) Logic.eq F.one F.add F.mul decode.
+Proof. apply field25519_64_and_homomorphisms. Qed.
diff --git a/src/SpecificGen/GF25519_64BoundedCommon.v b/src/SpecificGen/GF25519_64BoundedCommon.v
new file mode 100644
index 000000000..e6a3b6c20
--- /dev/null
+++ b/src/SpecificGen/GF25519_64BoundedCommon.v
@@ -0,0 +1,693 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF25519_64.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.HList.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+(* BEGIN common curve-specific definitions *)
+Definition bit_width : nat := Eval compute in Z.to_nat (GF25519_64.int_width).
+Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
+Definition bounds_exp : tuple Z length_fe25519_64
+ := Eval compute in
+ Tuple.from_list length_fe25519_64 limb_widths eq_refl.
+Definition bounds : tuple (Z * Z) length_fe25519_64
+ := Eval compute in
+ Tuple.map (fun e => b_of e) bounds_exp.
+Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
+ := Eval compute in Tuple.from_list _ wire_widths eq_refl.
+Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
+ := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
+(* 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.
+Arguments bit_width : simpl never.
+Global Opaque word64.
+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},
+ 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},
+ 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
+ := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
+ Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ : 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}
+ : 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 }.
+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)).
+Definition fe25519_64 :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => word_of e) bounds_exp.
+Definition wire_digits :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
+
+Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
+ := let res := Tuple.map2
+ (fun bounds v =>
+ let '(lower, upper) := bounds in
+ (lower <=? v) && (v <=? upper))%bool%Z
+ bounds x in
+ List.fold_right andb true (Tuple.to_list _ res).
+
+Definition is_bounded (x : SpecificGen.GF25519_64.fe25519_64) : bool
+ := is_bounded_gen (n:=length_fe25519_64) x bounds.
+
+Definition wire_digits_is_bounded (x : SpecificGen.GF25519_64.wire_digits) : bool
+ := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
+
+(* END basic types *)
+
+Section generic_destructuring.
+ Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
+ := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
+ end.
+ Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
+ := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => @app_on' A n'
+ end.
+ Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
+ Qed.
+ Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
+
+ Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
+ end.
+ Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => @app_on_h' A F n'
+ end.
+ Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
+ Qed.
+ Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
+
+ Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
+ : forall (f : tuple A (length wire_widths)), T f
+ := Eval compute in fun f => @app_on A (length wire_widths) T f P.
+ Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
+ := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
+ Definition app_fe25519_64W_dep {A T} (P : forall x : tuple A length_fe25519_64, T x)
+ : forall (f : tuple A length_fe25519_64), T f
+ := Eval compute in fun f => @app_on A length_fe25519_64 T f P.
+ Definition app_fe25519_64W {A T} (f : tuple A length_fe25519_64) (P : tuple A length_fe25519_64 -> T)
+ := Eval compute in @app_fe25519_64W_dep A (fun _ => T) P f.
+ Definition app_fe25519_64_dep {T} (P : forall x : fe25519_64, T x)
+ : forall f : fe25519_64, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe25519_64 bounds_exp T f P.
+ Definition app_fe25519_64 {T} (f : fe25519_64) (P : hlist (fun e => word_of e) bounds_exp -> T)
+ := Eval compute in @app_fe25519_64_dep (fun _ => T) P f.
+ Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
+ : forall f : wire_digits, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
+ Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
+ := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
+
+ Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
+ := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
+ Definition app_fe25519_64W_dep_correct {A T} f P : @app_fe25519_64W_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_fe25519_64W_correct {A T} f P : @app_fe25519_64W A T f P = P f
+ := @app_fe25519_64W_dep_correct A (fun _ => T) f P.
+ Definition app_fe25519_64_dep_correct {T} f P : @app_fe25519_64_dep T P f = P f
+ := app_on_h_correct (fun e => word_of e) bounds_exp f P.
+ Definition app_fe25519_64_correct {T} f P : @app_fe25519_64 T f P = P f
+ := @app_fe25519_64_dep_correct (fun _ => T) f P.
+ Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
+ := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
+ Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
+ := @app_wire_digits_dep_correct (fun _ => T) f P.
+
+ Definition appify2 {T} (op : fe25519_64W -> fe25519_64W -> T) (f g : fe25519_64W) :=
+ app_fe25519_64W f (fun f0 => (app_fe25519_64W g (fun g0 => op f0 g0))).
+
+ Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
+ Proof.
+ intros. cbv [appify2].
+ etransitivity; apply app_fe25519_64W_correct.
+ Qed.
+End generic_destructuring.
+
+Definition eta_fe25519_64W_sig (x : fe25519_64W) : { v : fe25519_64W | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+Defined.
+Definition eta_fe25519_64W (x : fe25519_64W) : fe25519_64W
+ := Eval cbv [proj1_sig eta_fe25519_64W_sig] in proj1_sig (eta_fe25519_64W_sig x).
+Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+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.
+Proof. apply NToWord_wordToN. Qed.
+Local Opaque word64.
+
+Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ rewrite N2Z.id, NToWord_wordToN; reflexivity.
+Qed.
+
+(* BEGIN precomputation. *)
+
+Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
+Definition word_bounded {lower upper} (v : bounded_word lower upper)
+ : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
+ := project_is_boundedT (proj2_sig v).
+Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
+ := exist _ proj_word (Build_is_boundedT word_bounded).
+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)
+ : bounded_word lower upper
+ := Build_bounded_word'
+ proj_word
+ (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
+ | 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.
+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.
+ match goal with
+ | [ |- context[andb ?x ?y] ]
+ => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
+ end;
+ intros; omega.
+Qed.
+Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
+Proof.
+ refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
+ abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
+Defined.
+Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
+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.
+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)).
+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)).
+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)).
+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
+ := Eval app_tuple_map in
+ app_fe25519_64W x (Tuple.map word64ize).
+Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map word64ize).
+
+(** TODO: Turn this into a lemma to speed up proofs *)
+Ltac unfold_is_bounded_in H :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe25519_64WToZ, wire_digitsWToZ in H;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe25519_64 List.length wire_widths] in H;
+ rewrite ?Bool.andb_true_iff in H.
+
+Ltac unfold_is_bounded :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe25519_64WToZ, wire_digitsWToZ;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe25519_64 List.length wire_widths];
+ 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.
+Local Opaque bit_width.
+
+Local Ltac prove_lt_bit_width :=
+ rewrite unfold_Pow2_64; cbv [Pow2_64]; 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.
+Qed.
+
+Lemma fe25519_64WToZ_ZToW x : is_bounded x = true -> fe25519_64WToZ (fe25519_64ZToW x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe25519_64WToZ fe25519_64ZToW].
+ intro H.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma fe25519_64W_word64ize_id x : fe25519_64W_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [fe25519_64W_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [wire_digitsW_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+
+Definition uncurry_unop_fe25519_64W {T} (op : fe25519_64W -> T)
+ := Eval cbv (*-[word64]*) 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).
+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)).
+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)).
+
+Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
+ := Eval cbv (*-[word64]*) 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).
+
+
+Definition proj1_fe25519_64W (x : fe25519_64) : fe25519_64W
+ := Eval app_tuple_map in
+ app_fe25519_64 x (HList.mapt (fun _ => (@proj_word _ _))).
+Coercion proj1_fe25519_64 (x : fe25519_64) : SpecificGen.GF25519_64.fe25519_64
+ := fe25519_64WToZ (proj1_fe25519_64W x).
+
+Lemma is_bounded_proj1_fe25519_64 (x : fe25519_64) : is_bounded (proj1_fe25519_64 x) = true.
+Proof.
+ revert x; refine (app_fe25519_64_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [is_bounded proj1_fe25519_64 proj1_fe25519_64W fe25519_64WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe25519_64 is_bounded_gen].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
+ := app_wire_digits x (HList.mapt (fun _ => proj_word)).
+Coercion proj1_wire_digits (x : wire_digits) : SpecificGen.GF25519_64.wire_digits
+ := wire_digitsWToZ (proj1_wire_digitsW x).
+
+Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
+Proof.
+ revert x; refine (app_wire_digits_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Local Ltac make_exist_W' x app_W_dep :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x);
+ cbv [tuple tuple' length_fe25519_64 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v H :=
+ first [ let v' := (eval cbv [snd fst] in (snd v)) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
+ | let v' := (eval cbv [snd fst] in v) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
+ let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
+ let T := type of H' in
+ let T := (eval cbv [id
+ List.fold_right List.map List.length List.app ListUtil.map2 List.rev
+ Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
+ fe25519_64 bounds fe25519_64WToZ length_fe25519_64
+ wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
+ let H' := constr:(H' : T) in
+ let v := (eval unfold x' in x') in
+ do_refine v H'.
+Local Ltac make_exist'' x exist_W ZToW :=
+ let H := fresh in
+ intro H; apply (exist_W (ZToW x));
+ abstract (
+ hnf in x; destruct_head' prod;
+ let H' := fresh in
+ pose proof H as H';
+ unfold_is_bounded_in H;
+ destruct_head' and; simpl in *;
+ Z.ltb_to_lt;
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ assumption
+ ).
+Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x) in *;
+ cbv [tuple tuple' length_fe25519_64 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v :=
+ first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
+ | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
+ let v := (eval unfold x' in (exist'' x' H)) in
+ do_refine v.
+
+Definition exist_fe25519_64W' (x : fe25519_64W) : is_bounded (fe25519_64WToZ x) = true -> fe25519_64.
+Proof. make_exist_W' x (@app_fe25519_64W_dep). Defined.
+Definition exist_fe25519_64W (x : fe25519_64W) : is_bounded (fe25519_64WToZ x) = true -> fe25519_64
+ := Eval cbv [app_fe25519_64W_dep exist_fe25519_64W' fe25519_64ZToW] in exist_fe25519_64W' x.
+Definition exist_fe25519_64'' (x : SpecificGen.GF25519_64.fe25519_64) : is_bounded x = true -> fe25519_64.
+Proof. make_exist'' x exist_fe25519_64W fe25519_64ZToW. Defined.
+Definition exist_fe25519_64' (x : SpecificGen.GF25519_64.fe25519_64) : is_bounded x = true -> fe25519_64.
+Proof. make_exist' x (@app_fe25519_64W_dep) exist_fe25519_64'' exist_fe25519_64W fe25519_64ZToW. Defined.
+Definition exist_fe25519_64 (x : SpecificGen.GF25519_64.fe25519_64) : is_bounded x = true -> fe25519_64
+ := Eval cbv [exist_fe25519_64' exist_fe25519_64W exist_fe25519_64' app_fe25519_64 app_fe25519_64W_dep] in
+ exist_fe25519_64' x.
+
+Lemma proj1_fe25519_64_exist_fe25519_64W x pf : proj1_fe25519_64 (exist_fe25519_64W x pf) = fe25519_64WToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe25519_64W_exist_fe25519_64 x pf : proj1_fe25519_64W (exist_fe25519_64 x pf) = fe25519_64ZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe25519_64_exist_fe25519_64 x pf : proj1_fe25519_64 (exist_fe25519_64 x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_fe25519_64 exist_fe25519_64 proj1_fe25519_64W fe25519_64WToZ proj_word Build_bounded_word Build_bounded_word'].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition exist_wire_digitsW' (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
+Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
+Definition exist_wire_digitsW (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
+ := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
+Definition exist_wire_digits'' (x : SpecificGen.GF25519_64.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits' (x : SpecificGen.GF25519_64.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits (x : SpecificGen.GF25519_64.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits
+ := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
+ exist_wire_digits' x.
+
+Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 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 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.
+End opt.
+
+Local Transparent bit_width.
+Local Ltac do_change lem :=
+ match lem with
+ | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
+ => 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
+ 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.
+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 := 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).
+Defined.
+Definition wire_digits_word64ize (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 := do_change lem in
+ let lem := (eval cbv [word64ize opt.word64ize 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).
+Defined.
+
+Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
+ : forall n : nat,
+ (n < length limb_widths)%nat
+ -> (0 <= nth_default 0 (Tuple.to_list length_fe25519_64 x) n <=
+ snd (b_of (nth_default (-1) limb_widths n)))%Z.
+Proof.
+ hnf in x; destruct_head' prod.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ unfold nth_default; simpl.
+ intros.
+ repeat match goal with
+ | [ |- context[nth_error _ ?x] ]
+ => is_var x; destruct x; simpl
+ end;
+ omega.
+Qed.
+
+(* END precomputation *)
+
+(* 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 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'.
+
+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)
+ (Hid : id = F id')
+ (Hop_bounded : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> is_bounded (op (F x) (F y)) = true)
+ (Hop : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> op (F x) (F y) = F (op' x y))
+ (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
+ : F (fold_chain_opt id' op' chain ls)
+ = fold_chain_opt id op chain (List.map F ls)
+ /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
+Proof.
+ rewrite !fold_chain_opt_correct.
+ revert dependent ls; induction chain as [|x xs IHxs]; intros.
+ { pose proof (Hls_bounded 0%nat).
+ destruct ls; simpl; split; trivial; congruence. }
+ { destruct x; simpl; unfold Let_In; simpl.
+ rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
+ { do 2 f_equal.
+ rewrite <- Hop, Hid by auto.
+ rewrite !map_nth_default_always.
+ split; try reflexivity.
+ apply (IHxs (_::_)).
+ intros [|?]; autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. }
+ { intros [|?]; simpl;
+ autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. } }
+Qed.
+
+Lemma encode_bounded x : is_bounded (encode x) = true.
+Proof.
+ pose proof (bounded_encode x).
+ generalize dependent (encode x).
+ intro t; compute in t; intros.
+ destruct_head' prod.
+ unfold Pow2Base.bounded in H.
+ cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params25519_64] in H.
+ repeat match type of H with
+ | context[nth_error (cons _ _) _]
+ => let H' := fresh in
+ pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
+ cbv beta iota in H'
+ end.
+ clear H.
+ simpl in *.
+ cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
+ unfold is_bounded.
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [is_bounded proj1_fe25519_64 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe25519_64].
+ repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
+Qed.
+
+Definition encode (x : F modulus) : fe25519_64
+ := exist_fe25519_64 (encode x) (encode_bounded x).
+
+Definition decode (x : fe25519_64) : F modulus
+ := ModularBaseSystem.decode (proj1_fe25519_64 x).
+
+Lemma proj1_fe25519_64_encode x
+ : proj1_fe25519_64 (encode x) = ModularBaseSystem.encode x.
+Proof.
+ cbv [encode].
+ generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
+ intros y pf; intros; hnf in y; destruct_head_hnf' prod.
+ cbv [proj1_fe25519_64 exist_fe25519_64 proj1_fe25519_64W Build_bounded_word Build_bounded_word' fe25519_64WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma decode_exist_fe25519_64 x pf
+ : decode (exist_fe25519_64 x pf) = ModularBaseSystem.decode x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [decode proj1_fe25519_64 exist_fe25519_64 proj1_fe25519_64W Build_bounded_word Build_bounded_word' fe25519_64WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition div (f g : fe25519_64) : fe25519_64
+ := exist_fe25519_64 (div (proj1_fe25519_64 f) (proj1_fe25519_64 g)) (encode_bounded _).
+
+Definition eq (f g : fe25519_64) : Prop := eq (proj1_fe25519_64 f) (proj1_fe25519_64 g).
+
+
+Notation ibinop_correct_and_bounded irop op
+ := (forall x y,
+ is_bounded (fe25519_64WToZ x) = true
+ -> is_bounded (fe25519_64WToZ y) = true
+ -> fe25519_64WToZ (irop x y) = op (fe25519_64WToZ x) (fe25519_64WToZ y)
+ /\ is_bounded (fe25519_64WToZ (irop x y)) = true) (only parsing).
+Notation iunop_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe25519_64WToZ x) = true
+ -> fe25519_64WToZ (irop x) = op (fe25519_64WToZ x)
+ /\ is_bounded (fe25519_64WToZ (irop x)) = true) (only parsing).
+Notation iunop_FEToZ_correct irop op
+ := (forall x,
+ is_bounded (fe25519_64WToZ x) = true
+ -> word64ToZ (irop x) = op (fe25519_64WToZ x)) (only parsing).
+Notation iunop_FEToWire_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe25519_64WToZ x) = true
+ -> wire_digitsWToZ (irop x) = op (fe25519_64WToZ x)
+ /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
+Notation iunop_WireToFE_correct_and_bounded irop op
+ := (forall x,
+ wire_digits_is_bounded (wire_digitsWToZ x) = true
+ -> fe25519_64WToZ (irop x) = op (wire_digitsWToZ x)
+ /\ is_bounded (fe25519_64WToZ (irop x)) = true) (only parsing).
+
+Definition prefreeze := GF25519_64.prefreeze.
diff --git a/src/SpecificGen/GF25519_64Reflective.v b/src/SpecificGen/GF25519_64Reflective.v
new file mode 100644
index 000000000..095def252
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective.v
@@ -0,0 +1,119 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Export Crypto.SpecificGen.GF25519_64.
+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.Reify.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.SpecificGen.GF25519_64Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+Definition radd : ExprBinOp := Eval vm_compute in rcarry_addW.
+Definition rsub : ExprBinOp := Eval vm_compute in rcarry_subW.
+Definition rmul : ExprBinOp := Eval vm_compute in rmulW.
+Definition ropp : ExprUnOp := Eval vm_compute in rcarry_oppW.
+Definition rprefreeze : ExprUnOp := Eval vm_compute in rprefreezeW.
+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.
+
+Declare Reduction asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+Ltac asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ 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).
+(*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).
+(*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).
+(*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).
+(*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).
+(*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_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).
+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).
+Definition interp_runpack_correct : interp_runpack = interp_uexpr_WireToFE runpack := eq_refl.
+
+Lemma radd_correct_and_bounded : binop_correct_and_bounded radd carry_add.
+Proof. exact rcarry_addW_correct_and_bounded. Qed.
+Lemma rsub_correct_and_bounded : binop_correct_and_bounded rsub carry_sub.
+Proof. exact rcarry_subW_correct_and_bounded. Qed.
+Lemma rmul_correct_and_bounded : binop_correct_and_bounded rmul mul.
+Proof. exact rmulW_correct_and_bounded. Qed.
+Lemma ropp_correct_and_bounded : unop_correct_and_bounded ropp carry_opp.
+Proof. exact rcarry_oppW_correct_and_bounded. Qed.
+Lemma rprefreeze_correct_and_bounded : unop_correct_and_bounded rprefreeze prefreeze.
+Proof. exact rprefreezeW_correct_and_bounded. Qed.
+Lemma rge_modulus_correct_and_bounded : unop_FEToZ_correct rge_modulus ge_modulus.
+Proof. exact rge_modulusW_correct_and_bounded. Qed.
+Lemma rpack_correct_and_bounded : unop_FEToWire_correct_and_bounded rpack pack.
+Proof. exact rpackW_correct_and_bounded. Qed.
+Lemma runpack_correct_and_bounded : unop_WireToFE_correct_and_bounded runpack unpack.
+Proof. exact runpackW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF25519_64Reflective/Common.v b/src/SpecificGen/GF25519_64Reflective/Common.v
new file mode 100644
index 000000000..444edcde4
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Common.v
@@ -0,0 +1,438 @@
+Require Export Coq.ZArith.ZArith.
+Require Export Coq.Strings.String.
+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.Reify.
+Require Export Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.InterpWfRel.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Reflection.MapInterpWf.
+Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.LetIn.
+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).
+
+Local Ltac make_type_from uncurried_op :=
+ let T := (type of uncurried_op) in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ exact rT.
+
+Definition ExprBinOpT : type base_type.
+Proof. make_type_from (uncurry_binop_fe25519_64 carry_add). Defined.
+Definition ExprUnOpT : type base_type.
+Proof. make_type_from (uncurry_unop_fe25519_64 carry_opp). Defined.
+Definition ExprUnOpFEToZT : type base_type.
+Proof. make_type_from (uncurry_unop_fe25519_64 ge_modulus). Defined.
+Definition ExprUnOpWireToFET : type base_type.
+Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Definition ExprUnOpFEToWireT : type base_type.
+Proof. make_type_from (uncurry_unop_fe25519_64 pack). Defined.
+Definition ExprBinOp : Type := Expr ExprBinOpT.
+Definition ExprUnOp : Type := Expr ExprUnOpT.
+Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
+Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
+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 :: ?xs)%list => let bs := bounds_from_list xs in
+ constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ end.
+
+Local Ltac make_bounds ls :=
+ compute;
+ let v := bounds_from_list (List.rev ls) in
+ let v := (eval compute in v) in
+ exact v.
+
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
+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).
+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).
+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).
+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).
+
+Notation binop_correct_and_bounded rop op
+ := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+Notation unop_correct_and_bounded rop op
+ := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
+Notation unop_FEToZ_correct rop op
+ := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
+Notation unop_FEToWire_correct_and_bounded rop op
+ := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
+Notation unop_WireToFE_correct_and_bounded rop op
+ := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
+
+Ltac rexpr_cbv :=
+ lazymatch goal with
+ | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ => let operf := head oper in
+ let uncurryf := head uncurry in
+ try cbv delta [T]; try cbv delta [oper];
+ try cbv beta iota delta [uncurryf]
+ end;
+ cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+
+Ltac reify_sig :=
+ rexpr_cbv; eexists; Reify_rhs; reflexivity.
+
+Local Notation rexpr_sig T uncurried_op :=
+ { rexprZ
+ | interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op }
+ (only parsing).
+
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (uncurry_binop_fe25519_64 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT (uncurry_unop_fe25519_64 op)) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT (uncurry_unop_fe25519_64 op)) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT (uncurry_unop_fe25519_64 op)) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET (uncurry_unop_wire_digits op)) (only parsing).
+
+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
+ 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))
+ (only parsing).
+
+Ltac app_tuples x y :=
+ let tx := type of x in
+ lazymatch (eval hnf in tx) with
+ | prod _ _ => let xs := app_tuples (snd x) y in
+ constr:((fst x, xs))
+ | _ => constr:((x, y))
+ end.
+
+Local Arguments Tuple.map2 : simpl never.
+Local Arguments Tuple.map : simpl never.
+
+Fixpoint args_to_bounded_helperT {n}
+ (v : Tuple.tuple' Word64.word64 n)
+ (bounds : Tuple.tuple' (Z * Z) n)
+ (pf : List.fold_right
+ andb true
+ (Tuple.to_list
+ _
+ (Tuple.map2
+ (n:=S 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)
+ (res : Type)
+ {struct n}
+ : Type.
+Proof.
+ refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
+ List.fold_right
+ _ _ (Tuple.to_list
+ _
+ (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
+ | 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
+ end v bounds pf).
+ { clear -pf0.
+ abstract (
+ destruct v', bounds'; simpl @fst;
+ rewrite Tuple.map_S in pf0;
+ simpl in pf0;
+ rewrite Tuple.map2_S in pf0;
+ simpl @List.fold_right in *;
+ rewrite Bool.andb_true_iff in pf0; tauto
+ ). }
+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.
+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 |}
+ | 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) |}))
+ end.
+ { clear -pf pf'.
+ unfold Tuple.map2, Tuple.map in pf; simpl in *.
+ abstract (
+ destruct bounds;
+ simpl in *;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+ { simpl in *.
+ clear -pf pf'.
+ abstract (
+ destruct bounds as [? [? ?] ], v; simpl in *;
+ rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
+ simpl in pf;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+Defined.
+
+Definition assoc_right''
+ := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
+
+Definition args_to_bounded {n} v bounds pf
+ := Eval cbv [args_to_bounded_helper assoc_right''] in
+ @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
+
+Local Ltac get_len T :=
+ match (eval hnf in T) with
+ | prod ?A ?B
+ => let a := get_len A in
+ let b := get_len B in
+ (eval compute in (a + b)%nat)
+ | _ => constr:(1%nat)
+ end.
+
+Local Ltac args_to_bounded x H :=
+ let x' := fresh in
+ set (x' := x);
+ compute in x;
+ let len := (let T := type of x in get_len T) in
+ destruct_head' prod;
+ let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
+ let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
+ let c := (eval cbv [all_binders_for ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
+ apply c; compute; clear;
+ try abstract (
+ repeat split;
+ solve [ reflexivity
+ | refine (fun v => match v with eq_refl => I end) ]
+ ).
+
+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).
+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).
+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).
+Proof.
+ let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
+ exact v.
+Defined.
+
+Ltac assoc_right_tuple x so_far :=
+ let t := type of x in
+ lazymatch (eval hnf in t) with
+ | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
+ assoc_right_tuple (fst x) so_far
+ | _ => lazymatch so_far with
+ | @None => x
+ | _ => constr:((x, so_far))
+ end
+ end.
+
+Local Ltac make_bounds_prop bounds orig_bounds :=
+ let bounds' := fresh "bounds'" in
+ let bounds_bad := fresh "bounds_bad" in
+ rename bounds into bounds_bad;
+ let boundsv := assoc_right_tuple bounds_bad (@None) in
+ pose boundsv as bounds;
+ pose orig_bounds as bounds';
+ repeat (refine (match fst bounds' with
+ | Some bounds' => let (l, u) := fst bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end && _)%bool;
+ destruct bounds' as [_ bounds'], bounds as [_ bounds]);
+ try exact (match bounds' with
+ | Some bounds' => let (l, u) := bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end).
+
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprBinOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToWireT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpWireToFET)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToZT)) : bool.
+Proof.
+ refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
+Defined.
+
+(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
+ various kinds of correct and boundedness, and abstract in Gallina
+ rather than Ltac *)
+
+Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
+ let Heq := fresh "Heq" in
+ let Hbounds0 := fresh "Hbounds0" in
+ let Hbounds1 := fresh "Hbounds1" in
+ let Hbounds2 := fresh "Hbounds2" in
+ pose proof (proj2_sig ropZ_sig) as Heq;
+ cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
+ curry_binop_fe25519_64W curry_unop_fe25519_64W curry_unop_wire_digitsW
+ curry_binop_fe25519_64 curry_unop_fe25519_64 curry_unop_wire_digits
+ uncurry_binop_fe25519_64W uncurry_unop_fe25519_64W uncurry_unop_wire_digitsW
+ uncurry_binop_fe25519_64 uncurry_unop_fe25519_64 uncurry_unop_wire_digits
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET
+ interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ cbv zeta in *;
+ simpl @fe25519_64WToZ; simpl @wire_digitsWToZ;
+ rewrite <- Heq; clear Heq;
+ destruct Hbounds as [Heq Hbounds];
+ change interp_op with (@Z.interp_op) in *;
+ 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;
+ specialize_by repeat first [ progress intros
+ | reflexivity
+ | assumption
+ | progress destruct_head' base_type
+ | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' and
+ | progress repeat apply conj ];
+ specialize (Hbounds_left args H0);
+ specialize (Hbounds_right args H0);
+ cbv beta in *;
+ lazymatch type of Hbounds_right with
+ | match ?e with _ => _ end
+ => lazymatch type of H1 with
+ | match ?e' with _ => _ end
+ => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
+ end
+ end;
+ 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;
+ match goal with
+ | [ |- fe25519_64WToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- wire_digitsWToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- _ = _ ]
+ => exact Hbounds_left
+ end;
+ change word64ToZ with Word64.word64ToZ in *;
+ (split; [ exact Hbounds_left | ]);
+ cbv [interp_flat_type] in *;
+ cbv [fst snd
+ binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good
+ ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds] in H1;
+ destruct_head' ZBounds.bounds;
+ unfold_is_bounded_in H1;
+ simpl @fe25519_64WToZ; simpl @wire_digitsWToZ;
+ unfold_is_bounded;
+ destruct_head' and;
+ Z.ltb_to_lt;
+ change Word64.word64ToZ with word64ToZ in *;
+ repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
+
+
+Ltac rexpr_correct :=
+ let ropW' := fresh in
+ let ropZ_sig := fresh in
+ intros ropW' ropZ_sig;
+ let wf_ropW := fresh "wf_ropW" in
+ assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
+ cbv zeta; repeat apply conj;
+ [ vm_compute; reflexivity
+ | apply @InterpRelWf;
+ [ | 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 compute_bounds opW bounds
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ (only parsing).
+
+
+Module Export PrettyPrinting.
+ (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
+ 8.5/8.6. Because [Set] is special and things break if
+ [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
+ to debug. *)
+ Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
+
+ Inductive result := yes | no | borked.
+
+ Definition ZBounds_to_bounds_on
+ := 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 |}
+ => in_range l u
+ | None
+ => overflow
+ end
+ end.
+
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result
+ := match t return interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result with
+ | Tbase TZ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
+ | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
+ | no, no => no
+ | yes, no | no, yes | yes, yes => yes
+ | borked, _ | _, borked => borked
+ end
+ end.
+
+ (** This gives a slightly easier to read version of the bounds *)
+ Notation compute_bounds_for_display opW bounds
+ := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
+ Notation sanity_compute opW bounds
+ := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
+ Notation sanity_check opW bounds
+ := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
+End PrettyPrinting.
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v b/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
new file mode 100644
index 000000000..cd4a5031b
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
@@ -0,0 +1,50 @@
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprBinOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_binop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall xy
+ (xy := (eta_fe25519_64W (fst xy), eta_fe25519_64W (snd xy)))
+ (Hxy : is_bounded (fe25519_64WToZ (fst xy)) = true
+ /\ is_bounded (fe25519_64WToZ (snd xy)) = true),
+ let Hx := let (Hx, Hy) := Hxy in Hx in
+ 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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall xy
+ (xy := (eta_fe25519_64W (fst xy), eta_fe25519_64W (snd xy)))
+ (Hxy : is_bounded (fe25519_64WToZ (fst xy)) = true
+ /\ is_bounded (fe25519_64WToZ (snd xy)) = true),
+ 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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => binop_bounds_good bounds = true
+ | None => False
+ end)
+ : binop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x y Hx Hy.
+ pose x as x'; pose y as y'.
+ hnf in x, y; destruct_head' prod.
+ specialize (H0 (x', y') (conj Hx Hy)).
+ specialize (H1 (x', y') (conj Hx Hy)).
+ let args := constr:(binop_args_to_bounded (x', y') Hx Hy) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
new file mode 100644
index 000000000..bfc3c3cc1
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe25519_64W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unop_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
new file mode 100644
index 000000000..f615067c6
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToWire_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe25519_64W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToWire_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToWire_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
new file mode 100644
index 000000000..079a0729c
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToZ_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe25519_64W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToZ_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToZ_correct (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
new file mode 100644
index 000000000..c58f70e91
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Common.
+Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpWireToFE_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_wire_digitsW x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopWireToFE_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_WireToFE_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified.v b/src/SpecificGen/GF25519_64Reflective/Reified.v
new file mode 100644
index 000000000..a5236884e
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified.v
@@ -0,0 +1,13 @@
+(** We split the reification up into separate files, one operation per
+ file, so that it can run in parallel. *)
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.Add.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.CarryAdd.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.Sub.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.CarrySub.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.Mul.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.Opp.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.CarryOpp.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.PreFreeze.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.GeModulus.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.Pack.
+Require Export Crypto.SpecificGen.GF25519_64Reflective.Reified.Unpack.
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/Add.v b/src/SpecificGen/GF25519_64Reflective/Reified/Add.v
new file mode 100644
index 000000000..1adf81d96
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/Add.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonBinOp.
+
+Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
+Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
+Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
+Proof. rexpr_correct. Qed.
+Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
+Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
+Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/CarryAdd.v b/src/SpecificGen/GF25519_64Reflective/Reified/CarryAdd.v
new file mode 100644
index 000000000..3a6871acc
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/CarryAdd.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonBinOp.
+
+Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
+Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
+Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_addW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/CarryOpp.v b/src/SpecificGen/GF25519_64Reflective/Reified/CarryOpp.v
new file mode 100644
index 000000000..bd584ff2e
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/CarryOpp.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOp.
+
+Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
+Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
+Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_oppW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/CarrySub.v b/src/SpecificGen/GF25519_64Reflective/Reified/CarrySub.v
new file mode 100644
index 000000000..5b7bda4e9
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/CarrySub.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonBinOp.
+
+Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
+Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
+Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_subW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/GeModulus.v b/src/SpecificGen/GF25519_64Reflective/Reified/GeModulus.v
new file mode 100644
index 000000000..e755a2238
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/GeModulus.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOpFEToZ.
+
+Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
+Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
+Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rge_modulusW_correct_and_bounded
+ := ExprUnOpFEToZ_correct_and_bounded
+ rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/Mul.v b/src/SpecificGen/GF25519_64Reflective/Reified/Mul.v
new file mode 100644
index 000000000..ba589a18d
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/Mul.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonBinOp.
+
+Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
+Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
+Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rmulW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
+Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
+Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/Opp.v b/src/SpecificGen/GF25519_64Reflective/Reified/Opp.v
new file mode 100644
index 000000000..6170329d5
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/Opp.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOp.
+
+Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
+Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
+Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
+Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
+Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/Pack.v b/src/SpecificGen/GF25519_64Reflective/Reified/Pack.v
new file mode 100644
index 000000000..c73b348ab
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/Pack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOpFEToWire.
+
+Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
+Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
+Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rpackW_correct_and_bounded
+ := ExprUnOpFEToWire_correct_and_bounded
+ rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v
new file mode 100644
index 000000000..457ffa715
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOp.
+
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF25519_64.prefreeze]. reify_sig. Defined.
+Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
+Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rprefreezeW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/Sub.v b/src/SpecificGen/GF25519_64Reflective/Reified/Sub.v
new file mode 100644
index 000000000..a0deda793
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/Sub.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonBinOp.
+
+Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
+Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
+Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
+Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
+Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/Unpack.v b/src/SpecificGen/GF25519_64Reflective/Reified/Unpack.v
new file mode 100644
index 000000000..2e1a40f31
--- /dev/null
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/Unpack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOpWireToFE.
+
+Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
+Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
+Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition runpackW_correct_and_bounded
+ := ExprUnOpWireToFE_correct_and_bounded
+ runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/SpecificGen/GF41417_32Bounded.v b/src/SpecificGen/GF41417_32Bounded.v
new file mode 100644
index 000000000..af6a28421
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Bounded.v
@@ -0,0 +1,462 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF41417_32.
+Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
+Require Import Crypto.SpecificGen.GF41417_32Reflective.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+
+Local Ltac bounded_t opW blem :=
+ apply blem; apply is_bounded_proj1_fe41417_32.
+Local Ltac bounded_wire_digits_t opW blem :=
+ apply blem; apply is_bounded_proj1_wire_digits.
+
+Local Ltac define_binop f g opW blem :=
+ refine (exist_fe41417_32W (opW (proj1_fe41417_32W f) (proj1_fe41417_32W g)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop f opW blem :=
+ refine (exist_fe41417_32W (opW (proj1_fe41417_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_FEToZ f opW :=
+ refine (opW (proj1_fe41417_32W f)).
+Local Ltac define_unop_FEToWire f opW blem :=
+ refine (exist_wire_digitsW (opW (proj1_fe41417_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_WireToFE f opW blem :=
+ refine (exist_fe41417_32W (opW (proj1_wire_digitsW f)) _);
+ 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 Arguments interp_radd / _ _.
+Local Arguments interp_rsub / _ _.
+Local Arguments interp_rmul / _ _.
+Local Arguments interp_ropp / _.
+Local Arguments interp_rprefreeze / _.
+Local Arguments interp_rge_modulus / _.
+Local Arguments interp_rpack / _.
+Local Arguments interp_runpack / _.
+Definition addW (f g : fe41417_32W) : fe41417_32W := Eval simpl in interp_radd f g.
+Definition subW (f g : fe41417_32W) : fe41417_32W := Eval simpl in interp_rsub f g.
+Definition mulW (f g : fe41417_32W) : fe41417_32W := Eval simpl in interp_rmul f g.
+Definition oppW (f : fe41417_32W) : fe41417_32W := Eval simpl in interp_ropp f.
+Definition prefreezeW (f : fe41417_32W) : fe41417_32W := Eval simpl in interp_rprefreeze f.
+Definition ge_modulusW (f : fe41417_32W) : word64 := Eval simpl in interp_rge_modulus f.
+Definition packW (f : fe41417_32W) : wire_digitsW := Eval simpl in interp_rpack f.
+Definition unpackW (f : wire_digitsW) : fe41417_32W := Eval simpl in interp_runpack f.
+
+Require Import ModularBaseSystemWord.
+Definition modulusW :=
+ Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z 10 GF41417_32.modulus_digits_)).
+
+Definition postfreeze : GF41417_32.fe41417_32 -> GF41417_32.fe41417_32 :=
+ GF41417_32.postfreeze.
+
+Lemma freeze_prepost_freeze : forall x, postfreeze (prefreeze x) = GF41417_32.freeze x. Proof. reflexivity. Qed.
+
+Definition postfreezeW : fe41417_32W -> fe41417_32W :=
+ (conditional_subtract_modulusW
+ (num_limbs := 10)
+ modulusW
+ ge_modulusW
+ (Interpretations.Word64.neg GF41417_32.int_width)
+ ).
+
+Definition freezeW (f : fe41417_32W) : fe41417_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
+
+Local Transparent Let_In.
+Definition powW (f : fe41417_32W) chain := fold_chain_opt (proj1_fe41417_32W one) mulW chain [f].
+Definition invW (f : fe41417_32W) : fe41417_32W
+ := Eval cbv -[Let_In fe41417_32W mulW] in powW f (chain inv_ec).
+
+Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
+ change opW with (interp_rop);
+ rewrite pre_rewrite;
+ intros; apply rop_cb; assumption.
+
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Proof. port_correct_and_bounded interp_radd_correct addW interp_radd radd_correct_and_bounded. Qed.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW carry_sub.
+Proof. port_correct_and_bounded interp_rsub_correct subW interp_rsub rsub_correct_and_bounded. Qed.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW mul.
+Proof. port_correct_and_bounded interp_rmul_correct mulW interp_rmul rmul_correct_and_bounded. Qed.
+Lemma oppW_correct_and_bounded : iunop_correct_and_bounded oppW carry_opp.
+Proof. port_correct_and_bounded interp_ropp_correct oppW interp_ropp ropp_correct_and_bounded. Qed.
+Lemma prefreezeW_correct_and_bounded : iunop_correct_and_bounded prefreezeW prefreeze.
+Proof. port_correct_and_bounded interp_rprefreeze_correct prefreezeW interp_rprefreeze rprefreeze_correct_and_bounded. Qed.
+Lemma ge_modulusW_correct : iunop_FEToZ_correct ge_modulusW ge_modulus.
+Proof. port_correct_and_bounded interp_rge_modulus_correct ge_modulusW interp_rge_modulus rge_modulus_correct_and_bounded. Qed.
+Lemma packW_correct_and_bounded : iunop_FEToWire_correct_and_bounded packW pack.
+Proof. port_correct_and_bounded interp_rpack_correct packW interp_rpack rpack_correct_and_bounded. Qed.
+Lemma unpackW_correct_and_bounded : iunop_WireToFE_correct_and_bounded unpackW unpack.
+Proof. port_correct_and_bounded interp_runpack_correct unpackW interp_runpack runpack_correct_and_bounded. Qed.
+
+(* TODO : move *)
+Lemma neg_range : forall x y, 0 <= x ->
+ 0 <= ModularBaseSystemListZOperations.neg x y < 2 ^ x.
+Proof.
+ intros.
+ split; auto using ModularBaseSystemListZOperationsProofs.neg_nonneg.
+ eapply Z.le_lt_trans; eauto using ModularBaseSystemListZOperationsProofs.neg_upperbound.
+ rewrite Z.ones_equiv.
+ omega.
+Qed.
+
+Ltac lower_bound_minus_ge_modulus :=
+ apply Z.le_0_sub;
+ 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;
+ etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
+ apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
+
+Ltac upper_bound_minus_ge_modulus :=
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | ];
+ eapply Z.le_lt_trans; [ eassumption | ];
+ instantiate; vm_compute; reflexivity.
+
+Lemma postfreezeW_correct_and_bounded : iunop_correct_and_bounded postfreezeW postfreeze.
+Proof.
+ intros x H.
+ pose proof (ge_modulusW_correct x H) as Hgm.
+ destruct_head_hnf' prod.
+ unfold_is_bounded_in H.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ cbv [postfreezeW].
+ cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
+ change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ rewrite Hgm.
+
+ cbv [modulusW Tuple.map].
+ cbv [on_tuple List.map to_list to_list' from_list from_list'
+ Tuple.map2 on_tuple2 ListUtil.map2 fe41417_32WToZ].
+ cbv [postfreeze GF41417_32.postfreeze].
+ cbv [Let_In].
+
+ split.
+ { match goal with
+ |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?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;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ | |- (_,_) = (_,_) => reflexivity
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ end. }
+
+
+ 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;
+ repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ try match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ | |- _ - _ <= _ => etransitivity; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | instantiate; assumption ]
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- ModularBaseSystemListZOperations.neg _ _ < _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- _ => vm_compute; (discriminate || reflexivity)
+ end.
+Qed.
+
+Lemma freezeW_correct_and_bounded : iunop_correct_and_bounded freezeW freeze.
+Proof.
+ intros f H; rewrite <- freeze_prepost_freeze.
+ change (freezeW f) with (postfreezeW (prefreezeW f)).
+ destruct (prefreezeW_correct_and_bounded f H) as [H0 H1].
+ destruct (postfreezeW_correct_and_bounded _ H1) as [H0' H1'].
+ rewrite H1', H0', H0; split; reflexivity.
+Qed.
+
+Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
+Proof.
+ cbv [powW pow].
+ intro x; intros; apply (fold_chain_opt_gen fe41417_32WToZ is_bounded [x]).
+ { reflexivity. }
+ { reflexivity. }
+ { intros; progress rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption.
+ apply mulW_correct_and_bounded; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption; reflexivity. }
+ { intros [|?]; autorewrite with simpl_nth_default;
+ (assumption || reflexivity). }
+Qed.
+
+Lemma invW_correct_and_bounded : iunop_correct_and_bounded invW inv.
+Proof.
+ intro f.
+ assert (H : forall f, invW f = powW f (chain inv_ec))
+ by abstract (cbv -[Let_In fe41417_32W mulW]; reflexivity).
+ rewrite H.
+ rewrite inv_correct.
+ cbv [inv_opt].
+ rewrite <- pow_correct.
+ apply powW_correct_and_bounded.
+Qed.
+
+Definition fieldwisebW_sig (f g : fe41417_32W)
+ : { b | b = GF41417_32.fieldwiseb (fe41417_32WToZ f) (fe41417_32WToZ g) }.
+Proof.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ cbv [GF41417_32.fieldwiseb fe41417_32WToZ].
+ rewrite ?word64eqb_Zeqb.
+ reflexivity.
+Defined.
+
+Definition fieldwisebW (f g : fe41417_32W) : bool :=
+ Eval cbv [proj1_sig fieldwisebW_sig appify2 app_fe41417_32W] in
+ appify2 (fun f g => proj1_sig (fieldwisebW_sig f g)) f g.
+
+Lemma fieldwisebW_correct f g
+ : fieldwisebW f g = GF41417_32.fieldwiseb (fe41417_32WToZ f) (fe41417_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (fieldwisebW_sig f' g')).
+Qed.
+
+Local Arguments freezeW : simpl never.
+Local Arguments fe41417_32WToZ !_ / .
+Local Opaque freezeW.
+
+Definition eqbW_sig (f g : fe41417_32W)
+ : { b | is_bounded (fe41417_32WToZ f) = true
+ -> is_bounded (fe41417_32WToZ g) = true
+ -> b = GF41417_32.eqb (fe41417_32WToZ f) (fe41417_32WToZ g) }.
+Proof.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded f pf)) as frf.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded g pf)) as frg.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ unfold GF41417_32.eqb.
+ simpl @fe41417_32WToZ in *; cbv beta iota.
+ intros.
+ rewrite <- frf, <- frg by assumption.
+ rewrite <- fieldwisebW_correct.
+ reflexivity.
+Defined.
+
+Definition eqbW (f g : fe41417_32W) : bool :=
+ Eval cbv [proj1_sig eqbW_sig appify2 app_fe41417_32W] in
+ appify2 (fun f g => proj1_sig (eqbW_sig f g)) f g.
+
+Lemma eqbW_correct f g
+ : is_bounded (fe41417_32WToZ f) = true
+ -> is_bounded (fe41417_32WToZ g) = true
+ -> eqbW f g = GF41417_32.eqb (fe41417_32WToZ f) (fe41417_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (eqbW_sig f' g')).
+Qed.
+
+(* TODO(jgross): use NToWord or such for this constant too *)
+Definition sqrt_m1W : fe41417_32W :=
+ Eval vm_compute in fe41417_32ZToW sqrt_m1.
+
+Definition GF41417_32sqrt x : GF41417_32.fe41417_32 :=
+ dlet powx := powW (fe41417_32ZToW x) (chain GF41417_32.sqrt_ec) in
+ GF41417_32.sqrt (fe41417_32WToZ powx) (fe41417_32WToZ (mulW powx powx)) x.
+
+Definition sqrtW_sig
+ : { sqrtW | iunop_correct_and_bounded sqrtW GF41417_32sqrt }.
+Proof.
+ eexists.
+ unfold GF41417_32sqrt, GF41417_32.sqrt.
+ intros.
+ rewrite !fe41417_32ZToW_WToZ.
+ split.
+ { etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody_eq; intros.
+ set_evars.
+ match goal with (* unfold the first dlet ... in, but only if it's binding a var *)
+ | [ |- ?x = dlet y := fe41417_32WToZ ?z in ?f ]
+ => is_var z; change (x = match fe41417_32WToZ z with y => f end)
+ end.
+ change sqrt_m1 with (fe41417_32WToZ sqrt_m1W).
+ rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded sqrt_m1W a X Y)), <- eqbW_correct, (pull_bool_if fe41417_32WToZ)
+ by repeat match goal with
+ | _ => progress subst
+ | [ |- is_bounded (fe41417_32WToZ ?op) = true ]
+ => lazymatch op with
+ | mulW _ _ => apply mulW_correct_and_bounded
+ | powW _ _ => apply powW_correct_and_bounded
+ | sqrt_m1W => vm_compute; reflexivity
+ | _ => assumption
+ end
+ end.
+ subst_evars; reflexivity.
+ } Unfocus.
+ lazymatch goal with
+ | [ |- context G[dlet x := ?v in fe41417_32WToZ (@?f x)] ]
+ => let G' := context G[fe41417_32WToZ (dlet x := v in f x)] in
+ cut G'; cbv beta;
+ [ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
+ end.
+ reflexivity. }
+ { cbv [Let_In].
+ break_if.
+ { apply powW_correct_and_bounded; assumption. }
+ { apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ].
+ apply powW_correct_and_bounded; assumption. } }
+Defined.
+
+Definition sqrtW (f : fe41417_32W) : fe41417_32W :=
+ Eval cbv [proj1_sig sqrtW_sig app_fe41417_32W] in
+ app_fe41417_32W f (proj1_sig sqrtW_sig).
+
+Lemma sqrtW_correct_and_bounded : iunop_correct_and_bounded sqrtW GF41417_32sqrt.
+Proof.
+ intro f.
+ set (f' := f).
+ hnf in f; destruct_head' prod.
+ assert (H : sqrtW f' = proj1_sig sqrtW_sig f')
+ by (subst f'; cbv beta iota delta [proj1_sig sqrtW_sig sqrtW]; reflexivity).
+ rewrite H.
+ exact (proj2_sig sqrtW_sig f').
+Qed.
+
+
+
+Definition add (f g : fe41417_32) : fe41417_32.
+Proof. define_binop f g addW addW_correct_and_bounded. Defined.
+Definition sub (f g : fe41417_32) : fe41417_32.
+Proof. define_binop f g subW subW_correct_and_bounded. Defined.
+Definition mul (f g : fe41417_32) : fe41417_32.
+Proof. define_binop f g mulW mulW_correct_and_bounded. Defined.
+Definition opp (f : fe41417_32) : fe41417_32.
+Proof. define_unop f oppW oppW_correct_and_bounded. Defined.
+Definition freeze (f : fe41417_32) : fe41417_32.
+Proof. define_unop f freezeW freezeW_correct_and_bounded. Defined.
+Definition ge_modulus (f : fe41417_32) : word64.
+Proof. define_unop_FEToZ f ge_modulusW. Defined.
+Definition pack (f : fe41417_32) : wire_digits.
+Proof. define_unop_FEToWire f packW packW_correct_and_bounded. Defined.
+Definition unpack (f : wire_digits) : fe41417_32.
+Proof. define_unop_WireToFE f unpackW unpackW_correct_and_bounded. Defined.
+
+Definition pow (f : fe41417_32) (chain : list (nat * nat)) : fe41417_32.
+Proof. define_unop f (fun x => powW x chain) powW_correct_and_bounded. Defined.
+Definition inv (f : fe41417_32) : fe41417_32.
+Proof. define_unop f invW invW_correct_and_bounded. Defined.
+Definition sqrt (f : fe41417_32) : fe41417_32.
+Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
+
+Local Ltac op_correct_t op opW_correct_and_bounded :=
+ cbv [op];
+ lazymatch goal with
+ | [ |- context[proj1_fe41417_32 (exist_fe41417_32W _ _)] ]
+ => rewrite proj1_fe41417_32_exist_fe41417_32W
+ | [ |- context[proj1_wire_digits (exist_wire_digitsW _ _)] ]
+ => rewrite proj1_wire_digits_exist_wire_digitsW
+ | _ => idtac
+ end;
+ apply opW_correct_and_bounded;
+ lazymatch goal with
+ | [ |- is_bounded _ = true ]
+ => apply is_bounded_proj1_fe41417_32
+ | [ |- wire_digits_is_bounded _ = true ]
+ => apply is_bounded_proj1_wire_digits
+ end.
+
+Lemma add_correct (f g : fe41417_32) : proj1_fe41417_32 (add f g) = carry_add (proj1_fe41417_32 f) (proj1_fe41417_32 g).
+Proof. op_correct_t add addW_correct_and_bounded. Qed.
+Lemma sub_correct (f g : fe41417_32) : proj1_fe41417_32 (sub f g) = carry_sub (proj1_fe41417_32 f) (proj1_fe41417_32 g).
+Proof. op_correct_t sub subW_correct_and_bounded. Qed.
+Lemma mul_correct (f g : fe41417_32) : proj1_fe41417_32 (mul f g) = GF41417_32.mul (proj1_fe41417_32 f) (proj1_fe41417_32 g).
+Proof. op_correct_t mul mulW_correct_and_bounded. Qed.
+Lemma opp_correct (f : fe41417_32) : proj1_fe41417_32 (opp f) = carry_opp (proj1_fe41417_32 f).
+Proof. op_correct_t opp oppW_correct_and_bounded. Qed.
+Lemma freeze_correct (f : fe41417_32) : proj1_fe41417_32 (freeze f) = GF41417_32.freeze (proj1_fe41417_32 f).
+Proof. op_correct_t freeze freezeW_correct_and_bounded. Qed.
+Lemma ge_modulus_correct (f : fe41417_32) : word64ToZ (ge_modulus f) = GF41417_32.ge_modulus (proj1_fe41417_32 f).
+Proof. op_correct_t ge_modulus ge_modulusW_correct. Qed.
+Lemma pack_correct (f : fe41417_32) : proj1_wire_digits (pack f) = GF41417_32.pack (proj1_fe41417_32 f).
+Proof. op_correct_t pack packW_correct_and_bounded. Qed.
+Lemma unpack_correct (f : wire_digits) : proj1_fe41417_32 (unpack f) = GF41417_32.unpack (proj1_wire_digits f).
+Proof. op_correct_t unpack unpackW_correct_and_bounded. Qed.
+Lemma pow_correct (f : fe41417_32) chain : proj1_fe41417_32 (pow f chain) = GF41417_32.pow (proj1_fe41417_32 f) chain.
+Proof. op_correct_t pow (powW_correct_and_bounded chain). Qed.
+Lemma inv_correct (f : fe41417_32) : proj1_fe41417_32 (inv f) = GF41417_32.inv (proj1_fe41417_32 f).
+Proof. op_correct_t inv invW_correct_and_bounded. Qed.
+Lemma sqrt_correct (f : fe41417_32) : proj1_fe41417_32 (sqrt f) = GF41417_32sqrt (proj1_fe41417_32 f).
+Proof. op_correct_t sqrt sqrtW_correct_and_bounded. Qed.
+
+Import Morphisms.
+
+Local Existing Instance prime_modulus.
+
+Lemma field41417_32_and_homomorphisms
+ : @field fe41417_32 eq zero one opp add sub mul inv div
+ /\ @Ring.is_homomorphism (F _) (@Logic.eq _) 1%F F.add F.mul fe41417_32 eq one add mul encode
+ /\ @Ring.is_homomorphism fe41417_32 eq one add mul (F _) (@Logic.eq _) 1%F F.add F.mul decode.
+Proof.
+ eapply @Field.field_and_homomorphism_from_redundant_representation.
+ { exact (F.field_modulo _). }
+ { cbv [decode encode]; intros; rewrite !proj1_fe41417_32_exist_fe41417_32; apply encode_rep. }
+ { reflexivity. }
+ { reflexivity. }
+ { reflexivity. }
+ { cbv [decode encode]; intros; rewrite opp_correct, carry_opp_correct, carry_opp_opt_correct, carry_opp_rep; apply opp_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite add_correct, carry_add_correct, carry_add_opt_correct, carry_add_rep; apply add_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite sub_correct, carry_sub_correct, carry_sub_opt_correct, carry_sub_rep; apply sub_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite mul_correct, GF41417_32.mul_correct, carry_mul_opt_correct by reflexivity; apply carry_mul_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite inv_correct, GF41417_32.inv_correct, inv_opt_correct by reflexivity; apply inv_rep; reflexivity. }
+ { cbv [decode encode div]; intros; rewrite !proj1_fe41417_32_exist_fe41417_32; apply encode_rep. }
+Qed.
+
+Global Instance field41417_32 : @field fe41417_32 eq zero one opp add sub mul inv div := proj1 field41417_32_and_homomorphisms.
+
+Local Opaque proj1_fe41417_32 exist_fe41417_32 proj1_fe41417_32W exist_fe41417_32W.
+Global Instance homomorphism_F41417_32_encode
+ : @Ring.is_homomorphism (F modulus) Logic.eq F.one F.add F.mul fe41417_32 eq one add mul encode.
+Proof. apply field41417_32_and_homomorphisms. Qed.
+
+Global Instance homomorphism_F41417_32_decode
+ : @Ring.is_homomorphism fe41417_32 eq one add mul (F modulus) Logic.eq F.one F.add F.mul decode.
+Proof. apply field41417_32_and_homomorphisms. Qed.
diff --git a/src/SpecificGen/GF41417_32BoundedCommon.v b/src/SpecificGen/GF41417_32BoundedCommon.v
new file mode 100644
index 000000000..2f87412c9
--- /dev/null
+++ b/src/SpecificGen/GF41417_32BoundedCommon.v
@@ -0,0 +1,693 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF41417_32.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.HList.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+(* BEGIN common curve-specific definitions *)
+Definition bit_width : nat := Eval compute in Z.to_nat (GF41417_32.int_width).
+Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
+Definition bounds_exp : tuple Z length_fe41417_32
+ := Eval compute in
+ Tuple.from_list length_fe41417_32 limb_widths eq_refl.
+Definition bounds : tuple (Z * Z) length_fe41417_32
+ := Eval compute in
+ Tuple.map (fun e => b_of e) bounds_exp.
+Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
+ := Eval compute in Tuple.from_list _ wire_widths eq_refl.
+Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
+ := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
+(* 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.
+Arguments bit_width : simpl never.
+Global Opaque word64.
+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},
+ 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},
+ 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
+ := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
+ Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ : 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}
+ : 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 }.
+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 fe41417_32W := Eval cbv (*-[word64]*) in (tuple word64 length_fe41417_32).
+Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
+Definition fe41417_32 :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => word_of e) bounds_exp.
+Definition wire_digits :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
+
+Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
+ := let res := Tuple.map2
+ (fun bounds v =>
+ let '(lower, upper) := bounds in
+ (lower <=? v) && (v <=? upper))%bool%Z
+ bounds x in
+ List.fold_right andb true (Tuple.to_list _ res).
+
+Definition is_bounded (x : SpecificGen.GF41417_32.fe41417_32) : bool
+ := is_bounded_gen (n:=length_fe41417_32) x bounds.
+
+Definition wire_digits_is_bounded (x : SpecificGen.GF41417_32.wire_digits) : bool
+ := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
+
+(* END basic types *)
+
+Section generic_destructuring.
+ Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
+ := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
+ end.
+ Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
+ := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => @app_on' A n'
+ end.
+ Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
+ Qed.
+ Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
+
+ Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
+ end.
+ Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => @app_on_h' A F n'
+ end.
+ Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
+ Qed.
+ Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
+
+ Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
+ : forall (f : tuple A (length wire_widths)), T f
+ := Eval compute in fun f => @app_on A (length wire_widths) T f P.
+ Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
+ := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
+ Definition app_fe41417_32W_dep {A T} (P : forall x : tuple A length_fe41417_32, T x)
+ : forall (f : tuple A length_fe41417_32), T f
+ := Eval compute in fun f => @app_on A length_fe41417_32 T f P.
+ Definition app_fe41417_32W {A T} (f : tuple A length_fe41417_32) (P : tuple A length_fe41417_32 -> T)
+ := Eval compute in @app_fe41417_32W_dep A (fun _ => T) P f.
+ Definition app_fe41417_32_dep {T} (P : forall x : fe41417_32, T x)
+ : forall f : fe41417_32, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe41417_32 bounds_exp T f P.
+ Definition app_fe41417_32 {T} (f : fe41417_32) (P : hlist (fun e => word_of e) bounds_exp -> T)
+ := Eval compute in @app_fe41417_32_dep (fun _ => T) P f.
+ Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
+ : forall f : wire_digits, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
+ Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
+ := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
+
+ Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
+ := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
+ Definition app_fe41417_32W_dep_correct {A T} f P : @app_fe41417_32W_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_fe41417_32W_correct {A T} f P : @app_fe41417_32W A T f P = P f
+ := @app_fe41417_32W_dep_correct A (fun _ => T) f P.
+ Definition app_fe41417_32_dep_correct {T} f P : @app_fe41417_32_dep T P f = P f
+ := app_on_h_correct (fun e => word_of e) bounds_exp f P.
+ Definition app_fe41417_32_correct {T} f P : @app_fe41417_32 T f P = P f
+ := @app_fe41417_32_dep_correct (fun _ => T) f P.
+ Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
+ := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
+ Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
+ := @app_wire_digits_dep_correct (fun _ => T) f P.
+
+ Definition appify2 {T} (op : fe41417_32W -> fe41417_32W -> T) (f g : fe41417_32W) :=
+ app_fe41417_32W f (fun f0 => (app_fe41417_32W g (fun g0 => op f0 g0))).
+
+ Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
+ Proof.
+ intros. cbv [appify2].
+ etransitivity; apply app_fe41417_32W_correct.
+ Qed.
+End generic_destructuring.
+
+Definition eta_fe41417_32W_sig (x : fe41417_32W) : { v : fe41417_32W | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+Defined.
+Definition eta_fe41417_32W (x : fe41417_32W) : fe41417_32W
+ := Eval cbv [proj1_sig eta_fe41417_32W_sig] in proj1_sig (eta_fe41417_32W_sig x).
+Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+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.
+Proof. apply NToWord_wordToN. Qed.
+Local Opaque word64.
+
+Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ rewrite N2Z.id, NToWord_wordToN; reflexivity.
+Qed.
+
+(* BEGIN precomputation. *)
+
+Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
+Definition word_bounded {lower upper} (v : bounded_word lower upper)
+ : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
+ := project_is_boundedT (proj2_sig v).
+Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
+ := exist _ proj_word (Build_is_boundedT word_bounded).
+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)
+ : bounded_word lower upper
+ := Build_bounded_word'
+ proj_word
+ (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
+ | 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.
+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.
+ match goal with
+ | [ |- context[andb ?x ?y] ]
+ => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
+ end;
+ intros; omega.
+Qed.
+Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
+Proof.
+ refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
+ abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
+Defined.
+Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
+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.
+Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe41417_32W app_fe41417_32 HList.mapt HList.mapt' Tuple.map on_tuple List.map List.app length_fe41417_32 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
+Definition fe41417_32WToZ (x : fe41417_32W) : SpecificGen.GF41417_32.fe41417_32
+ := Eval app_tuple_map in
+ app_fe41417_32W x (Tuple.map (fun v : word64 => v : Z)).
+Definition fe41417_32ZToW (x : SpecificGen.GF41417_32.fe41417_32) : fe41417_32W
+ := Eval app_tuple_map in
+ app_fe41417_32W x (Tuple.map (fun v : Z => v : word64)).
+Definition wire_digitsWToZ (x : wire_digitsW) : SpecificGen.GF41417_32.wire_digits
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
+Definition wire_digitsZToW (x : SpecificGen.GF41417_32.wire_digits) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
+Definition fe41417_32W_word64ize (x : fe41417_32W) : fe41417_32W
+ := Eval app_tuple_map in
+ app_fe41417_32W x (Tuple.map word64ize).
+Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map word64ize).
+
+(** TODO: Turn this into a lemma to speed up proofs *)
+Ltac unfold_is_bounded_in H :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe41417_32WToZ, wire_digitsWToZ in H;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe41417_32 List.length wire_widths] in H;
+ rewrite ?Bool.andb_true_iff in H.
+
+Ltac unfold_is_bounded :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe41417_32WToZ, wire_digitsWToZ;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe41417_32 List.length wire_widths];
+ 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.
+Local Opaque bit_width.
+
+Local Ltac prove_lt_bit_width :=
+ rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
+
+Lemma fe41417_32ZToW_WToZ (x : fe41417_32W) : fe41417_32ZToW (fe41417_32WToZ x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe41417_32WToZ fe41417_32ZToW].
+ rewrite !ZToWord64_word64ToZ; reflexivity.
+Qed.
+
+Lemma fe41417_32WToZ_ZToW x : is_bounded x = true -> fe41417_32WToZ (fe41417_32ZToW x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe41417_32WToZ fe41417_32ZToW].
+ intro H.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma fe41417_32W_word64ize_id x : fe41417_32W_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [fe41417_32W_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [wire_digitsW_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+
+Definition uncurry_unop_fe41417_32W {T} (op : fe41417_32W -> T)
+ := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe41417_32) op.
+Definition curry_unop_fe41417_32W {T} op : fe41417_32W -> T
+ := Eval cbv (*-[word64]*) in fun f => app_fe41417_32W f (Tuple.curry (n:=length_fe41417_32) op).
+Definition uncurry_binop_fe41417_32W {T} (op : fe41417_32W -> fe41417_32W -> T)
+ := Eval cbv (*-[word64]*) in uncurry_unop_fe41417_32W (fun f => uncurry_unop_fe41417_32W (op f)).
+Definition curry_binop_fe41417_32W {T} op : fe41417_32W -> fe41417_32W -> T
+ := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe41417_32W (curry_unop_fe41417_32W op f)).
+
+Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
+ := Eval cbv (*-[word64]*) 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).
+
+
+Definition proj1_fe41417_32W (x : fe41417_32) : fe41417_32W
+ := Eval app_tuple_map in
+ app_fe41417_32 x (HList.mapt (fun _ => (@proj_word _ _))).
+Coercion proj1_fe41417_32 (x : fe41417_32) : SpecificGen.GF41417_32.fe41417_32
+ := fe41417_32WToZ (proj1_fe41417_32W x).
+
+Lemma is_bounded_proj1_fe41417_32 (x : fe41417_32) : is_bounded (proj1_fe41417_32 x) = true.
+Proof.
+ revert x; refine (app_fe41417_32_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [is_bounded proj1_fe41417_32 proj1_fe41417_32W fe41417_32WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe41417_32 is_bounded_gen].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
+ := app_wire_digits x (HList.mapt (fun _ => proj_word)).
+Coercion proj1_wire_digits (x : wire_digits) : SpecificGen.GF41417_32.wire_digits
+ := wire_digitsWToZ (proj1_wire_digitsW x).
+
+Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
+Proof.
+ revert x; refine (app_wire_digits_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Local Ltac make_exist_W' x app_W_dep :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x);
+ cbv [tuple tuple' length_fe41417_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v H :=
+ first [ let v' := (eval cbv [snd fst] in (snd v)) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
+ | let v' := (eval cbv [snd fst] in v) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
+ let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
+ let T := type of H' in
+ let T := (eval cbv [id
+ List.fold_right List.map List.length List.app ListUtil.map2 List.rev
+ Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
+ fe41417_32 bounds fe41417_32WToZ length_fe41417_32
+ wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
+ let H' := constr:(H' : T) in
+ let v := (eval unfold x' in x') in
+ do_refine v H'.
+Local Ltac make_exist'' x exist_W ZToW :=
+ let H := fresh in
+ intro H; apply (exist_W (ZToW x));
+ abstract (
+ hnf in x; destruct_head' prod;
+ let H' := fresh in
+ pose proof H as H';
+ unfold_is_bounded_in H;
+ destruct_head' and; simpl in *;
+ Z.ltb_to_lt;
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ assumption
+ ).
+Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x) in *;
+ cbv [tuple tuple' length_fe41417_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v :=
+ first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
+ | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
+ let v := (eval unfold x' in (exist'' x' H)) in
+ do_refine v.
+
+Definition exist_fe41417_32W' (x : fe41417_32W) : is_bounded (fe41417_32WToZ x) = true -> fe41417_32.
+Proof. make_exist_W' x (@app_fe41417_32W_dep). Defined.
+Definition exist_fe41417_32W (x : fe41417_32W) : is_bounded (fe41417_32WToZ x) = true -> fe41417_32
+ := Eval cbv [app_fe41417_32W_dep exist_fe41417_32W' fe41417_32ZToW] in exist_fe41417_32W' x.
+Definition exist_fe41417_32'' (x : SpecificGen.GF41417_32.fe41417_32) : is_bounded x = true -> fe41417_32.
+Proof. make_exist'' x exist_fe41417_32W fe41417_32ZToW. Defined.
+Definition exist_fe41417_32' (x : SpecificGen.GF41417_32.fe41417_32) : is_bounded x = true -> fe41417_32.
+Proof. make_exist' x (@app_fe41417_32W_dep) exist_fe41417_32'' exist_fe41417_32W fe41417_32ZToW. Defined.
+Definition exist_fe41417_32 (x : SpecificGen.GF41417_32.fe41417_32) : is_bounded x = true -> fe41417_32
+ := Eval cbv [exist_fe41417_32' exist_fe41417_32W exist_fe41417_32' app_fe41417_32 app_fe41417_32W_dep] in
+ exist_fe41417_32' x.
+
+Lemma proj1_fe41417_32_exist_fe41417_32W x pf : proj1_fe41417_32 (exist_fe41417_32W x pf) = fe41417_32WToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe41417_32W_exist_fe41417_32 x pf : proj1_fe41417_32W (exist_fe41417_32 x pf) = fe41417_32ZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe41417_32_exist_fe41417_32 x pf : proj1_fe41417_32 (exist_fe41417_32 x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_fe41417_32 exist_fe41417_32 proj1_fe41417_32W fe41417_32WToZ proj_word Build_bounded_word Build_bounded_word'].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition exist_wire_digitsW' (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
+Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
+Definition exist_wire_digitsW (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
+ := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
+Definition exist_wire_digits'' (x : SpecificGen.GF41417_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits' (x : SpecificGen.GF41417_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits (x : SpecificGen.GF41417_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits
+ := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
+ exist_wire_digits' x.
+
+Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 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 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.
+End opt.
+
+Local Transparent bit_width.
+Local Ltac do_change lem :=
+ match lem with
+ | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
+ => 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
+ let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
+ do_change L'
+ end
+ | _ => lem
+ end.
+Definition fe41417_32_word64ize (x : fe41417_32) : fe41417_32.
+Proof.
+ set (x' := x).
+ hnf in x; destruct_head' prod.
+ let lem := constr:(exist_fe41417_32W (fe41417_32W_word64ize (proj1_fe41417_32W x'))) in
+ let lem := (eval cbv [proj1_fe41417_32W x' fe41417_32W_word64ize proj_word exist_fe41417_32W Build_bounded_word' Build_bounded_word] in lem) in
+ let lem := do_change lem in
+ refine (lem _);
+ change (is_bounded (fe41417_32WToZ (fe41417_32W_word64ize (proj1_fe41417_32W x'))) = true);
+ abstract (rewrite fe41417_32W_word64ize_id; apply is_bounded_proj1_fe41417_32).
+Defined.
+Definition wire_digits_word64ize (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 := do_change lem in
+ let lem := (eval cbv [word64ize opt.word64ize 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).
+Defined.
+
+Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
+ : forall n : nat,
+ (n < length limb_widths)%nat
+ -> (0 <= nth_default 0 (Tuple.to_list length_fe41417_32 x) n <=
+ snd (b_of (nth_default (-1) limb_widths n)))%Z.
+Proof.
+ hnf in x; destruct_head' prod.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ unfold nth_default; simpl.
+ intros.
+ repeat match goal with
+ | [ |- context[nth_error _ ?x] ]
+ => is_var x; destruct x; simpl
+ end;
+ omega.
+Qed.
+
+(* END precomputation *)
+
+(* Precompute constants *)
+
+Definition one' := Eval vm_compute in exist_fe41417_32 SpecificGen.GF41417_32.one_ eq_refl.
+Definition one := Eval cbv [one' fe41417_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe41417_32_word64ize one'.
+
+Definition zero' := Eval vm_compute in exist_fe41417_32 SpecificGen.GF41417_32.zero_ eq_refl.
+Definition zero := Eval cbv [zero' fe41417_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe41417_32_word64ize 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)
+ (Hid : id = F id')
+ (Hop_bounded : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> is_bounded (op (F x) (F y)) = true)
+ (Hop : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> op (F x) (F y) = F (op' x y))
+ (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
+ : F (fold_chain_opt id' op' chain ls)
+ = fold_chain_opt id op chain (List.map F ls)
+ /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
+Proof.
+ rewrite !fold_chain_opt_correct.
+ revert dependent ls; induction chain as [|x xs IHxs]; intros.
+ { pose proof (Hls_bounded 0%nat).
+ destruct ls; simpl; split; trivial; congruence. }
+ { destruct x; simpl; unfold Let_In; simpl.
+ rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
+ { do 2 f_equal.
+ rewrite <- Hop, Hid by auto.
+ rewrite !map_nth_default_always.
+ split; try reflexivity.
+ apply (IHxs (_::_)).
+ intros [|?]; autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. }
+ { intros [|?]; simpl;
+ autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. } }
+Qed.
+
+Lemma encode_bounded x : is_bounded (encode x) = true.
+Proof.
+ pose proof (bounded_encode x).
+ generalize dependent (encode x).
+ intro t; compute in t; intros.
+ destruct_head' prod.
+ unfold Pow2Base.bounded in H.
+ cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params41417_32] in H.
+ repeat match type of H with
+ | context[nth_error (cons _ _) _]
+ => let H' := fresh in
+ pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
+ cbv beta iota in H'
+ end.
+ clear H.
+ simpl in *.
+ cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
+ unfold is_bounded.
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [is_bounded proj1_fe41417_32 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe41417_32].
+ repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
+Qed.
+
+Definition encode (x : F modulus) : fe41417_32
+ := exist_fe41417_32 (encode x) (encode_bounded x).
+
+Definition decode (x : fe41417_32) : F modulus
+ := ModularBaseSystem.decode (proj1_fe41417_32 x).
+
+Lemma proj1_fe41417_32_encode x
+ : proj1_fe41417_32 (encode x) = ModularBaseSystem.encode x.
+Proof.
+ cbv [encode].
+ generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
+ intros y pf; intros; hnf in y; destruct_head_hnf' prod.
+ cbv [proj1_fe41417_32 exist_fe41417_32 proj1_fe41417_32W Build_bounded_word Build_bounded_word' fe41417_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma decode_exist_fe41417_32 x pf
+ : decode (exist_fe41417_32 x pf) = ModularBaseSystem.decode x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [decode proj1_fe41417_32 exist_fe41417_32 proj1_fe41417_32W Build_bounded_word Build_bounded_word' fe41417_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition div (f g : fe41417_32) : fe41417_32
+ := exist_fe41417_32 (div (proj1_fe41417_32 f) (proj1_fe41417_32 g)) (encode_bounded _).
+
+Definition eq (f g : fe41417_32) : Prop := eq (proj1_fe41417_32 f) (proj1_fe41417_32 g).
+
+
+Notation ibinop_correct_and_bounded irop op
+ := (forall x y,
+ is_bounded (fe41417_32WToZ x) = true
+ -> is_bounded (fe41417_32WToZ y) = true
+ -> fe41417_32WToZ (irop x y) = op (fe41417_32WToZ x) (fe41417_32WToZ y)
+ /\ is_bounded (fe41417_32WToZ (irop x y)) = true) (only parsing).
+Notation iunop_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe41417_32WToZ x) = true
+ -> fe41417_32WToZ (irop x) = op (fe41417_32WToZ x)
+ /\ is_bounded (fe41417_32WToZ (irop x)) = true) (only parsing).
+Notation iunop_FEToZ_correct irop op
+ := (forall x,
+ is_bounded (fe41417_32WToZ x) = true
+ -> word64ToZ (irop x) = op (fe41417_32WToZ x)) (only parsing).
+Notation iunop_FEToWire_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe41417_32WToZ x) = true
+ -> wire_digitsWToZ (irop x) = op (fe41417_32WToZ x)
+ /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
+Notation iunop_WireToFE_correct_and_bounded irop op
+ := (forall x,
+ wire_digits_is_bounded (wire_digitsWToZ x) = true
+ -> fe41417_32WToZ (irop x) = op (wire_digitsWToZ x)
+ /\ is_bounded (fe41417_32WToZ (irop x)) = true) (only parsing).
+
+Definition prefreeze := GF41417_32.prefreeze.
diff --git a/src/SpecificGen/GF41417_32Reflective.v b/src/SpecificGen/GF41417_32Reflective.v
new file mode 100644
index 000000000..d85e22870
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective.v
@@ -0,0 +1,119 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Export Crypto.SpecificGen.GF41417_32.
+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.Reify.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.SpecificGen.GF41417_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+Definition radd : ExprBinOp := Eval vm_compute in rcarry_addW.
+Definition rsub : ExprBinOp := Eval vm_compute in rcarry_subW.
+Definition rmul : ExprBinOp := Eval vm_compute in rmulW.
+Definition ropp : ExprUnOp := Eval vm_compute in rcarry_oppW.
+Definition rprefreeze : ExprUnOp := Eval vm_compute in rprefreezeW.
+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.
+
+Declare Reduction asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+Ltac asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+
+
+Definition interp_radd : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := Eval asm_interp in interp_bexpr (rword64ize radd).
+(*Print interp_radd.*)
+Definition interp_radd_correct : interp_radd = interp_bexpr radd := eq_refl.
+Definition interp_rsub : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rsub).
+(*Print interp_rsub.*)
+Definition interp_rsub_correct : interp_rsub = interp_bexpr rsub := eq_refl.
+Definition interp_rmul : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rmul).
+(*Print interp_rmul.*)
+Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
+Definition interp_ropp : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := Eval asm_interp in interp_uexpr (rword64ize ropp).
+(*Print interp_ropp.*)
+Definition interp_ropp_correct : interp_ropp = interp_uexpr ropp := eq_refl.
+Definition interp_rprefreeze : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := Eval asm_interp in interp_uexpr (rword64ize rprefreeze).
+(*Print interp_rprefreeze.*)
+Definition interp_rprefreeze_correct : interp_rprefreeze = interp_uexpr rprefreeze := eq_refl.
+
+Definition interp_rge_modulus : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.word64
+ := Eval asm_interp in interp_uexpr_FEToZ (rword64ize rge_modulus).
+Definition interp_rge_modulus_correct : interp_rge_modulus = interp_uexpr_FEToZ rge_modulus := eq_refl.
+
+Definition interp_rpack : SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.wire_digitsW
+ := Eval asm_interp in interp_uexpr_FEToWire (rword64ize rpack).
+Definition interp_rpack_correct : interp_rpack = interp_uexpr_FEToWire rpack := eq_refl.
+
+Definition interp_runpack : SpecificGen.GF41417_32BoundedCommon.wire_digitsW -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := Eval asm_interp in interp_uexpr_WireToFE (rword64ize 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.
+Proof. exact rcarry_addW_correct_and_bounded. Qed.
+Lemma rsub_correct_and_bounded : binop_correct_and_bounded rsub carry_sub.
+Proof. exact rcarry_subW_correct_and_bounded. Qed.
+Lemma rmul_correct_and_bounded : binop_correct_and_bounded rmul mul.
+Proof. exact rmulW_correct_and_bounded. Qed.
+Lemma ropp_correct_and_bounded : unop_correct_and_bounded ropp carry_opp.
+Proof. exact rcarry_oppW_correct_and_bounded. Qed.
+Lemma rprefreeze_correct_and_bounded : unop_correct_and_bounded rprefreeze prefreeze.
+Proof. exact rprefreezeW_correct_and_bounded. Qed.
+Lemma rge_modulus_correct_and_bounded : unop_FEToZ_correct rge_modulus ge_modulus.
+Proof. exact rge_modulusW_correct_and_bounded. Qed.
+Lemma rpack_correct_and_bounded : unop_FEToWire_correct_and_bounded rpack pack.
+Proof. exact rpackW_correct_and_bounded. Qed.
+Lemma runpack_correct_and_bounded : unop_WireToFE_correct_and_bounded runpack unpack.
+Proof. exact runpackW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF41417_32Reflective/Common.v b/src/SpecificGen/GF41417_32Reflective/Common.v
new file mode 100644
index 000000000..517c66d8a
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Common.v
@@ -0,0 +1,438 @@
+Require Export Coq.ZArith.ZArith.
+Require Export Coq.Strings.String.
+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.Reify.
+Require Export Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.InterpWfRel.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Reflection.MapInterpWf.
+Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.LetIn.
+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).
+
+Local Ltac make_type_from uncurried_op :=
+ let T := (type of uncurried_op) in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ exact rT.
+
+Definition ExprBinOpT : type base_type.
+Proof. make_type_from (uncurry_binop_fe41417_32 carry_add). Defined.
+Definition ExprUnOpT : type base_type.
+Proof. make_type_from (uncurry_unop_fe41417_32 carry_opp). Defined.
+Definition ExprUnOpFEToZT : type base_type.
+Proof. make_type_from (uncurry_unop_fe41417_32 ge_modulus). Defined.
+Definition ExprUnOpWireToFET : type base_type.
+Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Definition ExprUnOpFEToWireT : type base_type.
+Proof. make_type_from (uncurry_unop_fe41417_32 pack). Defined.
+Definition ExprBinOp : Type := Expr ExprBinOpT.
+Definition ExprUnOp : Type := Expr ExprUnOpT.
+Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
+Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
+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 :: ?xs)%list => let bs := bounds_from_list xs in
+ constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ end.
+
+Local Ltac make_bounds ls :=
+ compute;
+ let v := bounds_from_list (List.rev ls) in
+ let v := (eval compute in v) in
+ exact v.
+
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
+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).
+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).
+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).
+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).
+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).
+
+Notation binop_correct_and_bounded rop op
+ := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+Notation unop_correct_and_bounded rop op
+ := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
+Notation unop_FEToZ_correct rop op
+ := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
+Notation unop_FEToWire_correct_and_bounded rop op
+ := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
+Notation unop_WireToFE_correct_and_bounded rop op
+ := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
+
+Ltac rexpr_cbv :=
+ lazymatch goal with
+ | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ => let operf := head oper in
+ let uncurryf := head uncurry in
+ try cbv delta [T]; try cbv delta [oper];
+ try cbv beta iota delta [uncurryf]
+ end;
+ cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+
+Ltac reify_sig :=
+ rexpr_cbv; eexists; Reify_rhs; reflexivity.
+
+Local Notation rexpr_sig T uncurried_op :=
+ { rexprZ
+ | interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op }
+ (only parsing).
+
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (uncurry_binop_fe41417_32 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT (uncurry_unop_fe41417_32 op)) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT (uncurry_unop_fe41417_32 op)) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT (uncurry_unop_fe41417_32 op)) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET (uncurry_unop_wire_digits op)) (only parsing).
+
+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
+ 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))
+ (only parsing).
+
+Ltac app_tuples x y :=
+ let tx := type of x in
+ lazymatch (eval hnf in tx) with
+ | prod _ _ => let xs := app_tuples (snd x) y in
+ constr:((fst x, xs))
+ | _ => constr:((x, y))
+ end.
+
+Local Arguments Tuple.map2 : simpl never.
+Local Arguments Tuple.map : simpl never.
+
+Fixpoint args_to_bounded_helperT {n}
+ (v : Tuple.tuple' Word64.word64 n)
+ (bounds : Tuple.tuple' (Z * Z) n)
+ (pf : List.fold_right
+ andb true
+ (Tuple.to_list
+ _
+ (Tuple.map2
+ (n:=S 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)
+ (res : Type)
+ {struct n}
+ : Type.
+Proof.
+ refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
+ List.fold_right
+ _ _ (Tuple.to_list
+ _
+ (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
+ | 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
+ end v bounds pf).
+ { clear -pf0.
+ abstract (
+ destruct v', bounds'; simpl @fst;
+ rewrite Tuple.map_S in pf0;
+ simpl in pf0;
+ rewrite Tuple.map2_S in pf0;
+ simpl @List.fold_right in *;
+ rewrite Bool.andb_true_iff in pf0; tauto
+ ). }
+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.
+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 |}
+ | 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) |}))
+ end.
+ { clear -pf pf'.
+ unfold Tuple.map2, Tuple.map in pf; simpl in *.
+ abstract (
+ destruct bounds;
+ simpl in *;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+ { simpl in *.
+ clear -pf pf'.
+ abstract (
+ destruct bounds as [? [? ?] ], v; simpl in *;
+ rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
+ simpl in pf;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+Defined.
+
+Definition assoc_right''
+ := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
+
+Definition args_to_bounded {n} v bounds pf
+ := Eval cbv [args_to_bounded_helper assoc_right''] in
+ @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
+
+Local Ltac get_len T :=
+ match (eval hnf in T) with
+ | prod ?A ?B
+ => let a := get_len A in
+ let b := get_len B in
+ (eval compute in (a + b)%nat)
+ | _ => constr:(1%nat)
+ end.
+
+Local Ltac args_to_bounded x H :=
+ let x' := fresh in
+ set (x' := x);
+ compute in x;
+ let len := (let T := type of x in get_len T) in
+ destruct_head' prod;
+ let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
+ let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
+ let c := (eval cbv [all_binders_for ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
+ apply c; compute; clear;
+ try abstract (
+ repeat split;
+ solve [ reflexivity
+ | refine (fun v => match v with eq_refl => I end) ]
+ ).
+
+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).
+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).
+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).
+Proof.
+ let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
+ exact v.
+Defined.
+
+Ltac assoc_right_tuple x so_far :=
+ let t := type of x in
+ lazymatch (eval hnf in t) with
+ | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
+ assoc_right_tuple (fst x) so_far
+ | _ => lazymatch so_far with
+ | @None => x
+ | _ => constr:((x, so_far))
+ end
+ end.
+
+Local Ltac make_bounds_prop bounds orig_bounds :=
+ let bounds' := fresh "bounds'" in
+ let bounds_bad := fresh "bounds_bad" in
+ rename bounds into bounds_bad;
+ let boundsv := assoc_right_tuple bounds_bad (@None) in
+ pose boundsv as bounds;
+ pose orig_bounds as bounds';
+ repeat (refine (match fst bounds' with
+ | Some bounds' => let (l, u) := fst bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end && _)%bool;
+ destruct bounds' as [_ bounds'], bounds as [_ bounds]);
+ try exact (match bounds' with
+ | Some bounds' => let (l, u) := bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end).
+
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprBinOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToWireT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpWireToFET)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToZT)) : bool.
+Proof.
+ refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
+Defined.
+
+(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
+ various kinds of correct and boundedness, and abstract in Gallina
+ rather than Ltac *)
+
+Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
+ let Heq := fresh "Heq" in
+ let Hbounds0 := fresh "Hbounds0" in
+ let Hbounds1 := fresh "Hbounds1" in
+ let Hbounds2 := fresh "Hbounds2" in
+ pose proof (proj2_sig ropZ_sig) as Heq;
+ cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
+ curry_binop_fe41417_32W curry_unop_fe41417_32W curry_unop_wire_digitsW
+ curry_binop_fe41417_32 curry_unop_fe41417_32 curry_unop_wire_digits
+ uncurry_binop_fe41417_32W uncurry_unop_fe41417_32W uncurry_unop_wire_digitsW
+ uncurry_binop_fe41417_32 uncurry_unop_fe41417_32 uncurry_unop_wire_digits
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET
+ interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ cbv zeta in *;
+ simpl @fe41417_32WToZ; simpl @wire_digitsWToZ;
+ rewrite <- Heq; clear Heq;
+ destruct Hbounds as [Heq Hbounds];
+ change interp_op with (@Z.interp_op) in *;
+ 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;
+ specialize_by repeat first [ progress intros
+ | reflexivity
+ | assumption
+ | progress destruct_head' base_type
+ | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' and
+ | progress repeat apply conj ];
+ specialize (Hbounds_left args H0);
+ specialize (Hbounds_right args H0);
+ cbv beta in *;
+ lazymatch type of Hbounds_right with
+ | match ?e with _ => _ end
+ => lazymatch type of H1 with
+ | match ?e' with _ => _ end
+ => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
+ end
+ end;
+ 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;
+ match goal with
+ | [ |- fe41417_32WToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- wire_digitsWToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- _ = _ ]
+ => exact Hbounds_left
+ end;
+ change word64ToZ with Word64.word64ToZ in *;
+ (split; [ exact Hbounds_left | ]);
+ cbv [interp_flat_type] in *;
+ cbv [fst snd
+ binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good
+ ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds] in H1;
+ destruct_head' ZBounds.bounds;
+ unfold_is_bounded_in H1;
+ simpl @fe41417_32WToZ; simpl @wire_digitsWToZ;
+ unfold_is_bounded;
+ destruct_head' and;
+ Z.ltb_to_lt;
+ change Word64.word64ToZ with word64ToZ in *;
+ repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
+
+
+Ltac rexpr_correct :=
+ let ropW' := fresh in
+ let ropZ_sig := fresh in
+ intros ropW' ropZ_sig;
+ let wf_ropW := fresh "wf_ropW" in
+ assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
+ cbv zeta; repeat apply conj;
+ [ vm_compute; reflexivity
+ | apply @InterpRelWf;
+ [ | 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 compute_bounds opW bounds
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ (only parsing).
+
+
+Module Export PrettyPrinting.
+ (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
+ 8.5/8.6. Because [Set] is special and things break if
+ [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
+ to debug. *)
+ Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
+
+ Inductive result := yes | no | borked.
+
+ Definition ZBounds_to_bounds_on
+ := 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 |}
+ => in_range l u
+ | None
+ => overflow
+ end
+ end.
+
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result
+ := match t return interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result with
+ | Tbase TZ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
+ | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
+ | no, no => no
+ | yes, no | no, yes | yes, yes => yes
+ | borked, _ | _, borked => borked
+ end
+ end.
+
+ (** This gives a slightly easier to read version of the bounds *)
+ Notation compute_bounds_for_display opW bounds
+ := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
+ Notation sanity_compute opW bounds
+ := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
+ Notation sanity_check opW bounds
+ := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
+End PrettyPrinting.
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v b/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
new file mode 100644
index 000000000..93a83cf51
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
@@ -0,0 +1,50 @@
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprBinOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_binop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall xy
+ (xy := (eta_fe41417_32W (fst xy), eta_fe41417_32W (snd xy)))
+ (Hxy : is_bounded (fe41417_32WToZ (fst xy)) = true
+ /\ is_bounded (fe41417_32WToZ (snd xy)) = true),
+ let Hx := let (Hx, Hy) := Hxy in Hx in
+ 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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall xy
+ (xy := (eta_fe41417_32W (fst xy), eta_fe41417_32W (snd xy)))
+ (Hxy : is_bounded (fe41417_32WToZ (fst xy)) = true
+ /\ is_bounded (fe41417_32WToZ (snd xy)) = true),
+ 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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => binop_bounds_good bounds = true
+ | None => False
+ end)
+ : binop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x y Hx Hy.
+ pose x as x'; pose y as y'.
+ hnf in x, y; destruct_head' prod.
+ specialize (H0 (x', y') (conj Hx Hy)).
+ specialize (H1 (x', y') (conj Hx Hy)).
+ let args := constr:(binop_args_to_bounded (x', y') Hx Hy) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
new file mode 100644
index 000000000..c616b1a77
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe41417_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unop_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
new file mode 100644
index 000000000..7391e98bf
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToWire_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe41417_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToWire_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToWire_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
new file mode 100644
index 000000000..6d991dbb4
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToZ_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe41417_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToZ_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToZ_correct (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
new file mode 100644
index 000000000..bb91ea89f
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpWireToFE_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_wire_digitsW x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopWireToFE_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_WireToFE_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified.v b/src/SpecificGen/GF41417_32Reflective/Reified.v
new file mode 100644
index 000000000..ad8291b55
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified.v
@@ -0,0 +1,13 @@
+(** We split the reification up into separate files, one operation per
+ file, so that it can run in parallel. *)
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.Add.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.CarryAdd.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.Sub.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.CarrySub.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.Mul.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.Opp.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.CarryOpp.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.PreFreeze.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.GeModulus.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.Pack.
+Require Export Crypto.SpecificGen.GF41417_32Reflective.Reified.Unpack.
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/Add.v b/src/SpecificGen/GF41417_32Reflective/Reified/Add.v
new file mode 100644
index 000000000..ee249b9b1
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/Add.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonBinOp.
+
+Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
+Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
+Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
+Proof. rexpr_correct. Qed.
+Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
+Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
+Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/CarryAdd.v b/src/SpecificGen/GF41417_32Reflective/Reified/CarryAdd.v
new file mode 100644
index 000000000..db3d447ea
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/CarryAdd.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonBinOp.
+
+Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
+Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
+Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_addW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/CarryOpp.v b/src/SpecificGen/GF41417_32Reflective/Reified/CarryOpp.v
new file mode 100644
index 000000000..9ba6183b9
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/CarryOpp.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOp.
+
+Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
+Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
+Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_oppW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/CarrySub.v b/src/SpecificGen/GF41417_32Reflective/Reified/CarrySub.v
new file mode 100644
index 000000000..ccd18a264
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/CarrySub.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonBinOp.
+
+Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
+Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
+Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_subW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/GeModulus.v b/src/SpecificGen/GF41417_32Reflective/Reified/GeModulus.v
new file mode 100644
index 000000000..e1985d4c1
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/GeModulus.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOpFEToZ.
+
+Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
+Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
+Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rge_modulusW_correct_and_bounded
+ := ExprUnOpFEToZ_correct_and_bounded
+ rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/Mul.v b/src/SpecificGen/GF41417_32Reflective/Reified/Mul.v
new file mode 100644
index 000000000..3c3b38b85
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/Mul.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonBinOp.
+
+Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
+Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
+Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rmulW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
+Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
+Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/Opp.v b/src/SpecificGen/GF41417_32Reflective/Reified/Opp.v
new file mode 100644
index 000000000..4672002ac
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/Opp.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOp.
+
+Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
+Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
+Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
+Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
+Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/Pack.v b/src/SpecificGen/GF41417_32Reflective/Reified/Pack.v
new file mode 100644
index 000000000..d4fac5c2a
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/Pack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOpFEToWire.
+
+Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
+Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
+Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rpackW_correct_and_bounded
+ := ExprUnOpFEToWire_correct_and_bounded
+ rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v
new file mode 100644
index 000000000..889c6ca49
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOp.
+
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF41417_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
+Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rprefreezeW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/Sub.v b/src/SpecificGen/GF41417_32Reflective/Reified/Sub.v
new file mode 100644
index 000000000..fd25f599d
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/Sub.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonBinOp.
+
+Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
+Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
+Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
+Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
+Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/Unpack.v b/src/SpecificGen/GF41417_32Reflective/Reified/Unpack.v
new file mode 100644
index 000000000..f2f8881b7
--- /dev/null
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/Unpack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOpWireToFE.
+
+Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
+Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
+Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition runpackW_correct_and_bounded
+ := ExprUnOpWireToFE_correct_and_bounded
+ runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/SpecificGen/GF5211_32Bounded.v b/src/SpecificGen/GF5211_32Bounded.v
new file mode 100644
index 000000000..3574afadc
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Bounded.v
@@ -0,0 +1,462 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF5211_32.
+Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
+Require Import Crypto.SpecificGen.GF5211_32Reflective.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+
+Local Ltac bounded_t opW blem :=
+ apply blem; apply is_bounded_proj1_fe5211_32.
+Local Ltac bounded_wire_digits_t opW blem :=
+ apply blem; apply is_bounded_proj1_wire_digits.
+
+Local Ltac define_binop f g opW blem :=
+ refine (exist_fe5211_32W (opW (proj1_fe5211_32W f) (proj1_fe5211_32W g)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop f opW blem :=
+ refine (exist_fe5211_32W (opW (proj1_fe5211_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_FEToZ f opW :=
+ refine (opW (proj1_fe5211_32W f)).
+Local Ltac define_unop_FEToWire f opW blem :=
+ refine (exist_wire_digitsW (opW (proj1_fe5211_32W f)) _);
+ abstract bounded_t opW blem.
+Local Ltac define_unop_WireToFE f opW blem :=
+ refine (exist_fe5211_32W (opW (proj1_wire_digitsW f)) _);
+ 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 Arguments interp_radd / _ _.
+Local Arguments interp_rsub / _ _.
+Local Arguments interp_rmul / _ _.
+Local Arguments interp_ropp / _.
+Local Arguments interp_rprefreeze / _.
+Local Arguments interp_rge_modulus / _.
+Local Arguments interp_rpack / _.
+Local Arguments interp_runpack / _.
+Definition addW (f g : fe5211_32W) : fe5211_32W := Eval simpl in interp_radd f g.
+Definition subW (f g : fe5211_32W) : fe5211_32W := Eval simpl in interp_rsub f g.
+Definition mulW (f g : fe5211_32W) : fe5211_32W := Eval simpl in interp_rmul f g.
+Definition oppW (f : fe5211_32W) : fe5211_32W := Eval simpl in interp_ropp f.
+Definition prefreezeW (f : fe5211_32W) : fe5211_32W := Eval simpl in interp_rprefreeze f.
+Definition ge_modulusW (f : fe5211_32W) : word64 := Eval simpl in interp_rge_modulus f.
+Definition packW (f : fe5211_32W) : wire_digitsW := Eval simpl in interp_rpack f.
+Definition unpackW (f : wire_digitsW) : fe5211_32W := Eval simpl in interp_runpack f.
+
+Require Import ModularBaseSystemWord.
+Definition modulusW :=
+ Eval cbv - [ZToWord64] in (Tuple.map ZToWord64 (Tuple.from_list_default 0%Z 10 GF5211_32.modulus_digits_)).
+
+Definition postfreeze : GF5211_32.fe5211_32 -> GF5211_32.fe5211_32 :=
+ GF5211_32.postfreeze.
+
+Lemma freeze_prepost_freeze : forall x, postfreeze (prefreeze x) = GF5211_32.freeze x. Proof. reflexivity. Qed.
+
+Definition postfreezeW : fe5211_32W -> fe5211_32W :=
+ (conditional_subtract_modulusW
+ (num_limbs := 10)
+ modulusW
+ ge_modulusW
+ (Interpretations.Word64.neg GF5211_32.int_width)
+ ).
+
+Definition freezeW (f : fe5211_32W) : fe5211_32W := Eval cbv beta delta [prefreezeW postfreezeW] in postfreezeW (prefreezeW f).
+
+Local Transparent Let_In.
+Definition powW (f : fe5211_32W) chain := fold_chain_opt (proj1_fe5211_32W one) mulW chain [f].
+Definition invW (f : fe5211_32W) : fe5211_32W
+ := Eval cbv -[Let_In fe5211_32W mulW] in powW f (chain inv_ec).
+
+Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
+ change opW with (interp_rop);
+ rewrite pre_rewrite;
+ intros; apply rop_cb; assumption.
+
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Proof. port_correct_and_bounded interp_radd_correct addW interp_radd radd_correct_and_bounded. Qed.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW carry_sub.
+Proof. port_correct_and_bounded interp_rsub_correct subW interp_rsub rsub_correct_and_bounded. Qed.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW mul.
+Proof. port_correct_and_bounded interp_rmul_correct mulW interp_rmul rmul_correct_and_bounded. Qed.
+Lemma oppW_correct_and_bounded : iunop_correct_and_bounded oppW carry_opp.
+Proof. port_correct_and_bounded interp_ropp_correct oppW interp_ropp ropp_correct_and_bounded. Qed.
+Lemma prefreezeW_correct_and_bounded : iunop_correct_and_bounded prefreezeW prefreeze.
+Proof. port_correct_and_bounded interp_rprefreeze_correct prefreezeW interp_rprefreeze rprefreeze_correct_and_bounded. Qed.
+Lemma ge_modulusW_correct : iunop_FEToZ_correct ge_modulusW ge_modulus.
+Proof. port_correct_and_bounded interp_rge_modulus_correct ge_modulusW interp_rge_modulus rge_modulus_correct_and_bounded. Qed.
+Lemma packW_correct_and_bounded : iunop_FEToWire_correct_and_bounded packW pack.
+Proof. port_correct_and_bounded interp_rpack_correct packW interp_rpack rpack_correct_and_bounded. Qed.
+Lemma unpackW_correct_and_bounded : iunop_WireToFE_correct_and_bounded unpackW unpack.
+Proof. port_correct_and_bounded interp_runpack_correct unpackW interp_runpack runpack_correct_and_bounded. Qed.
+
+(* TODO : move *)
+Lemma neg_range : forall x y, 0 <= x ->
+ 0 <= ModularBaseSystemListZOperations.neg x y < 2 ^ x.
+Proof.
+ intros.
+ split; auto using ModularBaseSystemListZOperationsProofs.neg_nonneg.
+ eapply Z.le_lt_trans; eauto using ModularBaseSystemListZOperationsProofs.neg_upperbound.
+ rewrite Z.ones_equiv.
+ omega.
+Qed.
+
+Ltac lower_bound_minus_ge_modulus :=
+ apply Z.le_0_sub;
+ 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;
+ etransitivity; try apply Z.land_upper_bound_r; instantiate; try omega;
+ apply Z.ones_nonneg; instantiate; vm_compute; discriminate.
+
+Ltac upper_bound_minus_ge_modulus :=
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | ];
+ eapply Z.le_lt_trans; [ eassumption | ];
+ instantiate; vm_compute; reflexivity.
+
+Lemma postfreezeW_correct_and_bounded : iunop_correct_and_bounded postfreezeW postfreeze.
+Proof.
+ intros x H.
+ pose proof (ge_modulusW_correct x H) as Hgm.
+ destruct_head_hnf' prod.
+ unfold_is_bounded_in H.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ cbv [postfreezeW].
+ cbv [conditional_subtract_modulusW Interpretations.Word64.neg].
+ change word64ToZ with Interpretations.Word64.word64ToZ in *.
+ rewrite Hgm.
+
+ cbv [modulusW Tuple.map].
+ cbv [on_tuple List.map to_list to_list' from_list from_list'
+ Tuple.map2 on_tuple2 ListUtil.map2 fe5211_32WToZ].
+ cbv [postfreeze GF5211_32.postfreeze].
+ cbv [Let_In].
+
+ split.
+ { match goal with
+ |- (_,word64ToZ (_ ^- (Interpretations.Word64.ZToWord64 ?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;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ | |- (_,_) = (_,_) => reflexivity
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ end. }
+
+
+ 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;
+ repeat match goal with |- _ /\ _ => split; Z.ltb_to_lt end;
+ try match goal with
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ < 2 ^ _ => apply neg_range; omega
+ | |- 0 <= _ < 2 ^ Z.of_nat _ => vm_compute; split; [refine (fun x => match x with eq_refl => I end) | reflexivity]
+ | |- 0 <= _ &' _ => apply Z.land_nonneg; right; omega
+ end;
+ try solve [
+ (apply Z.log2_lt_pow2_alt; [ vm_compute; reflexivity | ]);
+ eapply Z.le_lt_trans; try apply Z.land_upper_bound_r; try apply neg_range; instantiate; try match goal with |- ?G => not has_evar G; vm_compute; discriminate end; reflexivity];
+ try match goal with
+ | |- 0 <= _ - _ => lower_bound_minus_ge_modulus
+ | |- Z.log2 (_ - _) < _ => upper_bound_minus_ge_modulus
+ | |- _ - _ <= _ => etransitivity; [ apply Z.le_sub_nonneg; apply Z.land_nonneg; right; omega | instantiate; assumption ]
+ | |- 0 <= ModularBaseSystemListZOperations.neg _ _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- ModularBaseSystemListZOperations.neg _ _ < _ =>
+ apply neg_range; vm_compute; discriminate
+ | |- _ => vm_compute; (discriminate || reflexivity)
+ end.
+Qed.
+
+Lemma freezeW_correct_and_bounded : iunop_correct_and_bounded freezeW freeze.
+Proof.
+ intros f H; rewrite <- freeze_prepost_freeze.
+ change (freezeW f) with (postfreezeW (prefreezeW f)).
+ destruct (prefreezeW_correct_and_bounded f H) as [H0 H1].
+ destruct (postfreezeW_correct_and_bounded _ H1) as [H0' H1'].
+ rewrite H1', H0', H0; split; reflexivity.
+Qed.
+
+Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
+Proof.
+ cbv [powW pow].
+ intro x; intros; apply (fold_chain_opt_gen fe5211_32WToZ is_bounded [x]).
+ { reflexivity. }
+ { reflexivity. }
+ { intros; progress rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption.
+ apply mulW_correct_and_bounded; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded _ _ X Y)) by assumption; reflexivity. }
+ { intros [|?]; autorewrite with simpl_nth_default;
+ (assumption || reflexivity). }
+Qed.
+
+Lemma invW_correct_and_bounded : iunop_correct_and_bounded invW inv.
+Proof.
+ intro f.
+ assert (H : forall f, invW f = powW f (chain inv_ec))
+ by abstract (cbv -[Let_In fe5211_32W mulW]; reflexivity).
+ rewrite H.
+ rewrite inv_correct.
+ cbv [inv_opt].
+ rewrite <- pow_correct.
+ apply powW_correct_and_bounded.
+Qed.
+
+Definition fieldwisebW_sig (f g : fe5211_32W)
+ : { b | b = GF5211_32.fieldwiseb (fe5211_32WToZ f) (fe5211_32WToZ g) }.
+Proof.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ cbv [GF5211_32.fieldwiseb fe5211_32WToZ].
+ rewrite ?word64eqb_Zeqb.
+ reflexivity.
+Defined.
+
+Definition fieldwisebW (f g : fe5211_32W) : bool :=
+ Eval cbv [proj1_sig fieldwisebW_sig appify2 app_fe5211_32W] in
+ appify2 (fun f g => proj1_sig (fieldwisebW_sig f g)) f g.
+
+Lemma fieldwisebW_correct f g
+ : fieldwisebW f g = GF5211_32.fieldwiseb (fe5211_32WToZ f) (fe5211_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (fieldwisebW_sig f' g')).
+Qed.
+
+Local Arguments freezeW : simpl never.
+Local Arguments fe5211_32WToZ !_ / .
+Local Opaque freezeW.
+
+Definition eqbW_sig (f g : fe5211_32W)
+ : { b | is_bounded (fe5211_32WToZ f) = true
+ -> is_bounded (fe5211_32WToZ g) = true
+ -> b = GF5211_32.eqb (fe5211_32WToZ f) (fe5211_32WToZ g) }.
+Proof.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded f pf)) as frf.
+ pose proof (fun pf => proj1 (freezeW_correct_and_bounded g pf)) as frg.
+ hnf in f, g; destruct_head' prod.
+ eexists.
+ unfold GF5211_32.eqb.
+ simpl @fe5211_32WToZ in *; cbv beta iota.
+ intros.
+ rewrite <- frf, <- frg by assumption.
+ rewrite <- fieldwisebW_correct.
+ reflexivity.
+Defined.
+
+Definition eqbW (f g : fe5211_32W) : bool :=
+ Eval cbv [proj1_sig eqbW_sig appify2 app_fe5211_32W] in
+ appify2 (fun f g => proj1_sig (eqbW_sig f g)) f g.
+
+Lemma eqbW_correct f g
+ : is_bounded (fe5211_32WToZ f) = true
+ -> is_bounded (fe5211_32WToZ g) = true
+ -> eqbW f g = GF5211_32.eqb (fe5211_32WToZ f) (fe5211_32WToZ g).
+Proof.
+ set (f' := f); set (g' := g).
+ hnf in f, g; destruct_head' prod.
+ exact (proj2_sig (eqbW_sig f' g')).
+Qed.
+
+(* TODO(jgross): use NToWord or such for this constant too *)
+Definition sqrt_m1W : fe5211_32W :=
+ Eval vm_compute in fe5211_32ZToW sqrt_m1.
+
+Definition GF5211_32sqrt x : GF5211_32.fe5211_32 :=
+ dlet powx := powW (fe5211_32ZToW x) (chain GF5211_32.sqrt_ec) in
+ GF5211_32.sqrt (fe5211_32WToZ powx) (fe5211_32WToZ (mulW powx powx)) x.
+
+Definition sqrtW_sig
+ : { sqrtW | iunop_correct_and_bounded sqrtW GF5211_32sqrt }.
+Proof.
+ eexists.
+ unfold GF5211_32sqrt, GF5211_32.sqrt.
+ intros.
+ rewrite !fe5211_32ZToW_WToZ.
+ split.
+ { etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody_eq; intros.
+ set_evars.
+ match goal with (* unfold the first dlet ... in, but only if it's binding a var *)
+ | [ |- ?x = dlet y := fe5211_32WToZ ?z in ?f ]
+ => is_var z; change (x = match fe5211_32WToZ z with y => f end)
+ end.
+ change sqrt_m1 with (fe5211_32WToZ sqrt_m1W).
+ rewrite <- (fun X Y => proj1 (mulW_correct_and_bounded sqrt_m1W a X Y)), <- eqbW_correct, (pull_bool_if fe5211_32WToZ)
+ by repeat match goal with
+ | _ => progress subst
+ | [ |- is_bounded (fe5211_32WToZ ?op) = true ]
+ => lazymatch op with
+ | mulW _ _ => apply mulW_correct_and_bounded
+ | powW _ _ => apply powW_correct_and_bounded
+ | sqrt_m1W => vm_compute; reflexivity
+ | _ => assumption
+ end
+ end.
+ subst_evars; reflexivity.
+ } Unfocus.
+ lazymatch goal with
+ | [ |- context G[dlet x := ?v in fe5211_32WToZ (@?f x)] ]
+ => let G' := context G[fe5211_32WToZ (dlet x := v in f x)] in
+ cut G'; cbv beta;
+ [ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
+ end.
+ reflexivity. }
+ { cbv [Let_In].
+ break_if.
+ { apply powW_correct_and_bounded; assumption. }
+ { apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ].
+ apply powW_correct_and_bounded; assumption. } }
+Defined.
+
+Definition sqrtW (f : fe5211_32W) : fe5211_32W :=
+ Eval cbv [proj1_sig sqrtW_sig app_fe5211_32W] in
+ app_fe5211_32W f (proj1_sig sqrtW_sig).
+
+Lemma sqrtW_correct_and_bounded : iunop_correct_and_bounded sqrtW GF5211_32sqrt.
+Proof.
+ intro f.
+ set (f' := f).
+ hnf in f; destruct_head' prod.
+ assert (H : sqrtW f' = proj1_sig sqrtW_sig f')
+ by (subst f'; cbv beta iota delta [proj1_sig sqrtW_sig sqrtW]; reflexivity).
+ rewrite H.
+ exact (proj2_sig sqrtW_sig f').
+Qed.
+
+
+
+Definition add (f g : fe5211_32) : fe5211_32.
+Proof. define_binop f g addW addW_correct_and_bounded. Defined.
+Definition sub (f g : fe5211_32) : fe5211_32.
+Proof. define_binop f g subW subW_correct_and_bounded. Defined.
+Definition mul (f g : fe5211_32) : fe5211_32.
+Proof. define_binop f g mulW mulW_correct_and_bounded. Defined.
+Definition opp (f : fe5211_32) : fe5211_32.
+Proof. define_unop f oppW oppW_correct_and_bounded. Defined.
+Definition freeze (f : fe5211_32) : fe5211_32.
+Proof. define_unop f freezeW freezeW_correct_and_bounded. Defined.
+Definition ge_modulus (f : fe5211_32) : word64.
+Proof. define_unop_FEToZ f ge_modulusW. Defined.
+Definition pack (f : fe5211_32) : wire_digits.
+Proof. define_unop_FEToWire f packW packW_correct_and_bounded. Defined.
+Definition unpack (f : wire_digits) : fe5211_32.
+Proof. define_unop_WireToFE f unpackW unpackW_correct_and_bounded. Defined.
+
+Definition pow (f : fe5211_32) (chain : list (nat * nat)) : fe5211_32.
+Proof. define_unop f (fun x => powW x chain) powW_correct_and_bounded. Defined.
+Definition inv (f : fe5211_32) : fe5211_32.
+Proof. define_unop f invW invW_correct_and_bounded. Defined.
+Definition sqrt (f : fe5211_32) : fe5211_32.
+Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
+
+Local Ltac op_correct_t op opW_correct_and_bounded :=
+ cbv [op];
+ lazymatch goal with
+ | [ |- context[proj1_fe5211_32 (exist_fe5211_32W _ _)] ]
+ => rewrite proj1_fe5211_32_exist_fe5211_32W
+ | [ |- context[proj1_wire_digits (exist_wire_digitsW _ _)] ]
+ => rewrite proj1_wire_digits_exist_wire_digitsW
+ | _ => idtac
+ end;
+ apply opW_correct_and_bounded;
+ lazymatch goal with
+ | [ |- is_bounded _ = true ]
+ => apply is_bounded_proj1_fe5211_32
+ | [ |- wire_digits_is_bounded _ = true ]
+ => apply is_bounded_proj1_wire_digits
+ end.
+
+Lemma add_correct (f g : fe5211_32) : proj1_fe5211_32 (add f g) = carry_add (proj1_fe5211_32 f) (proj1_fe5211_32 g).
+Proof. op_correct_t add addW_correct_and_bounded. Qed.
+Lemma sub_correct (f g : fe5211_32) : proj1_fe5211_32 (sub f g) = carry_sub (proj1_fe5211_32 f) (proj1_fe5211_32 g).
+Proof. op_correct_t sub subW_correct_and_bounded. Qed.
+Lemma mul_correct (f g : fe5211_32) : proj1_fe5211_32 (mul f g) = GF5211_32.mul (proj1_fe5211_32 f) (proj1_fe5211_32 g).
+Proof. op_correct_t mul mulW_correct_and_bounded. Qed.
+Lemma opp_correct (f : fe5211_32) : proj1_fe5211_32 (opp f) = carry_opp (proj1_fe5211_32 f).
+Proof. op_correct_t opp oppW_correct_and_bounded. Qed.
+Lemma freeze_correct (f : fe5211_32) : proj1_fe5211_32 (freeze f) = GF5211_32.freeze (proj1_fe5211_32 f).
+Proof. op_correct_t freeze freezeW_correct_and_bounded. Qed.
+Lemma ge_modulus_correct (f : fe5211_32) : word64ToZ (ge_modulus f) = GF5211_32.ge_modulus (proj1_fe5211_32 f).
+Proof. op_correct_t ge_modulus ge_modulusW_correct. Qed.
+Lemma pack_correct (f : fe5211_32) : proj1_wire_digits (pack f) = GF5211_32.pack (proj1_fe5211_32 f).
+Proof. op_correct_t pack packW_correct_and_bounded. Qed.
+Lemma unpack_correct (f : wire_digits) : proj1_fe5211_32 (unpack f) = GF5211_32.unpack (proj1_wire_digits f).
+Proof. op_correct_t unpack unpackW_correct_and_bounded. Qed.
+Lemma pow_correct (f : fe5211_32) chain : proj1_fe5211_32 (pow f chain) = GF5211_32.pow (proj1_fe5211_32 f) chain.
+Proof. op_correct_t pow (powW_correct_and_bounded chain). Qed.
+Lemma inv_correct (f : fe5211_32) : proj1_fe5211_32 (inv f) = GF5211_32.inv (proj1_fe5211_32 f).
+Proof. op_correct_t inv invW_correct_and_bounded. Qed.
+Lemma sqrt_correct (f : fe5211_32) : proj1_fe5211_32 (sqrt f) = GF5211_32sqrt (proj1_fe5211_32 f).
+Proof. op_correct_t sqrt sqrtW_correct_and_bounded. Qed.
+
+Import Morphisms.
+
+Local Existing Instance prime_modulus.
+
+Lemma field5211_32_and_homomorphisms
+ : @field fe5211_32 eq zero one opp add sub mul inv div
+ /\ @Ring.is_homomorphism (F _) (@Logic.eq _) 1%F F.add F.mul fe5211_32 eq one add mul encode
+ /\ @Ring.is_homomorphism fe5211_32 eq one add mul (F _) (@Logic.eq _) 1%F F.add F.mul decode.
+Proof.
+ eapply @Field.field_and_homomorphism_from_redundant_representation.
+ { exact (F.field_modulo _). }
+ { cbv [decode encode]; intros; rewrite !proj1_fe5211_32_exist_fe5211_32; apply encode_rep. }
+ { reflexivity. }
+ { reflexivity. }
+ { reflexivity. }
+ { cbv [decode encode]; intros; rewrite opp_correct, carry_opp_correct, carry_opp_opt_correct, carry_opp_rep; apply opp_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite add_correct, carry_add_correct, carry_add_opt_correct, carry_add_rep; apply add_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite sub_correct, carry_sub_correct, carry_sub_opt_correct, carry_sub_rep; apply sub_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite mul_correct, GF5211_32.mul_correct, carry_mul_opt_correct by reflexivity; apply carry_mul_rep; reflexivity. }
+ { cbv [decode encode]; intros; rewrite inv_correct, GF5211_32.inv_correct, inv_opt_correct by reflexivity; apply inv_rep; reflexivity. }
+ { cbv [decode encode div]; intros; rewrite !proj1_fe5211_32_exist_fe5211_32; apply encode_rep. }
+Qed.
+
+Global Instance field5211_32 : @field fe5211_32 eq zero one opp add sub mul inv div := proj1 field5211_32_and_homomorphisms.
+
+Local Opaque proj1_fe5211_32 exist_fe5211_32 proj1_fe5211_32W exist_fe5211_32W.
+Global Instance homomorphism_F5211_32_encode
+ : @Ring.is_homomorphism (F modulus) Logic.eq F.one F.add F.mul fe5211_32 eq one add mul encode.
+Proof. apply field5211_32_and_homomorphisms. Qed.
+
+Global Instance homomorphism_F5211_32_decode
+ : @Ring.is_homomorphism fe5211_32 eq one add mul (F modulus) Logic.eq F.one F.add F.mul decode.
+Proof. apply field5211_32_and_homomorphisms. Qed.
diff --git a/src/SpecificGen/GF5211_32BoundedCommon.v b/src/SpecificGen/GF5211_32BoundedCommon.v
new file mode 100644
index 000000000..bd2394540
--- /dev/null
+++ b/src/SpecificGen/GF5211_32BoundedCommon.v
@@ -0,0 +1,693 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.SpecificGen.GF5211_32.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Util.HList.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+(* BEGIN common curve-specific definitions *)
+Definition bit_width : nat := Eval compute in Z.to_nat (GF5211_32.int_width).
+Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
+Definition bounds_exp : tuple Z length_fe5211_32
+ := Eval compute in
+ Tuple.from_list length_fe5211_32 limb_widths eq_refl.
+Definition bounds : tuple (Z * Z) length_fe5211_32
+ := Eval compute in
+ Tuple.map (fun e => b_of e) bounds_exp.
+Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
+ := Eval compute in Tuple.from_list _ wire_widths eq_refl.
+Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
+ := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
+(* 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.
+Arguments bit_width : simpl never.
+Global Opaque word64.
+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},
+ 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},
+ 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
+ := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
+ Definition Build_is_boundedT {lower upper} {proj_word : word64}
+ : 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}
+ : 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 }.
+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 fe5211_32W := Eval cbv (*-[word64]*) in (tuple word64 length_fe5211_32).
+Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
+Definition fe5211_32 :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => word_of e) bounds_exp.
+Definition wire_digits :=
+ Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
+ hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
+
+Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
+ := let res := Tuple.map2
+ (fun bounds v =>
+ let '(lower, upper) := bounds in
+ (lower <=? v) && (v <=? upper))%bool%Z
+ bounds x in
+ List.fold_right andb true (Tuple.to_list _ res).
+
+Definition is_bounded (x : SpecificGen.GF5211_32.fe5211_32) : bool
+ := is_bounded_gen (n:=length_fe5211_32) x bounds.
+
+Definition wire_digits_is_bounded (x : SpecificGen.GF5211_32.wire_digits) : bool
+ := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
+
+(* END basic types *)
+
+Section generic_destructuring.
+ Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
+ := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
+ end.
+ Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
+ := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
+ | O => fun T v P => P v
+ | S n' => @app_on' A n'
+ end.
+ Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
+ Qed.
+ Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
+
+ Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
+ end.
+ Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
+ := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
+ | O => fun ts T v P => P v
+ | S n' => @app_on_h' A F n'
+ end.
+ Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
+ Proof.
+ induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
+ Qed.
+ Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
+ Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
+
+ Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
+ : forall (f : tuple A (length wire_widths)), T f
+ := Eval compute in fun f => @app_on A (length wire_widths) T f P.
+ Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
+ := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
+ Definition app_fe5211_32W_dep {A T} (P : forall x : tuple A length_fe5211_32, T x)
+ : forall (f : tuple A length_fe5211_32), T f
+ := Eval compute in fun f => @app_on A length_fe5211_32 T f P.
+ Definition app_fe5211_32W {A T} (f : tuple A length_fe5211_32) (P : tuple A length_fe5211_32 -> T)
+ := Eval compute in @app_fe5211_32W_dep A (fun _ => T) P f.
+ Definition app_fe5211_32_dep {T} (P : forall x : fe5211_32, T x)
+ : forall f : fe5211_32, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe5211_32 bounds_exp T f P.
+ Definition app_fe5211_32 {T} (f : fe5211_32) (P : hlist (fun e => word_of e) bounds_exp -> T)
+ := Eval compute in @app_fe5211_32_dep (fun _ => T) P f.
+ Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
+ : forall f : wire_digits, T f
+ := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
+ Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
+ := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
+
+ Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
+ := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
+ Definition app_fe5211_32W_dep_correct {A T} f P : @app_fe5211_32W_dep A T P f = P f
+ := app_on_correct f P.
+ Definition app_fe5211_32W_correct {A T} f P : @app_fe5211_32W A T f P = P f
+ := @app_fe5211_32W_dep_correct A (fun _ => T) f P.
+ Definition app_fe5211_32_dep_correct {T} f P : @app_fe5211_32_dep T P f = P f
+ := app_on_h_correct (fun e => word_of e) bounds_exp f P.
+ Definition app_fe5211_32_correct {T} f P : @app_fe5211_32 T f P = P f
+ := @app_fe5211_32_dep_correct (fun _ => T) f P.
+ Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
+ := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
+ Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
+ := @app_wire_digits_dep_correct (fun _ => T) f P.
+
+ Definition appify2 {T} (op : fe5211_32W -> fe5211_32W -> T) (f g : fe5211_32W) :=
+ app_fe5211_32W f (fun f0 => (app_fe5211_32W g (fun g0 => op f0 g0))).
+
+ Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
+ Proof.
+ intros. cbv [appify2].
+ etransitivity; apply app_fe5211_32W_correct.
+ Qed.
+End generic_destructuring.
+
+Definition eta_fe5211_32W_sig (x : fe5211_32W) : { v : fe5211_32W | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+Defined.
+Definition eta_fe5211_32W (x : fe5211_32W) : fe5211_32W
+ := Eval cbv [proj1_sig eta_fe5211_32W_sig] in proj1_sig (eta_fe5211_32W_sig x).
+Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
+Proof.
+ eexists; symmetry.
+ repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
+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.
+Proof. apply NToWord_wordToN. Qed.
+Local Opaque word64.
+
+Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ 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.
+Proof.
+ intros; unfold word64ToZ, ZToWord64.
+ rewrite N2Z.id, NToWord_wordToN; reflexivity.
+Qed.
+
+(* BEGIN precomputation. *)
+
+Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
+Definition word_bounded {lower upper} (v : bounded_word lower upper)
+ : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
+ := project_is_boundedT (proj2_sig v).
+Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
+ := exist _ proj_word (Build_is_boundedT word_bounded).
+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)
+ : bounded_word lower upper
+ := Build_bounded_word'
+ proj_word
+ (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
+ | 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.
+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.
+ match goal with
+ | [ |- context[andb ?x ?y] ]
+ => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
+ end;
+ intros; omega.
+Qed.
+Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
+Proof.
+ refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
+ abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
+Defined.
+Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
+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.
+Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe5211_32W app_fe5211_32 HList.mapt HList.mapt' Tuple.map on_tuple List.map List.app length_fe5211_32 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
+Definition fe5211_32WToZ (x : fe5211_32W) : SpecificGen.GF5211_32.fe5211_32
+ := Eval app_tuple_map in
+ app_fe5211_32W x (Tuple.map (fun v : word64 => v : Z)).
+Definition fe5211_32ZToW (x : SpecificGen.GF5211_32.fe5211_32) : fe5211_32W
+ := Eval app_tuple_map in
+ app_fe5211_32W x (Tuple.map (fun v : Z => v : word64)).
+Definition wire_digitsWToZ (x : wire_digitsW) : SpecificGen.GF5211_32.wire_digits
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
+Definition wire_digitsZToW (x : SpecificGen.GF5211_32.wire_digits) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
+Definition fe5211_32W_word64ize (x : fe5211_32W) : fe5211_32W
+ := Eval app_tuple_map in
+ app_fe5211_32W x (Tuple.map word64ize).
+Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
+ := Eval app_tuple_map in
+ app_wire_digitsW x (Tuple.map word64ize).
+
+(** TODO: Turn this into a lemma to speed up proofs *)
+Ltac unfold_is_bounded_in H :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe5211_32WToZ, wire_digitsWToZ in H;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe5211_32 List.length wire_widths] in H;
+ rewrite ?Bool.andb_true_iff in H.
+
+Ltac unfold_is_bounded :=
+ unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe5211_32WToZ, wire_digitsWToZ;
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map fold_right List.rev List.app length_fe5211_32 List.length wire_widths];
+ 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.
+Local Opaque bit_width.
+
+Local Ltac prove_lt_bit_width :=
+ rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
+
+Lemma fe5211_32ZToW_WToZ (x : fe5211_32W) : fe5211_32ZToW (fe5211_32WToZ x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe5211_32WToZ fe5211_32ZToW].
+ rewrite !ZToWord64_word64ToZ; reflexivity.
+Qed.
+
+Lemma fe5211_32WToZ_ZToW x : is_bounded x = true -> fe5211_32WToZ (fe5211_32ZToW x) = x.
+Proof.
+ hnf in x; destruct_head' prod; cbv [fe5211_32WToZ fe5211_32ZToW].
+ intro H.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma fe5211_32W_word64ize_id x : fe5211_32W_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [fe5211_32W_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [wire_digitsW_word64ize];
+ repeat apply f_equal2; apply word64ize_id.
+Qed.
+
+Definition uncurry_unop_fe5211_32W {T} (op : fe5211_32W -> T)
+ := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe5211_32) op.
+Definition curry_unop_fe5211_32W {T} op : fe5211_32W -> T
+ := Eval cbv (*-[word64]*) in fun f => app_fe5211_32W f (Tuple.curry (n:=length_fe5211_32) op).
+Definition uncurry_binop_fe5211_32W {T} (op : fe5211_32W -> fe5211_32W -> T)
+ := Eval cbv (*-[word64]*) in uncurry_unop_fe5211_32W (fun f => uncurry_unop_fe5211_32W (op f)).
+Definition curry_binop_fe5211_32W {T} op : fe5211_32W -> fe5211_32W -> T
+ := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe5211_32W (curry_unop_fe5211_32W op f)).
+
+Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
+ := Eval cbv (*-[word64]*) 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).
+
+
+Definition proj1_fe5211_32W (x : fe5211_32) : fe5211_32W
+ := Eval app_tuple_map in
+ app_fe5211_32 x (HList.mapt (fun _ => (@proj_word _ _))).
+Coercion proj1_fe5211_32 (x : fe5211_32) : SpecificGen.GF5211_32.fe5211_32
+ := fe5211_32WToZ (proj1_fe5211_32W x).
+
+Lemma is_bounded_proj1_fe5211_32 (x : fe5211_32) : is_bounded (proj1_fe5211_32 x) = true.
+Proof.
+ revert x; refine (app_fe5211_32_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [is_bounded proj1_fe5211_32 proj1_fe5211_32W fe5211_32WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe5211_32 is_bounded_gen].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
+ := app_wire_digits x (HList.mapt (fun _ => proj_word)).
+Coercion proj1_wire_digits (x : wire_digits) : SpecificGen.GF5211_32.wire_digits
+ := wire_digitsWToZ (proj1_wire_digitsW x).
+
+Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
+Proof.
+ revert x; refine (app_wire_digits_dep _); intro x.
+ hnf in x; destruct_head' prod; destruct_head' bounded_word.
+ cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [fold_right List.map].
+ cbv beta in *.
+ repeat split; auto using project_is_boundedT.
+Qed.
+
+Local Ltac make_exist_W' x app_W_dep :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x);
+ cbv [tuple tuple' length_fe5211_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v H :=
+ first [ let v' := (eval cbv [snd fst] in (snd v)) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
+ | let v' := (eval cbv [snd fst] in v) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
+ let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
+ let T := type of H' in
+ let T := (eval cbv [id
+ List.fold_right List.map List.length List.app ListUtil.map2 List.rev
+ Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
+ fe5211_32 bounds fe5211_32WToZ length_fe5211_32
+ wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
+ let H' := constr:(H' : T) in
+ let v := (eval unfold x' in x') in
+ do_refine v H'.
+Local Ltac make_exist'' x exist_W ZToW :=
+ let H := fresh in
+ intro H; apply (exist_W (ZToW x));
+ abstract (
+ hnf in x; destruct_head' prod;
+ let H' := fresh in
+ pose proof H as H';
+ unfold_is_bounded_in H;
+ destruct_head' and; simpl in *;
+ Z.ltb_to_lt;
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
+ assumption
+ ).
+Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
+ let H := fresh in
+ revert x; refine (@app_W_dep _ _ _); intros x H;
+ let x' := fresh in
+ set (x' := x) in *;
+ cbv [tuple tuple' length_fe5211_32 List.length wire_widths] in x;
+ destruct_head' prod;
+ let rec do_refine v :=
+ first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
+ refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
+ | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
+ refine (_, Build_bounded_word v' _);
+ [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
+ let v := (eval unfold x' in (exist'' x' H)) in
+ do_refine v.
+
+Definition exist_fe5211_32W' (x : fe5211_32W) : is_bounded (fe5211_32WToZ x) = true -> fe5211_32.
+Proof. make_exist_W' x (@app_fe5211_32W_dep). Defined.
+Definition exist_fe5211_32W (x : fe5211_32W) : is_bounded (fe5211_32WToZ x) = true -> fe5211_32
+ := Eval cbv [app_fe5211_32W_dep exist_fe5211_32W' fe5211_32ZToW] in exist_fe5211_32W' x.
+Definition exist_fe5211_32'' (x : SpecificGen.GF5211_32.fe5211_32) : is_bounded x = true -> fe5211_32.
+Proof. make_exist'' x exist_fe5211_32W fe5211_32ZToW. Defined.
+Definition exist_fe5211_32' (x : SpecificGen.GF5211_32.fe5211_32) : is_bounded x = true -> fe5211_32.
+Proof. make_exist' x (@app_fe5211_32W_dep) exist_fe5211_32'' exist_fe5211_32W fe5211_32ZToW. Defined.
+Definition exist_fe5211_32 (x : SpecificGen.GF5211_32.fe5211_32) : is_bounded x = true -> fe5211_32
+ := Eval cbv [exist_fe5211_32' exist_fe5211_32W exist_fe5211_32' app_fe5211_32 app_fe5211_32W_dep] in
+ exist_fe5211_32' x.
+
+Lemma proj1_fe5211_32_exist_fe5211_32W x pf : proj1_fe5211_32 (exist_fe5211_32W x pf) = fe5211_32WToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe5211_32W_exist_fe5211_32 x pf : proj1_fe5211_32W (exist_fe5211_32 x pf) = fe5211_32ZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_fe5211_32_exist_fe5211_32 x pf : proj1_fe5211_32 (exist_fe5211_32 x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_fe5211_32 exist_fe5211_32 proj1_fe5211_32W fe5211_32WToZ proj_word Build_bounded_word Build_bounded_word'].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition exist_wire_digitsW' (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
+Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
+Definition exist_wire_digitsW (x : wire_digitsW)
+ : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
+ := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
+Definition exist_wire_digits'' (x : SpecificGen.GF5211_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits' (x : SpecificGen.GF5211_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits.
+Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
+Definition exist_wire_digits (x : SpecificGen.GF5211_32.wire_digits)
+ : wire_digits_is_bounded x = true -> wire_digits
+ := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
+ exist_wire_digits' x.
+
+Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
+Proof. now hnf in x; destruct_head' prod. Qed.
+Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 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 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.
+End opt.
+
+Local Transparent bit_width.
+Local Ltac do_change lem :=
+ match lem with
+ | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
+ => 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
+ let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
+ do_change L'
+ end
+ | _ => lem
+ end.
+Definition fe5211_32_word64ize (x : fe5211_32) : fe5211_32.
+Proof.
+ set (x' := x).
+ hnf in x; destruct_head' prod.
+ let lem := constr:(exist_fe5211_32W (fe5211_32W_word64ize (proj1_fe5211_32W x'))) in
+ let lem := (eval cbv [proj1_fe5211_32W x' fe5211_32W_word64ize proj_word exist_fe5211_32W Build_bounded_word' Build_bounded_word] in lem) in
+ let lem := do_change lem in
+ refine (lem _);
+ change (is_bounded (fe5211_32WToZ (fe5211_32W_word64ize (proj1_fe5211_32W x'))) = true);
+ abstract (rewrite fe5211_32W_word64ize_id; apply is_bounded_proj1_fe5211_32).
+Defined.
+Definition wire_digits_word64ize (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 := do_change lem in
+ let lem := (eval cbv [word64ize opt.word64ize 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).
+Defined.
+
+Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
+ : forall n : nat,
+ (n < length limb_widths)%nat
+ -> (0 <= nth_default 0 (Tuple.to_list length_fe5211_32 x) n <=
+ snd (b_of (nth_default (-1) limb_widths n)))%Z.
+Proof.
+ hnf in x; destruct_head' prod.
+ unfold_is_bounded_in H; destruct_head' and.
+ Z.ltb_to_lt.
+ unfold nth_default; simpl.
+ intros.
+ repeat match goal with
+ | [ |- context[nth_error _ ?x] ]
+ => is_var x; destruct x; simpl
+ end;
+ omega.
+Qed.
+
+(* END precomputation *)
+
+(* Precompute constants *)
+
+Definition one' := Eval vm_compute in exist_fe5211_32 SpecificGen.GF5211_32.one_ eq_refl.
+Definition one := Eval cbv [one' fe5211_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe5211_32_word64ize one'.
+
+Definition zero' := Eval vm_compute in exist_fe5211_32 SpecificGen.GF5211_32.zero_ eq_refl.
+Definition zero := Eval cbv [zero' fe5211_32_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe5211_32_word64ize 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)
+ (Hid : id = F id')
+ (Hop_bounded : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> is_bounded (op (F x) (F y)) = true)
+ (Hop : forall x y, is_bounded (F x) = true
+ -> is_bounded (F y) = true
+ -> op (F x) (F y) = F (op' x y))
+ (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
+ : F (fold_chain_opt id' op' chain ls)
+ = fold_chain_opt id op chain (List.map F ls)
+ /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
+Proof.
+ rewrite !fold_chain_opt_correct.
+ revert dependent ls; induction chain as [|x xs IHxs]; intros.
+ { pose proof (Hls_bounded 0%nat).
+ destruct ls; simpl; split; trivial; congruence. }
+ { destruct x; simpl; unfold Let_In; simpl.
+ rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
+ { do 2 f_equal.
+ rewrite <- Hop, Hid by auto.
+ rewrite !map_nth_default_always.
+ split; try reflexivity.
+ apply (IHxs (_::_)).
+ intros [|?]; autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. }
+ { intros [|?]; simpl;
+ autorewrite with simpl_nth_default; auto.
+ rewrite <- Hop; auto. } }
+Qed.
+
+Lemma encode_bounded x : is_bounded (encode x) = true.
+Proof.
+ pose proof (bounded_encode x).
+ generalize dependent (encode x).
+ intro t; compute in t; intros.
+ destruct_head' prod.
+ unfold Pow2Base.bounded in H.
+ cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params5211_32] in H.
+ repeat match type of H with
+ | context[nth_error (cons _ _) _]
+ => let H' := fresh in
+ pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
+ cbv beta iota in H'
+ end.
+ clear H.
+ simpl in *.
+ cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
+ unfold is_bounded.
+ apply fold_right_andb_true_iff_fold_right_and_True.
+ cbv [is_bounded proj1_fe5211_32 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe5211_32].
+ repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
+Qed.
+
+Definition encode (x : F modulus) : fe5211_32
+ := exist_fe5211_32 (encode x) (encode_bounded x).
+
+Definition decode (x : fe5211_32) : F modulus
+ := ModularBaseSystem.decode (proj1_fe5211_32 x).
+
+Lemma proj1_fe5211_32_encode x
+ : proj1_fe5211_32 (encode x) = ModularBaseSystem.encode x.
+Proof.
+ cbv [encode].
+ generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
+ intros y pf; intros; hnf in y; destruct_head_hnf' prod.
+ cbv [proj1_fe5211_32 exist_fe5211_32 proj1_fe5211_32W Build_bounded_word Build_bounded_word' fe5211_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Lemma decode_exist_fe5211_32 x pf
+ : decode (exist_fe5211_32 x pf) = ModularBaseSystem.decode x.
+Proof.
+ hnf in x; destruct_head' prod.
+ cbv [decode proj1_fe5211_32 exist_fe5211_32 proj1_fe5211_32W Build_bounded_word Build_bounded_word' fe5211_32WToZ proj_word].
+ unfold_is_bounded_in pf.
+ destruct_head' and.
+ Z.ltb_to_lt.
+ rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
+ reflexivity.
+Qed.
+
+Definition div (f g : fe5211_32) : fe5211_32
+ := exist_fe5211_32 (div (proj1_fe5211_32 f) (proj1_fe5211_32 g)) (encode_bounded _).
+
+Definition eq (f g : fe5211_32) : Prop := eq (proj1_fe5211_32 f) (proj1_fe5211_32 g).
+
+
+Notation ibinop_correct_and_bounded irop op
+ := (forall x y,
+ is_bounded (fe5211_32WToZ x) = true
+ -> is_bounded (fe5211_32WToZ y) = true
+ -> fe5211_32WToZ (irop x y) = op (fe5211_32WToZ x) (fe5211_32WToZ y)
+ /\ is_bounded (fe5211_32WToZ (irop x y)) = true) (only parsing).
+Notation iunop_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe5211_32WToZ x) = true
+ -> fe5211_32WToZ (irop x) = op (fe5211_32WToZ x)
+ /\ is_bounded (fe5211_32WToZ (irop x)) = true) (only parsing).
+Notation iunop_FEToZ_correct irop op
+ := (forall x,
+ is_bounded (fe5211_32WToZ x) = true
+ -> word64ToZ (irop x) = op (fe5211_32WToZ x)) (only parsing).
+Notation iunop_FEToWire_correct_and_bounded irop op
+ := (forall x,
+ is_bounded (fe5211_32WToZ x) = true
+ -> wire_digitsWToZ (irop x) = op (fe5211_32WToZ x)
+ /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
+Notation iunop_WireToFE_correct_and_bounded irop op
+ := (forall x,
+ wire_digits_is_bounded (wire_digitsWToZ x) = true
+ -> fe5211_32WToZ (irop x) = op (wire_digitsWToZ x)
+ /\ is_bounded (fe5211_32WToZ (irop x)) = true) (only parsing).
+
+Definition prefreeze := GF5211_32.prefreeze.
diff --git a/src/SpecificGen/GF5211_32Reflective.v b/src/SpecificGen/GF5211_32Reflective.v
new file mode 100644
index 000000000..332d5c7e9
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective.v
@@ -0,0 +1,119 @@
+Require Import Crypto.BaseSystem.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
+Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystem.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Export Crypto.SpecificGen.GF5211_32.
+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.Reify.
+Require Import Crypto.Reflection.Z.Syntax.
+Require Import Crypto.SpecificGen.GF5211_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.
+Require Import Bedrock.Word Crypto.Util.WordUtil.
+Require Import Coq.Lists.List Crypto.Util.ListUtil.
+Require Import Crypto.Tactics.VerdiTactics.
+Require Import Crypto.Util.ZUtil.
+Require Import Crypto.Util.Tuple.
+Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.LetIn.
+Require Import Crypto.Util.Notations.
+Require Import Crypto.Util.Bool.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Algebra.
+Import ListNotations.
+Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
+Local Open Scope Z.
+
+Definition radd : ExprBinOp := Eval vm_compute in rcarry_addW.
+Definition rsub : ExprBinOp := Eval vm_compute in rcarry_subW.
+Definition rmul : ExprBinOp := Eval vm_compute in rmulW.
+Definition ropp : ExprUnOp := Eval vm_compute in rcarry_oppW.
+Definition rprefreeze : ExprUnOp := Eval vm_compute in rprefreezeW.
+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.
+
+Declare Reduction asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+Ltac asm_interp
+ := cbv beta iota delta
+ [id
+ 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
+ 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
+ Interp interp interp_flat_type interpf interp_flat_type fst snd].
+
+
+Definition interp_radd : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := Eval asm_interp in interp_bexpr (rword64ize radd).
+(*Print interp_radd.*)
+Definition interp_radd_correct : interp_radd = interp_bexpr radd := eq_refl.
+Definition interp_rsub : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rsub).
+(*Print interp_rsub.*)
+Definition interp_rsub_correct : interp_rsub = interp_bexpr rsub := eq_refl.
+Definition interp_rmul : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := Eval asm_interp in interp_bexpr (rword64ize rmul).
+(*Print interp_rmul.*)
+Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
+Definition interp_ropp : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := Eval asm_interp in interp_uexpr (rword64ize ropp).
+(*Print interp_ropp.*)
+Definition interp_ropp_correct : interp_ropp = interp_uexpr ropp := eq_refl.
+Definition interp_rprefreeze : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := Eval asm_interp in interp_uexpr (rword64ize rprefreeze).
+(*Print interp_rprefreeze.*)
+Definition interp_rprefreeze_correct : interp_rprefreeze = interp_uexpr rprefreeze := eq_refl.
+
+Definition interp_rge_modulus : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.word64
+ := Eval asm_interp in interp_uexpr_FEToZ (rword64ize rge_modulus).
+Definition interp_rge_modulus_correct : interp_rge_modulus = interp_uexpr_FEToZ rge_modulus := eq_refl.
+
+Definition interp_rpack : SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.wire_digitsW
+ := Eval asm_interp in interp_uexpr_FEToWire (rword64ize rpack).
+Definition interp_rpack_correct : interp_rpack = interp_uexpr_FEToWire rpack := eq_refl.
+
+Definition interp_runpack : SpecificGen.GF5211_32BoundedCommon.wire_digitsW -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := Eval asm_interp in interp_uexpr_WireToFE (rword64ize 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.
+Proof. exact rcarry_addW_correct_and_bounded. Qed.
+Lemma rsub_correct_and_bounded : binop_correct_and_bounded rsub carry_sub.
+Proof. exact rcarry_subW_correct_and_bounded. Qed.
+Lemma rmul_correct_and_bounded : binop_correct_and_bounded rmul mul.
+Proof. exact rmulW_correct_and_bounded. Qed.
+Lemma ropp_correct_and_bounded : unop_correct_and_bounded ropp carry_opp.
+Proof. exact rcarry_oppW_correct_and_bounded. Qed.
+Lemma rprefreeze_correct_and_bounded : unop_correct_and_bounded rprefreeze prefreeze.
+Proof. exact rprefreezeW_correct_and_bounded. Qed.
+Lemma rge_modulus_correct_and_bounded : unop_FEToZ_correct rge_modulus ge_modulus.
+Proof. exact rge_modulusW_correct_and_bounded. Qed.
+Lemma rpack_correct_and_bounded : unop_FEToWire_correct_and_bounded rpack pack.
+Proof. exact rpackW_correct_and_bounded. Qed.
+Lemma runpack_correct_and_bounded : unop_WireToFE_correct_and_bounded runpack unpack.
+Proof. exact runpackW_correct_and_bounded. Qed.
diff --git a/src/SpecificGen/GF5211_32Reflective/Common.v b/src/SpecificGen/GF5211_32Reflective/Common.v
new file mode 100644
index 000000000..a2b9a7c63
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Common.v
@@ -0,0 +1,438 @@
+Require Export Coq.ZArith.ZArith.
+Require Export Coq.Strings.String.
+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.Reify.
+Require Export Crypto.Reflection.Z.Syntax.
+Require Import Crypto.Reflection.InterpWfRel.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Reflection.MapInterpWf.
+Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.LetIn.
+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).
+
+Local Ltac make_type_from uncurried_op :=
+ let T := (type of uncurried_op) in
+ let T := (eval compute in T) in
+ let rT := reify_type T in
+ exact rT.
+
+Definition ExprBinOpT : type base_type.
+Proof. make_type_from (uncurry_binop_fe5211_32 carry_add). Defined.
+Definition ExprUnOpT : type base_type.
+Proof. make_type_from (uncurry_unop_fe5211_32 carry_opp). Defined.
+Definition ExprUnOpFEToZT : type base_type.
+Proof. make_type_from (uncurry_unop_fe5211_32 ge_modulus). Defined.
+Definition ExprUnOpWireToFET : type base_type.
+Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Definition ExprUnOpFEToWireT : type base_type.
+Proof. make_type_from (uncurry_unop_fe5211_32 pack). Defined.
+Definition ExprBinOp : Type := Expr ExprBinOpT.
+Definition ExprUnOp : Type := Expr ExprUnOpT.
+Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
+Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
+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 :: ?xs)%list => let bs := bounds_from_list xs in
+ constr:((Some {| ZBounds.lower := fst x ; ZBounds.upper := snd x |}, bs))
+ end.
+
+Local Ltac make_bounds ls :=
+ compute;
+ let v := bounds_from_list (List.rev ls) in
+ let v := (eval compute in v) in
+ exact v.
+
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
+Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
+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).
+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).
+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).
+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).
+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).
+
+Notation binop_correct_and_bounded rop op
+ := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+Notation unop_correct_and_bounded rop op
+ := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
+Notation unop_FEToZ_correct rop op
+ := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
+Notation unop_FEToWire_correct_and_bounded rop op
+ := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
+Notation unop_WireToFE_correct_and_bounded rop op
+ := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
+
+Ltac rexpr_cbv :=
+ lazymatch goal with
+ | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ => let operf := head oper in
+ let uncurryf := head uncurry in
+ try cbv delta [T]; try cbv delta [oper];
+ try cbv beta iota delta [uncurryf]
+ end;
+ cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+
+Ltac reify_sig :=
+ rexpr_cbv; eexists; Reify_rhs; reflexivity.
+
+Local Notation rexpr_sig T uncurried_op :=
+ { rexprZ
+ | interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op }
+ (only parsing).
+
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (uncurry_binop_fe5211_32 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT (uncurry_unop_fe5211_32 op)) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT (uncurry_unop_fe5211_32 op)) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT (uncurry_unop_fe5211_32 op)) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET (uncurry_unop_wire_digits op)) (only parsing).
+
+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
+ 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))
+ (only parsing).
+
+Ltac app_tuples x y :=
+ let tx := type of x in
+ lazymatch (eval hnf in tx) with
+ | prod _ _ => let xs := app_tuples (snd x) y in
+ constr:((fst x, xs))
+ | _ => constr:((x, y))
+ end.
+
+Local Arguments Tuple.map2 : simpl never.
+Local Arguments Tuple.map : simpl never.
+
+Fixpoint args_to_bounded_helperT {n}
+ (v : Tuple.tuple' Word64.word64 n)
+ (bounds : Tuple.tuple' (Z * Z) n)
+ (pf : List.fold_right
+ andb true
+ (Tuple.to_list
+ _
+ (Tuple.map2
+ (n:=S 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)
+ (res : Type)
+ {struct n}
+ : Type.
+Proof.
+ refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
+ List.fold_right
+ _ _ (Tuple.to_list
+ _
+ (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
+ | 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
+ end v bounds pf).
+ { clear -pf0.
+ abstract (
+ destruct v', bounds'; simpl @fst;
+ rewrite Tuple.map_S in pf0;
+ simpl in pf0;
+ rewrite Tuple.map2_S in pf0;
+ simpl @List.fold_right in *;
+ rewrite Bool.andb_true_iff in pf0; tauto
+ ). }
+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.
+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 |}
+ | 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) |}))
+ end.
+ { clear -pf pf'.
+ unfold Tuple.map2, Tuple.map in pf; simpl in *.
+ abstract (
+ destruct bounds;
+ simpl in *;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+ { simpl in *.
+ clear -pf pf'.
+ abstract (
+ destruct bounds as [? [? ?] ], v; simpl in *;
+ rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
+ simpl in pf;
+ rewrite !Bool.andb_true_iff in pf;
+ destruct_head' and;
+ Z.ltb_to_lt; auto
+ ). }
+Defined.
+
+Definition assoc_right''
+ := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
+
+Definition args_to_bounded {n} v bounds pf
+ := Eval cbv [args_to_bounded_helper assoc_right''] in
+ @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
+
+Local Ltac get_len T :=
+ match (eval hnf in T) with
+ | prod ?A ?B
+ => let a := get_len A in
+ let b := get_len B in
+ (eval compute in (a + b)%nat)
+ | _ => constr:(1%nat)
+ end.
+
+Local Ltac args_to_bounded x H :=
+ let x' := fresh in
+ set (x' := x);
+ compute in x;
+ let len := (let T := type of x in get_len T) in
+ destruct_head' prod;
+ let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
+ let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
+ let c := (eval cbv [all_binders_for ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
+ apply c; compute; clear;
+ try abstract (
+ repeat split;
+ solve [ reflexivity
+ | refine (fun v => match v with eq_refl => I end) ]
+ ).
+
+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).
+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).
+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).
+Proof.
+ let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
+ exact v.
+Defined.
+
+Ltac assoc_right_tuple x so_far :=
+ let t := type of x in
+ lazymatch (eval hnf in t) with
+ | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
+ assoc_right_tuple (fst x) so_far
+ | _ => lazymatch so_far with
+ | @None => x
+ | _ => constr:((x, so_far))
+ end
+ end.
+
+Local Ltac make_bounds_prop bounds orig_bounds :=
+ let bounds' := fresh "bounds'" in
+ let bounds_bad := fresh "bounds_bad" in
+ rename bounds into bounds_bad;
+ let boundsv := assoc_right_tuple bounds_bad (@None) in
+ pose boundsv as bounds;
+ pose orig_bounds as bounds';
+ repeat (refine (match fst bounds' with
+ | Some bounds' => let (l, u) := fst bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end && _)%bool;
+ destruct bounds' as [_ bounds'], bounds as [_ bounds]);
+ try exact (match bounds' with
+ | Some bounds' => let (l, u) := bounds in
+ let (l', u') := bounds' in
+ ((l' <=? l) && (u <=? u'))%Z%bool
+ | None => false
+ end).
+
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprBinOpT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToWireT)) : bool.
+Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpWireToFET)) : bool.
+Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
+(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders ExprUnOpFEToZT)) : bool.
+Proof.
+ refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
+Defined.
+
+(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
+ various kinds of correct and boundedness, and abstract in Gallina
+ rather than Ltac *)
+
+Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
+ let Heq := fresh "Heq" in
+ let Hbounds0 := fresh "Hbounds0" in
+ let Hbounds1 := fresh "Hbounds1" in
+ let Hbounds2 := fresh "Hbounds2" in
+ pose proof (proj2_sig ropZ_sig) as Heq;
+ cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE
+ curry_binop_fe5211_32W curry_unop_fe5211_32W curry_unop_wire_digitsW
+ curry_binop_fe5211_32 curry_unop_fe5211_32 curry_unop_wire_digits
+ uncurry_binop_fe5211_32W uncurry_unop_fe5211_32W uncurry_unop_wire_digitsW
+ uncurry_binop_fe5211_32 uncurry_unop_fe5211_32 uncurry_unop_wire_digits
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET
+ interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ cbv zeta in *;
+ simpl @fe5211_32WToZ; simpl @wire_digitsWToZ;
+ rewrite <- Heq; clear Heq;
+ destruct Hbounds as [Heq Hbounds];
+ change interp_op with (@Z.interp_op) in *;
+ 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;
+ specialize_by repeat first [ progress intros
+ | reflexivity
+ | assumption
+ | progress destruct_head' base_type
+ | progress destruct_head' BoundedWord64.BoundedWord
+ | progress destruct_head' and
+ | progress repeat apply conj ];
+ specialize (Hbounds_left args H0);
+ specialize (Hbounds_right args H0);
+ cbv beta in *;
+ lazymatch type of Hbounds_right with
+ | match ?e with _ => _ end
+ => lazymatch type of H1 with
+ | match ?e' with _ => _ end
+ => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
+ end
+ end;
+ 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;
+ match goal with
+ | [ |- fe5211_32WToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- wire_digitsWToZ ?x = _ /\ _ ]
+ => destruct x; destruct_head_hnf' prod
+ | [ |- _ = _ ]
+ => exact Hbounds_left
+ end;
+ change word64ToZ with Word64.word64ToZ in *;
+ (split; [ exact Hbounds_left | ]);
+ cbv [interp_flat_type] in *;
+ cbv [fst snd
+ binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good
+ ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds] in H1;
+ destruct_head' ZBounds.bounds;
+ unfold_is_bounded_in H1;
+ simpl @fe5211_32WToZ; simpl @wire_digitsWToZ;
+ unfold_is_bounded;
+ destruct_head' and;
+ Z.ltb_to_lt;
+ change Word64.word64ToZ with word64ToZ in *;
+ repeat apply conj; Z.ltb_to_lt; try omega; try reflexivity.
+
+
+Ltac rexpr_correct :=
+ let ropW' := fresh in
+ let ropZ_sig := fresh in
+ intros ropW' ropZ_sig;
+ let wf_ropW := fresh "wf_ropW" in
+ assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
+ cbv zeta; repeat apply conj;
+ [ vm_compute; reflexivity
+ | apply @InterpRelWf;
+ [ | 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 compute_bounds opW bounds
+ := (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp (@ZBounds.of_word64) opW)) bounds)
+ (only parsing).
+
+
+Module Export PrettyPrinting.
+ (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
+ 8.5/8.6. Because [Set] is special and things break if
+ [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
+ to debug. *)
+ Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
+
+ Inductive result := yes | no | borked.
+
+ Definition ZBounds_to_bounds_on
+ := 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 |}
+ => in_range l u
+ | None
+ => overflow
+ end
+ end.
+
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result
+ := match t return interp_flat_type (fun t => match t with TZ => bounds_on end) t -> result with
+ | Tbase TZ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
+ | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
+ | no, no => no
+ | yes, no | no, yes | yes, yes => yes
+ | borked, _ | _, borked => borked
+ end
+ end.
+
+ (** This gives a slightly easier to read version of the bounds *)
+ Notation compute_bounds_for_display opW bounds
+ := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
+ Notation sanity_compute opW bounds
+ := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
+ Notation sanity_check opW bounds
+ := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
+End PrettyPrinting.
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v b/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
new file mode 100644
index 000000000..a6a6e0f28
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
@@ -0,0 +1,50 @@
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprBinOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_binop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall xy
+ (xy := (eta_fe5211_32W (fst xy), eta_fe5211_32W (snd xy)))
+ (Hxy : is_bounded (fe5211_32WToZ (fst xy)) = true
+ /\ is_bounded (fe5211_32WToZ (snd xy)) = true),
+ let Hx := let (Hx, Hy) := Hxy in Hx in
+ 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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall xy
+ (xy := (eta_fe5211_32W (fst xy), eta_fe5211_32W (snd xy)))
+ (Hxy : is_bounded (fe5211_32WToZ (fst xy)) = true
+ /\ is_bounded (fe5211_32WToZ (snd xy)) = true),
+ 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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => binop_bounds_good bounds = true
+ | None => False
+ end)
+ : binop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x y Hx Hy.
+ pose x as x'; pose y as y'.
+ hnf in x, y; destruct_head' prod.
+ specialize (H0 (x', y') (conj Hx Hy)).
+ specialize (H1 (x', y') (conj Hx Hy)).
+ let args := constr:(binop_args_to_bounded (x', y') Hx Hy) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
new file mode 100644
index 000000000..9a1d5bdef
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOp_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe5211_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unop_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
new file mode 100644
index 000000000..14f3d3dd7
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToWire_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe5211_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToWire_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToWire_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
new file mode 100644
index 000000000..ab44167a6
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpFEToZ_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_fe5211_32W x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopFEToZ_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_FEToZ_correct (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unop_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
new file mode 100644
index 000000000..9948ea3a0
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
@@ -0,0 +1,44 @@
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Common.
+Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
+Require Import Crypto.Reflection.Z.Interpretations.
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.Application.
+Require Import Crypto.Reflection.MapInterp.
+Require Import Crypto.Util.Tactics.
+
+Local Opaque Interp.
+Lemma ExprUnOpWireToFE_correct_and_bounded
+ ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
+ (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
+ (H0 : forall x
+ (x := eta_wire_digitsW x)
+ (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))
+ (LiftOption.to' (Some args)))
+ with
+ | Some _ => True
+ | None => False
+ end)
+ (H1 : forall x
+ (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
+ match LiftOption.of'
+ (ApplyInterpedAll (Interp (@ZBounds.interp_op) (MapInterp ZBounds.of_word64 ropW)) (LiftOption.to' (Some x')))
+ with
+ | Some bounds => unopWireToFE_bounds_good bounds = true
+ | None => False
+ end)
+ : unop_WireToFE_correct_and_bounded (MapInterp (fun _ x => x) ropW) op.
+Proof.
+ intros x Hx.
+ pose x as x'.
+ hnf in x; destruct_head' prod.
+ specialize (H0 x' Hx).
+ specialize (H1 x' Hx).
+ let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
+ t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
+Qed.
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified.v b/src/SpecificGen/GF5211_32Reflective/Reified.v
new file mode 100644
index 000000000..a783188c9
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified.v
@@ -0,0 +1,13 @@
+(** We split the reification up into separate files, one operation per
+ file, so that it can run in parallel. *)
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.Add.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.CarryAdd.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.Sub.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.CarrySub.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.Mul.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.Opp.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.CarryOpp.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.PreFreeze.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.GeModulus.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.Pack.
+Require Export Crypto.SpecificGen.GF5211_32Reflective.Reified.Unpack.
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/Add.v b/src/SpecificGen/GF5211_32Reflective/Reified/Add.v
new file mode 100644
index 000000000..5ddd88689
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/Add.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonBinOp.
+
+Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
+Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
+Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
+Proof. rexpr_correct. Qed.
+Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
+Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
+Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/CarryAdd.v b/src/SpecificGen/GF5211_32Reflective/Reified/CarryAdd.v
new file mode 100644
index 000000000..3b17ad60c
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/CarryAdd.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonBinOp.
+
+Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
+Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
+Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_addW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
+Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/CarryOpp.v b/src/SpecificGen/GF5211_32Reflective/Reified/CarryOpp.v
new file mode 100644
index 000000000..faf4f7533
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/CarryOpp.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOp.
+
+Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
+Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
+Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_oppW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
+Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/CarrySub.v b/src/SpecificGen/GF5211_32Reflective/Reified/CarrySub.v
new file mode 100644
index 000000000..ad75f4530
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/CarrySub.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonBinOp.
+
+Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
+Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
+Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rcarry_subW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
+Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/GeModulus.v b/src/SpecificGen/GF5211_32Reflective/Reified/GeModulus.v
new file mode 100644
index 000000000..68b630622
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/GeModulus.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOpFEToZ.
+
+Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
+Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
+Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rge_modulusW_correct_and_bounded
+ := ExprUnOpFEToZ_correct_and_bounded
+ rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
+Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/Mul.v b/src/SpecificGen/GF5211_32Reflective/Reified/Mul.v
new file mode 100644
index 000000000..b6d2170de
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/Mul.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonBinOp.
+
+Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
+Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
+Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rmulW_correct_and_bounded
+ := ExprBinOp_correct_and_bounded
+ rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
+Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
+Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/Opp.v b/src/SpecificGen/GF5211_32Reflective/Reified/Opp.v
new file mode 100644
index 000000000..ae3389010
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/Opp.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOp.
+
+Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
+Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
+Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
+Proof. rexpr_correct. Qed.
+Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
+Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
+Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/Pack.v b/src/SpecificGen/GF5211_32Reflective/Reified/Pack.v
new file mode 100644
index 000000000..108ce3578
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/Pack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOpFEToWire.
+
+Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
+Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
+Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rpackW_correct_and_bounded
+ := ExprUnOpFEToWire_correct_and_bounded
+ rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
+Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v
new file mode 100644
index 000000000..4788f9f1a
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOp.
+
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF5211_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
+Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition rprefreezeW_correct_and_bounded
+ := ExprUnOp_correct_and_bounded
+ rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
+Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/Sub.v b/src/SpecificGen/GF5211_32Reflective/Reified/Sub.v
new file mode 100644
index 000000000..f0c1f5d27
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/Sub.v
@@ -0,0 +1,12 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonBinOp.
+
+Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
+Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
+Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
+Proof. rexpr_correct. Qed.
+Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
+
+Local Open Scope string_scope.
+Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
+Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
+Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/Unpack.v b/src/SpecificGen/GF5211_32Reflective/Reified/Unpack.v
new file mode 100644
index 000000000..1040ec6d2
--- /dev/null
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/Unpack.v
@@ -0,0 +1,17 @@
+Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOpWireToFE.
+
+Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
+Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
+Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
+Proof. rexpr_correct. Qed.
+Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
+Local Obligation Tactic := intros; vm_compute; constructor.
+Program Definition runpackW_correct_and_bounded
+ := ExprUnOpWireToFE_correct_and_bounded
+ runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
+ _ _.
+
+Local Open Scope string_scope.
+Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
+Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/SpecificGen/copy_bounds.sh b/src/SpecificGen/copy_bounds.sh
new file mode 100755
index 000000000..5c883ad18
--- /dev/null
+++ b/src/SpecificGen/copy_bounds.sh
@@ -0,0 +1,16 @@
+#!/bin/bash
+
+DIR="$( cd "$( dirname "${BASH_SOURCE[0]}" )" && pwd )"
+
+cd "$DIR"
+
+FILENAME="$1"
+V_FILE_STEM="${FILENAME%.*}"
+
+for ORIG in $(git ls-files "../Specific/**GF25519*.v" | grep -v "../Specific/GF25519.v"); do
+ NEW="$(echo "$ORIG" | sed s'|^../Specific|.|g' | sed s"|25519|${V_FILE_STEM}|g")"
+ 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 $?
+done