aboutsummaryrefslogtreecommitdiff
path: root/src/SpecificGen
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2017-01-11 21:02:50 -0500
committerGravatar Jason Gross <jasongross9@gmail.com>2017-03-01 11:45:47 -0500
commit6b3048c37ad348dc88ecc03ef892ecfb121bfa7f (patch)
tree351e5438c5664ab0caf08b9d5054f296ff4aa2ee /src/SpecificGen
parent80dc66a34fbf031bfac1214ccbb3bb1dcdef3d39 (diff)
Switch to fully uncurried form for reflection
This will eventually make a number of proofs easier. Unfortunately, the correctness lemmas for AddCoordinates and LadderStep no longer work (because of different arities), and there's a proof in Experiments/Ed25519 that I've admitted. The correctness lemmas will be easy to re-add when we have a more general version that handle arbitrary type shapes.
Diffstat (limited to 'src/SpecificGen')
-rw-r--r--src/SpecificGen/GF2213_32Bounded.v128
-rw-r--r--src/SpecificGen/GF2213_32BoundedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF2213_32BoundedCommon.v44
-rw-r--r--src/SpecificGen/GF2213_32BoundedExtendedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF2213_32Reflective.v8
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Common.v440
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Common9_4Op.v32
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonBinOp.v20
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOp.v7
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v7
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v7
-rw-r--r--src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v7
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/AddCoordinates.v188
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/LadderStep.v207
-rw-r--r--src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v2
-rw-r--r--src/SpecificGen/GF2213_32ReflectiveAddCoordinates.v5
-rw-r--r--src/SpecificGen/GF2519_32Bounded.v128
-rw-r--r--src/SpecificGen/GF2519_32BoundedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF2519_32BoundedCommon.v44
-rw-r--r--src/SpecificGen/GF2519_32BoundedExtendedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF2519_32Reflective.v8
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Common.v440
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Common9_4Op.v32
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonBinOp.v20
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOp.v7
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v7
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v7
-rw-r--r--src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v7
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/AddCoordinates.v188
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/LadderStep.v207
-rw-r--r--src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v2
-rw-r--r--src/SpecificGen/GF2519_32ReflectiveAddCoordinates.v5
-rw-r--r--src/SpecificGen/GF25519_32Bounded.v128
-rw-r--r--src/SpecificGen/GF25519_32BoundedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF25519_32BoundedCommon.v44
-rw-r--r--src/SpecificGen/GF25519_32BoundedExtendedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF25519_32Reflective.v8
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Common.v440
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Common9_4Op.v32
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonBinOp.v20
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOp.v7
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v7
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v7
-rw-r--r--src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v7
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/AddCoordinates.v188
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/LadderStep.v207
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v2
-rw-r--r--src/SpecificGen/GF25519_32ReflectiveAddCoordinates.v5
-rw-r--r--src/SpecificGen/GF25519_64Bounded.v128
-rw-r--r--src/SpecificGen/GF25519_64BoundedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF25519_64BoundedCommon.v44
-rw-r--r--src/SpecificGen/GF25519_64BoundedExtendedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF25519_64Reflective.v8
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Common.v440
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Common9_4Op.v32
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonBinOp.v20
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOp.v7
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v7
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v7
-rw-r--r--src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v7
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/AddCoordinates.v188
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/LadderStep.v207
-rw-r--r--src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v2
-rw-r--r--src/SpecificGen/GF25519_64ReflectiveAddCoordinates.v5
-rw-r--r--src/SpecificGen/GF41417_32Bounded.v128
-rw-r--r--src/SpecificGen/GF41417_32BoundedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF41417_32BoundedCommon.v44
-rw-r--r--src/SpecificGen/GF41417_32BoundedExtendedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF41417_32Reflective.v8
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Common.v440
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Common9_4Op.v32
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonBinOp.v20
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOp.v7
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v7
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v7
-rw-r--r--src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v7
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/AddCoordinates.v188
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/LadderStep.v207
-rw-r--r--src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v2
-rw-r--r--src/SpecificGen/GF41417_32ReflectiveAddCoordinates.v5
-rw-r--r--src/SpecificGen/GF5211_32Bounded.v128
-rw-r--r--src/SpecificGen/GF5211_32BoundedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF5211_32BoundedCommon.v44
-rw-r--r--src/SpecificGen/GF5211_32BoundedExtendedAddCoordinates.v3
-rw-r--r--src/SpecificGen/GF5211_32Reflective.v8
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Common.v440
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Common9_4Op.v32
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonBinOp.v20
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOp.v7
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v7
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v7
-rw-r--r--src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v7
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/AddCoordinates.v188
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/LadderStep.v207
-rw-r--r--src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v2
-rw-r--r--src/SpecificGen/GF5211_32ReflectiveAddCoordinates.v5
96 files changed, 3618 insertions, 3030 deletions
diff --git a/src/SpecificGen/GF2213_32Bounded.v b/src/SpecificGen/GF2213_32Bounded.v
index c98f234a0..c207dc557 100644
--- a/src/SpecificGen/GF2213_32Bounded.v
+++ b/src/SpecificGen/GF2213_32Bounded.v
@@ -25,13 +25,23 @@ Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.
Local Open Scope Z.
+Local Ltac cbv_tuple_map :=
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'].
+
+Local Ltac post_bounded_t :=
+ (* much pain and hackery to work around [Defined] taking forever *)
+ cbv_tuple_map;
+ let blem' := fresh "blem'" in
+ let is_bounded_lem := fresh "is_bounded_lem" in
+ intros is_bounded_lem blem';
+ apply blem'; repeat apply conj; apply is_bounded_lem.
Local Ltac bounded_t opW blem :=
- apply blem; apply is_bounded_proj1_fe2213_32.
+ generalize blem; generalize is_bounded_proj1_fe2213_32; post_bounded_t.
Local Ltac bounded_wire_digits_t opW blem :=
- apply blem; apply is_bounded_proj1_wire_digits.
+ generalize blem; generalize is_bounded_proj1_wire_digits; post_bounded_t.
Local Ltac define_binop f g opW blem :=
- refine (exist_fe2213_32W (opW (proj1_fe2213_32W f) (proj1_fe2213_32W g)) _);
+ 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)) _);
@@ -47,17 +57,17 @@ Local Ltac define_unop_WireToFE f 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_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 addW (f : fe2213_32W * fe2213_32W) : fe2213_32W := Eval simpl in interp_radd f.
+Definition subW (f : fe2213_32W * fe2213_32W) : fe2213_32W := Eval simpl in interp_rsub f.
+Definition mulW (f : fe2213_32W * fe2213_32W) : fe2213_32W := Eval simpl in interp_rmul f.
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.
@@ -86,7 +96,7 @@ Definition freezeW (f : fe2213_32W) : fe2213_32W := Eval cbv beta delta [prefree
Local Transparent Let_In.
(* Wrapper to allow extracted code to not unfold [mulW] *)
Definition mulW_noinline := mulW.
-Definition powW (f : fe2213_32W) chain := fold_chain_opt (proj1_fe2213_32W one) mulW_noinline chain [f].
+Definition powW (f : fe2213_32W) chain := fold_chain_opt (proj1_fe2213_32W one) (fun f g => mulW_noinline (f, g)) chain [f].
Definition invW (f : fe2213_32W) : fe2213_32W
:= Eval cbv -[Let_In fe2213_32W mulW_noinline] in powW f (chain inv_ec).
@@ -95,11 +105,11 @@ Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
rewrite pre_rewrite;
intros; apply rop_cb; assumption.
-Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW (Curry.curry2 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.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW (Curry.curry2 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.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW (Curry.curry2 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.
@@ -142,6 +152,7 @@ Proof.
cbv [modulusW Tuple.map].
cbv [on_tuple List.map to_list to_list' from_list from_list'
+ HList.hlistP HList.hlistP'
Tuple.map2 on_tuple2 ListUtil.map2 fe2213_32WToZ length_fe2213_32].
cbv [postfreeze GF2213_32.postfreeze].
cbv [Let_In].
@@ -203,7 +214,8 @@ Proof.
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.
+ split; [ | assumption ].
+ rewrite H0', H0; reflexivity.
Qed.
Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
@@ -212,9 +224,11 @@ Proof.
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; pose proof (fun k0 k1 X Y => proj1 (mulW_correct_and_bounded (k0, k1) (conj X Y))) as H'.
+ cbv [Curry.curry2 Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list'] in H'.
+ rewrite <- H' by assumption.
+ apply mulW_correct_and_bounded; split; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded (_, _) (conj X Y))) by assumption; reflexivity. }
{ intros [|?]; autorewrite with simpl_nth_default;
(assumption || reflexivity). }
Qed.
@@ -269,8 +283,10 @@ Proof.
unfold GF2213_32.eqb.
simpl @fe2213_32WToZ in *; cbv beta iota.
intros.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe2213_32WToZ] in *.
rewrite <- frf, <- frg by assumption.
- rewrite <- fieldwisebW_correct.
+ etransitivity; [ eapply fieldwisebW_correct | ].
+ cbv [fe2213_32WToZ].
reflexivity.
Defined.
@@ -297,7 +313,7 @@ Proof.
lazymatch (eval cbv delta [GF2213_32.sqrt] in GF2213_32.sqrt) with
| (fun powf powf_squared f => dlet a := powf in _)
=> exact (dlet powx := powW (fe2213_32ZToW x) (chain GF2213_32.sqrt_ec) in
- GF2213_32.sqrt (fe2213_32WToZ powx) (fe2213_32WToZ (mulW_noinline powx powx)) x)
+ GF2213_32.sqrt (fe2213_32WToZ powx) (fe2213_32WToZ (mulW_noinline (powx, powx))) x)
| (fun f => pow f _)
=> exact (GF2213_32.sqrt x)
end.
@@ -324,21 +340,41 @@ Proof.
=> 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
- | mulW_noinline _ _ => apply mulW_correct_and_bounded
- | powW _ _ => apply powW_correct_and_bounded
- | sqrt_m1W => vm_compute; reflexivity
- | _ => assumption
- end
- end;
- subst_evars; reflexivity
+ pose proof (fun X Y => proj1 (mulW_correct_and_bounded (sqrt_m1W, a) (conj X Y))) as correctness;
+ let cbv_in_all _ := (cbv [fe2213_32WToZ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe2213_32WToZ Curry.curry2 HList.hlistP HList.hlistP'] in *; idtac) in
+ cbv_in_all ();
+ let solver _ := (repeat match goal with
+ | _ => progress subst
+ | _ => progress unfold fst, snd
+ | _ => progress cbv_in_all ()
+ | [ |- ?x /\ ?x ] => cut x; [ intro; split; assumption | ]
+ | [ |- is_bounded ?op = true ]
+ => let H := fresh in
+ lazymatch op with
+ | context[mulW (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[mulW_noinline (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[powW _ _] => pose proof powW_correct_and_bounded as H
+ | context[sqrt_m1W] => vm_compute; reflexivity
+ | _ => assumption
+ end;
+ cbv_in_all ();
+ apply H
+ end) in
+ rewrite <- correctness by solver (); clear correctness;
+ let lem := fresh in
+ pose proof eqbW_correct as lem; cbv_in_all (); rewrite <- lem by solver (); clear lem;
+ pose proof (pull_bool_if fe2213_32WToZ) as lem; cbv_in_all (); rewrite lem by solver (); clear lem;
+ subst_evars; reflexivity
end.
} Unfocus.
+ assert (Hfold : forall x, fe2213_32WToZ x = fe2213_32WToZ x) by reflexivity.
+ unfold fe2213_32WToZ at 2 in Hfold.
+ etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody; [ reflexivity | intro ].
+ apply Hfold.
+ } Unfocus.
+ clear Hfold.
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
@@ -346,15 +382,22 @@ Proof.
[ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
| _ => idtac
end;
- reflexivity. }
- { cbv [Let_In];
+ reflexivity.
+ }
+
+ { cbv [Let_In HList.hlistP HList.hlistP'];
try break_if;
repeat lazymatch goal with
| [ |- is_bounded (?WToZ (powW _ _)) = true ]
=> apply powW_correct_and_bounded; assumption
- | [ |- is_bounded (?WToZ (mulW _ _)) = true ]
- => apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ]
- end. }
+ | [ |- is_bounded (snd (?WToZ (_, powW _ _))) = true ]
+ => generalize powW_correct_and_bounded;
+ cbv [snd Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'];
+ let H := fresh in intro H; apply H; assumption
+ | [ |- is_bounded (?WToZ (mulW (_, _))) = true ]
+ => apply mulW_correct_and_bounded; split; [ vm_compute; reflexivity | ]
+ end.
+ }
Defined.
Definition sqrtW (f : fe2213_32W) : fe2213_32W :=
@@ -394,7 +437,7 @@ 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.
+Proof. define_unop f invW (fun x p => proj2 (invW_correct_and_bounded x p)). Defined.
Definition sqrt (f : fe2213_32) : fe2213_32.
Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
@@ -407,7 +450,12 @@ Local Ltac op_correct_t op opW_correct_and_bounded :=
=> rewrite proj1_wire_digits_exist_wire_digitsW
| _ => idtac
end;
- apply opW_correct_and_bounded;
+ generalize opW_correct_and_bounded;
+ cbv_tuple_map;
+ cbv [fst snd];
+ let H := fresh in
+ intro H; apply H;
+ repeat match goal with |- and _ _ => apply conj end;
lazymatch goal with
| [ |- is_bounded _ = true ]
=> apply is_bounded_proj1_fe2213_32
@@ -434,7 +482,7 @@ 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.
+Proof. op_correct_t inv (fun x p => proj1 (invW_correct_and_bounded x p)). 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.
diff --git a/src/SpecificGen/GF2213_32BoundedAddCoordinates.v b/src/SpecificGen/GF2213_32BoundedAddCoordinates.v
index 6079aadc9..86d5b1db7 100644
--- a/src/SpecificGen/GF2213_32BoundedAddCoordinates.v
+++ b/src/SpecificGen/GF2213_32BoundedAddCoordinates.v
@@ -14,7 +14,7 @@ Local Ltac define_binop f g 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_coordinates / _ _ _ _ _ _ _ _ _.
+(*Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
Definition add_coordinatesW (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe2213_32W) : Tuple.tuple fe2213_32W 4
:= Eval simpl in interp_radd_coordinates x0 x1 x2 x3 x4 x5 x6 x7 x8.
@@ -75,3 +75,4 @@ Lemma add_coordinates_correct (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe2213_32)
(proj1_fe2213_32 x7)
(proj1_fe2213_32 x8).
Proof. op_correct_t add_coordinates add_coordinatesW_correct_and_bounded. Qed.
+*)
diff --git a/src/SpecificGen/GF2213_32BoundedCommon.v b/src/SpecificGen/GF2213_32BoundedCommon.v
index f9b215ac9..c1b109cc2 100644
--- a/src/SpecificGen/GF2213_32BoundedCommon.v
+++ b/src/SpecificGen/GF2213_32BoundedCommon.v
@@ -322,14 +322,14 @@ Ltac unfold_is_bounded_in' H :=
end.
Ltac preunfold_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.
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe2213_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
Ltac unfold_is_bounded_in H :=
preunfold_is_bounded_in H;
unfold_is_bounded_in' H.
Ltac preunfold_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].
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe2213_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
Ltac unfold_is_bounded :=
preunfold_is_bounded;
@@ -724,7 +724,7 @@ Notation in_op_correct_and_bounded k irop op
/\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe2213_32WToZ irop))%type)
(only parsing).
-Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
+(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
: forall (irop : Tower.tower_nd fe2213_32W (Tuple.tuple fe2213_32W count_out) count_in)
(op : Tower.tower_nd GF2213_32.fe2213_32 (Tuple.tuple GF2213_32.fe2213_32 count_out) count_in)
(cont : Prop -> Prop),
@@ -792,18 +792,14 @@ Qed.
Definition inm_op_correct_and_bounded1 count_in irop op
:= Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.
-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).
+ inm_op_correct_and_bounded count_in 1 irop op.*)
+Notation inm_op_correct_and_bounded n m irop op
+ := ((forall x,
+ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe2213_32WToZ x)
+ -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe2213_32WToZ x))))
+ (only parsing).
+Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
+Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe2213_32WToZ x) = true
@@ -818,20 +814,6 @@ Notation iunop_WireToFE_correct_and_bounded irop op
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).
-Notation i9top_correct_and_bounded k irop op
- := ((forall x0 x1 x2 x3 x4 x5 x6 x7 x8,
- is_bounded (fe2213_32WToZ x0) = true
- -> is_bounded (fe2213_32WToZ x1) = true
- -> is_bounded (fe2213_32WToZ x2) = true
- -> is_bounded (fe2213_32WToZ x3) = true
- -> is_bounded (fe2213_32WToZ x4) = true
- -> is_bounded (fe2213_32WToZ x5) = true
- -> is_bounded (fe2213_32WToZ x6) = true
- -> is_bounded (fe2213_32WToZ x7) = true
- -> is_bounded (fe2213_32WToZ x8) = true
- -> (Tuple.map (n:=k) fe2213_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)
- = op (fe2213_32WToZ x0) (fe2213_32WToZ x1) (fe2213_32WToZ x2) (fe2213_32WToZ x3) (fe2213_32WToZ x4) (fe2213_32WToZ x5) (fe2213_32WToZ x6) (fe2213_32WToZ x7) (fe2213_32WToZ x8))
- * HList.hlist (fun v => is_bounded v = true) (Tuple.map (n:=k) fe2213_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)))%type)
- (only parsing).
+Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-Definition prefreeze := GF2213_32.prefreeze.
+Notation prefreeze := GF2213_32.prefreeze.
diff --git a/src/SpecificGen/GF2213_32BoundedExtendedAddCoordinates.v b/src/SpecificGen/GF2213_32BoundedExtendedAddCoordinates.v
index 58c297578..9ef98db72 100644
--- a/src/SpecificGen/GF2213_32BoundedExtendedAddCoordinates.v
+++ b/src/SpecificGen/GF2213_32BoundedExtendedAddCoordinates.v
@@ -5,7 +5,7 @@ Require Import Crypto.SpecificGen.GF2213_32ExtendedAddCoordinates.
Require Import Crypto.SpecificGen.GF2213_32BoundedAddCoordinates.
Require Import Crypto.Util.Tuple.
Require Import Crypto.Util.Tactics.
-
+(*
Lemma fieldwise_eq_extended_add_coordinates_full' twice_d P10 P11 P12 P13 P20 P21 P22 P23
: Tuple.fieldwise
(n:=4) GF2213_32BoundedCommon.eq
@@ -65,3 +65,4 @@ Proof.
destruct_head' prod.
rewrite <- fieldwise_eq_extended_add_coordinates_full'; reflexivity.
Qed.
+*)
diff --git a/src/SpecificGen/GF2213_32Reflective.v b/src/SpecificGen/GF2213_32Reflective.v
index 078dc2c4a..e871b3a91 100644
--- a/src/SpecificGen/GF2213_32Reflective.v
+++ b/src/SpecificGen/GF2213_32Reflective.v
@@ -45,6 +45,7 @@ Declare Reduction asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2213_32W curry_unop_fe2213_32W curry_unop_wire_digitsW curry_9op_fe2213_32W
WordW.interp_op WordW.interp_base_type
@@ -56,6 +57,7 @@ Ltac asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2213_32W curry_unop_fe2213_32W curry_unop_wire_digitsW curry_9op_fe2213_32W
WordW.interp_op WordW.interp_base_type
@@ -65,15 +67,15 @@ Ltac asm_interp
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
+Definition interp_radd : SpecificGen.GF2213_32BoundedCommon.fe2213_32W * SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rsub : SpecificGen.GF2213_32BoundedCommon.fe2213_32W * SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rmul : SpecificGen.GF2213_32BoundedCommon.fe2213_32W * SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
:= Eval asm_interp in interp_bexpr rmul.
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
diff --git a/src/SpecificGen/GF2213_32Reflective/Common.v b/src/SpecificGen/GF2213_32Reflective/Common.v
index bcaa5d064..f1f4921b4 100644
--- a/src/SpecificGen/GF2213_32Reflective/Common.v
+++ b/src/SpecificGen/GF2213_32Reflective/Common.v
@@ -7,7 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Wf.
Require Import Crypto.Reflection.ExprInversion.
+Require Import Crypto.Reflection.Tuple.
Require Import Crypto.Reflection.Relations.
+Require Import Crypto.Reflection.Eta.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -15,13 +17,14 @@ Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Z.Syntax.Equality.
Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.Curry.
Require Import Crypto.Util.Tower.
Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Notations.
Notation Expr := (Expr base_type op).
@@ -43,30 +46,24 @@ Defined.
Definition Expr_n_OpT (count_out : nat) : flat_type base_type
:= Eval cbv [Syntax.tuple Syntax.tuple' fe2213_32T] in
Syntax.tuple fe2213_32T count_out.
-Fixpoint Expr_nm_OpT (count_in count_out : nat) : type base_type
- := match count_in with
- | 0 => Expr_n_OpT count_out
- | S n => SmartArrow fe2213_32T (Expr_nm_OpT n count_out)
- end.
+Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
+ := Eval cbv [Syntax.tuple Syntax.tuple' fe2213_32T Expr_n_OpT] in
+ Arrow (Syntax.tuple fe2213_32T count_in) (Expr_n_OpT count_out).
Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from (uncurry_unop_fe2213_32 ge_modulus). Defined.
+Proof. make_type_from ge_modulus. Defined.
Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Proof. make_type_from unpack. Defined.
Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from (uncurry_unop_fe2213_32 pack). Defined.
+Proof. make_type_from pack. Defined.
Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
Definition ExprArgT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpT.
+ := Eval compute in domain ExprUnOpT.
Definition ExprArgWireT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpFEToWireT.
-Definition ExprArgRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpT.
-Definition ExprArgWireRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Tbase TZ).
+ := Eval compute in domain ExprUnOpWireToFET.
+Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
@@ -75,99 +72,67 @@ Definition ExprBinOp : Type := Expr ExprBinOpT.
Definition ExprUnOp : Type := Expr ExprUnOpT.
Definition Expr4Op : Type := Expr Expr4OpT.
Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr ExprArgT.
-Definition ExprArgWire : Type := Expr ExprArgWireT.
-Definition ExprArgRev : Type := Expr ExprArgRevT.
-Definition ExprArgWireRev : Type := Expr ExprArgWireRevT.
+Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
+Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
Definition expr_nm_Op count_in count_out var : Type
:= expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Tbase TZ).
+Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) ExprArgT.
-Definition exprArgWire var : Type := expr base_type op (var:=var) ExprArgWireT.
-Definition exprArgRev var : Type := expr base_type op (var:=var) ExprArgRevT.
-Definition exprArgWireRev var : Type := expr base_type op (var:=var) ExprArgWireRevT.
-
-Local Ltac bounds_from_list_cps ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(fun T (extra : T) => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, extra))
- | (?x :: ?xs)%list => let bs := bounds_from_list_cps xs in
- constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
- end.
-
-Local Ltac make_bounds_cps ls extra :=
- let v := bounds_from_list_cps (List.rev ls) in
- let v := (eval compute in v) in
- exact (v _ extra).
-
-Local Ltac bounds_from_list ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
- | (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| Bounds.lower := fst x ; Bounds.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.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out : interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type.
-Proof.
- refine match count_in return interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type with
- | 0 => tt
- | S n => let v := interp_all_binders_for_to' (Expr_nm_Op_bounds n count_out) in
- interp_all_binders_for_of' _
- end; simpl.
- make_bounds_cps (Tuple.to_list _ bounds) v.
-Defined.
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
+Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
+
+Definition make_bound (x : Z * Z) : ZBounds.t
+ := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
+
+Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
+ := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
+ | 0 => tt
+ | S n
+ => let b := (Tuple.map make_bound bounds) in
+ let bs := Expr_nm_Op_bounds n count_out in
+ match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
+ | 0 => fun _ => b
+ | S n' => fun bs => (bs, b)
+ end bs
+ end.
+Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
:= Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
:= Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
:= Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
+Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
+ := Tuple.map make_bound wire_digit_bounds.
-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 (@WordW.interp_op) e).
+Definition interp_bexpr : ExprBinOp -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W * SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- := fun e => curry_unop_fe2213_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.word64
- := fun e => curry_unop_fe2213_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W -> SpecificGen.GF2213_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe2213_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF2213_32BoundedCommon.wire_digitsW -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_9_4expr : Expr9_4Op
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
- -> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
+ -> Tuple.tuple SpecificGen.GF2213_32BoundedCommon.fe2213_32W 9
-> Tuple.tuple SpecificGen.GF2213_32BoundedCommon.fe2213_32W 4
- := fun e => curry_9op_fe2213_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+ := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 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
@@ -181,40 +146,39 @@ Notation op9_4_correct_and_bounded rop op
Ltac rexpr_cbv :=
lazymatch goal with
- | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
=> let operf := head oper in
let uncurryf := head uncurry in
try cbv delta [T]; try cbv delta [oper];
try cbv beta iota delta [uncurryf]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
+ => let operf := head oper in
+ try cbv delta [T]; try cbv delta [oper]
end;
- cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+ cbv beta iota delta [interp_flat_type interp_base_type zero_ GF2213_32.fe2213_32 GF2213_32.wire_digits].
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 }
+ | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
(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 rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT (uncurry_9op_fe2213_32 op)) (only parsing).
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
+Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT 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 := ropW' in
- let ropZ := ropW' in
- let ropBounds := ropW' in
- let ropBoundedWordW := ropW' in
- ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
+ ropW' = proj1_sig ropZ_sig
+ /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
(only parsing).
Ltac app_tuples x y :=
@@ -227,7 +191,7 @@ Ltac app_tuples x y :=
Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
-
+(*
Fixpoint args_to_bounded_helperT {n}
(v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
@@ -299,14 +263,15 @@ Proof.
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
@@ -327,6 +292,7 @@ Ltac assoc_right_tuple x so_far :=
end
end.
+(*
Local Ltac make_args x :=
let x' := fresh "x'" in
compute in x |- *;
@@ -338,11 +304,6 @@ Local Ltac make_args x :=
let xv := assoc_right_tuple x'' (@None) in
refine (SmartVarf (xv : interp_flat_type _ t')).
-Definition unop_make_args {var} (x : exprArg var) : exprArgRev var.
-Proof. make_args x. Defined.
-Definition unop_wire_make_args {var} (x : exprArgWire var) : exprArgWireRev var.
-Proof. make_args x. Defined.
-
Local Ltac args_to_bounded x H :=
let x' := fresh in
set (x' := x);
@@ -351,29 +312,138 @@ Local Ltac args_to_bounded x H :=
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
+ let c := (eval cbv [domain 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) ]
).
+ *)
+
+Section gen.
+ Definition bounds_are_good_gen
+ {n : nat} (bounds : Tuple.tuple (Z * Z) n)
+ := let res :=
+ Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
+ in
+ List.fold_right andb true (Tuple.to_list n res).
+ Definition unop_args_to_bounded'
+ (bs : Z * Z)
+ (Hbs : bounds_are_good_gen (n:=1) bs = true)
+ (x : word64)
+ (H : is_bounded_gen (Tuple.map (fun v : word64 => (v : Z)) (n:=1) x) bs = true)
+ : BoundedWordW.BoundedWord.
+ Proof.
+ refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
+ unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
+ abstract (
+ destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
+ repeat apply conj; assumption
+ ).
+ Defined.
+ Fixpoint n_op_args_to_bounded'
+ n
+ : forall (bs : Tuple.tuple' (Z * Z) n)
+ (Hbs : bounds_are_good_gen (n:=S n) bs = true)
+ (x : Tuple.tuple' word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
+ Proof.
+ destruct n as [|n']; simpl in *.
+ { exact unop_args_to_bounded'. }
+ { refine (fun bs Hbs x H
+ => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
+ @unop_args_to_bounded' (snd bs) _ (snd x) _));
+ clear n_op_args_to_bounded';
+ simpl in *;
+ [ clear x H | clear Hbs | clear x H | clear Hbs ];
+ unfold bounds_are_good_gen, is_bounded_gen in *;
+ abstract (
+ repeat first [ progress simpl in *
+ | assumption
+ | reflexivity
+ | progress Bool.split_andb
+ | progress destruct_head prod
+ | match goal with
+ | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
+ end
+ | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
+ | progress break_match_hyps
+ | rewrite Bool.andb_true_iff; apply conj
+ | unfold Tuple.map, Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
+ ). }
+ Defined.
+
+ Definition n_op_args_to_bounded
+ n
+ : forall (bs : Tuple.tuple (Z * Z) n)
+ (Hbs : bounds_are_good_gen bs = true)
+ (x : Tuple.tuple word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
+ := match n with
+ | 0 => fun _ _ _ _ => tt
+ | S n' => @n_op_args_to_bounded' n'
+ end.
-Definition unop_args_to_bounded (x : fe2213_32W) (H : is_bounded (fe2213_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
+ Fixpoint nm_op_args_to_bounded' n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple' (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => @n_op_args_to_bounded m bs Hbs
+ | S n' => fun x H
+ => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
+ @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
+ end.
+ Definition nm_op_args_to_bounded n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => fun _ _ => tt
+ | S n' => @nm_op_args_to_bounded' n' m bs Hbs
+ end.
+End gen.
+
+Local Ltac get_inner_len T :=
+ lazymatch T with
+ | (?T * _)%type => get_inner_len T
+ | ?T => get_len T
+ end.
+Local Ltac get_outer_len T :=
+ lazymatch T with
+ | (?A * ?B)%type => let a := get_outer_len A in
+ let b := get_outer_len B in
+ (eval compute in (a + b)%nat)
+ | ?T => constr:(1%nat)
+ end.
+Local Ltac args_to_bounded x H :=
+ let t := type of x in
+ let m := get_inner_len t in
+ let n := get_outer_len t in
+ let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
+ let H := (eval cbv beta in (H eq_refl)) in
+ exact H.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe2213_32W * fe2213_32W)
(H : is_bounded (fe2213_32WToZ (fst x)) = true)
(H' : is_bounded (fe2213_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
-Proof.
- let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
- exact v.
-Defined.
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
+Proof. args_to_bounded x (conj H H'). Defined.
+Definition unop_args_to_bounded (x : fe2213_32W) (H : is_bounded (fe2213_32WToZ x) = true)
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain 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 _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
+Proof. args_to_bounded x H. Defined.
Definition op9_args_to_bounded (x : fe2213_32W * fe2213_32W * fe2213_32W * fe2213_32W * fe2213_32W * fe2213_32W * fe2213_32W * fe2213_32W * fe2213_32W)
(H0 : is_bounded (fe2213_32WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
(H1 : is_bounded (fe2213_32WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
@@ -384,58 +454,39 @@ Definition op9_args_to_bounded (x : fe2213_32W * fe2213_32W * fe2213_32W * fe221
(H6 : is_bounded (fe2213_32WToZ (snd (fst (fst x)))) = true)
(H7 : is_bounded (fe2213_32WToZ (snd (fst x))) = true)
(H8 : is_bounded (fe2213_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for Expr9_4OpT).
-Proof.
- let v := constr:(unop_args_to_bounded _ H8) in
- let v := app_tuples (unop_args_to_bounded _ H7) v in
- let v := app_tuples (unop_args_to_bounded _ H6) v in
- let v := app_tuples (unop_args_to_bounded _ H5) v in
- let v := app_tuples (unop_args_to_bounded _ H4) v in
- let v := app_tuples (unop_args_to_bounded _ H3) v in
- let v := app_tuples (unop_args_to_bounded _ H2) v in
- let v := app_tuples (unop_args_to_bounded _ H1) v in
- let v := app_tuples (unop_args_to_bounded _ H0) v in
- exact v.
-Defined.
-
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
+Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
+Local Ltac make_bounds_prop' bounds bounds' :=
+ first [ refine (andb _ _);
+ [ destruct bounds' as [bounds' _], bounds as [bounds _]
+ | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
+ try make_bounds_prop' bounds bounds'
+ | 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) ].
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.
+ make_bounds_prop' bounds bounds'.
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
Proof.
refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders Expr9_4OpT)) : bool.
+Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-
-Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
+(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe2213_32 f k)).
@@ -460,12 +511,11 @@ Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ v
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe2213_32 f k)).
-
+*)
(* 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
@@ -473,23 +523,25 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
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 interp_9_4expr
+ interp_flat_type_eta interp_flat_type_eta_gen
curry_binop_fe2213_32W curry_unop_fe2213_32W curry_unop_wire_digitsW curry_9op_fe2213_32W
curry_binop_fe2213_32 curry_unop_fe2213_32 curry_unop_wire_digits curry_9op_fe2213_32
uncurry_binop_fe2213_32W uncurry_unop_fe2213_32W uncurry_unop_wire_digitsW uncurry_9op_fe2213_32W
uncurry_binop_fe2213_32 uncurry_unop_fe2213_32 uncurry_unop_wire_digits uncurry_9op_fe2213_32
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT
- interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] 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 *;
+ change word64 with WordW.wordW in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
+ | progress unfold RelationClasses.Reflexive
| reflexivity
| assumption
| progress destruct_head' base_type
@@ -509,23 +561,33 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
+ Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
+ simpl @interp_flat_type in *;
+ (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
+ change (WordW.interp_base_type TZ) with v in *);
+ cbv beta iota zeta in *;
lazymatch goal with
| [ |- fe2213_32WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
| [ |- wire_digitsWToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
+ | [ |- (Tuple.map fe2213_32WToZ ?x = _) /\ _ ]
+ => generalize dependent x; intros
| [ |- ((Tuple.map fe2213_32WToZ ?x = _) * _)%type ]
=> generalize dependent x; intros
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type remove_all_binders] in *;
+ cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
destruct_head' prod;
change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
+ make_bound
+ Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
destruct_head' ZBounds.bounds;
@@ -534,12 +596,12 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
destruct_head' and;
Z.ltb_to_lt;
change WordW.wordWToZ with word64ToZ in *;
- cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd];
+ cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe2213_32WToZ HList.hlistP HList.hlistP'];
+ cbv [WordW.bit_width BitSize64.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
repeat split; unfold_is_bounded;
Z.ltb_to_lt;
try omega; try reflexivity.
-
Ltac rexpr_correct :=
let ropW' := fresh in
let ropZ_sig := fresh in
@@ -555,9 +617,13 @@ Ltac rexpr_correct :=
Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) opW) bounds)
+ := (Interp (@ZBounds.interp_op) opW bounds)
(only parsing).
+Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
+
+Ltac prove_rexpr_wfT
+ := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
Module Export PrettyPrinting.
(* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
@@ -569,23 +635,21 @@ Module Export PrettyPrinting.
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 {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end
+ := fun (t : base_type) (x : ZBounds.interp_base_type t)
+ => match x with
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
+ => in_range l u
+ | None
+ => overflow
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
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
+ := match t return interp_flat_type _ t -> result with
+ | Tbase _ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
| Unit => fun _ => no
| Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
| no, no => no
diff --git a/src/SpecificGen/GF2213_32Reflective/Common9_4Op.v b/src/SpecificGen/GF2213_32Reflective/Common9_4Op.v
index fb8b27f18..587eff7a7 100644
--- a/src/SpecificGen/GF2213_32Reflective/Common9_4Op.v
+++ b/src/SpecificGen/GF2213_32Reflective/Common9_4Op.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -42,8 +41,8 @@ Lemma Expr9_4Op_correct_and_bounded
let (Hx7, Hx8) := (eta_and Hx78) in
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -80,29 +79,24 @@ Lemma Expr9_4Op_correct_and_bounded
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => op9_4_bounds_good bounds = true
| None => False
end)
: op9_4_correct_and_bounded ropW op.
Proof.
- intros x0 x1 x2 x3 x4 x5 x6 x7 x8.
- intros Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8.
- pose x0 as x0'.
- pose x1 as x1'.
- pose x2 as x2'.
- pose x3 as x3'.
- pose x4 as x4'.
- pose x5 as x5'.
- pose x6 as x6'.
- pose x7 as x7'.
- pose x8 as x8'.
- hnf in x0, x1, x2, x3, x4, x5, x6, x7, x8; destruct_head' prod.
- specialize (H0 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ intros xs Hxs.
+ pose xs as xs'.
+ compute in xs.
+ destruct_head' prod.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
+ pose Hxs as Hxs'.
+ destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
+ specialize (H0 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ specialize (H1 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded (x0', x1', x2', x3', x4', x5', x6', x7', x8') Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
+ Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.64s) (successful) *)*)
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v b/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
index 8ad7f623c..3d38818f4 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonBinOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -18,8 +17,8 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -33,18 +32,19 @@ Lemma ExprBinOp_correct_and_bounded
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
end)
: binop_correct_and_bounded 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
+ intros xy HxHy.
+ pose xy as xy'.
+ compute in xy; destruct_head' prod.
+ specialize (H0 xy' HxHy).
+ specialize (H1 xy' HxHy).
+ destruct HxHy as [Hx Hy].
+ let args := constr:(binop_args_to_bounded xy' 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
index 40b39edf4..8f8035405 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOp_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
index 60c0726dc..76f1635df 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToWire.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
index bc5a6e265..95dbd10cd 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpFEToZ.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe2213_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
index e4f57d7bb..45094aa32 100644
--- a/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF2213_32Reflective/CommonUnOpWireToFE.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2213_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
let args := unopWireToFE_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/AddCoordinates.v b/src/SpecificGen/GF2213_32Reflective/Reified/AddCoordinates.v
index 84ac36861..0230b3427 100644
--- a/src/SpecificGen/GF2213_32Reflective/Reified/AddCoordinates.v
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/AddCoordinates.v
@@ -7,7 +7,6 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.ExprInversion.
Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
@@ -17,7 +16,10 @@ Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
+Require Import Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.Mul.
@@ -33,24 +35,28 @@ Require Import Bedrock.Word.
Definition radd_coordinatesZ' var twice_d P1 P2
:= @Extended.add_coordinates_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
twice_d _
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
P1 P2.
+Local Notation eta x := (fst x, snd x).
+
Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => radd_coordinatesZ'
- var (Return twice_d)
- (Return P10, Return P11, Return P12, Return P13)
- (Return P20, Return P21, Return P22, Return P23))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
+ => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
+ := twice_d_P1_P2 in
+ radd_coordinatesZ'
+ var (SmartVarf twice_d)
+ (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
+ (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
Definition add_coordinates
:= fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
@@ -59,70 +65,60 @@ Definition add_coordinates
twice_d (P10, P11, P12, P13) (P20, P21, P22, P23).
Definition uncurried_add_coordinates
- := apply9_nd
- (@uncurry_unop_fe2213_32)
- add_coordinates.
+ := fun twice_d_P1_P2
+ => let twice_d := fst twice_d_P1_P2 in
+ let (P1, P2) := eta (snd twice_d_P1_P2) in
+ @Extended.add_coordinates
+ _ add sub mul
+ twice_d P1 P2.
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
let H := fresh in
pose k as k';
assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
+ repeat (unfold k'; change k with k'; unfold interpf_step);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF2213_32.fe2213_32 -> GF2213_32.fe2213_32 -> GF2213_32.fe2213_32)
- (x' y' : GF2213_32.fe2213_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe2213_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
+Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe2213_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF2213_32.fe2213_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe2213_32 Apply' interp fst snd].
- rewrite (invert_Abs_Some (e:=f) eq_refl).
+ cbv [radd_coordinatesZ''].
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type] in x.
+ destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
+ | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
end.
- reflexivity.
-Qed.
-
-Lemma radd_coordinatesZ_sigP : rexpr_sigP Expr9_4OpT uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT radd_coordinatesZ'' uncurried_add_coordinates uncurry_unop_fe2213_32 SmartAbs radd_coordinatesZ' exprArg Extended.add_coordinates_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe2213_32 add_coordinates].
- intros.
- unfold invert_Return at 13 14 15 16.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [radd_coordinatesZ' add_coordinates Extended.add_coordinates_gen uncurried_add_coordinates SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -132,48 +128,56 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 13 14 15; unfold interpf_step.
- cbv beta iota delta [Extended.add_coordinates].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition radd_coordinatesZ_sig : rexpr_9_4op_sig add_coordinates.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv beta iota delta [Extended.add_coordinates interp_flat_type interp_base_type GF2213_32.fe2213_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- eexists.
- apply radd_coordinatesZ_sigP.
-Defined.
+ exact radd_coordinatesZ_sigP'.
+Qed.
+Definition radd_coordinatesZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
+
+Definition radd_coordinates_input_bounds
+ := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW Expr9Op_bounds.
+Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition radd_coordinatesW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- radd_coordinatesW add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
+ radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW Expr9Op_bounds).
+Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/LadderStep.v b/src/SpecificGen/GF2213_32Reflective/Reified/LadderStep.v
index f4fa2daff..7136ce80a 100644
--- a/src/SpecificGen/GF2213_32Reflective/Reified/LadderStep.v
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/LadderStep.v
@@ -7,8 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Relations.
Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -18,6 +19,7 @@ Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
Require Import Crypto.Spec.MxDH.
+Require Import Crypto.SpecificGen.GF2213_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF2213_32Reflective.Reified.Mul.
@@ -30,50 +32,80 @@ Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
Require Import Bedrock.Word.
-(* XXX FIXME: Remove dummy code *)
-Definition rladderstepZ' var (T:=_) (dummy0 dummy1 dummy2 a24 x0 : T) P1 P2
+Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
:= @MxDH.ladderstep_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
a24
_
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
x0
P1 P2.
Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21
- => rladderstepZ'
- var (Return dummy0) (Return dummy1) (Return dummy2)
- (Return a24) (Return x0)
- (Return P10, Return P11)
- (Return P20, Return P21))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
+ => let '(a24, x0, ((P10, P11), (P20, P21)))
+ := a24_x0_P1_P2 in
+ rladderstepZ'
+ var (SmartVarf a24) (SmartVarf x0)
+ (SmartVarf P10, SmartVarf P11)
+ (SmartVarf P20, SmartVarf P21)))).
-Definition ladderstep (T := _)
- := fun (dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21 : T)
- => @MxDH.ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
+Local Notation eta x := (fst x, snd x).
+
+Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
+ Eval cbv beta delta [MxDH.ladderstep_gen] in
+ @MxDH.ladderstep_gen
+ F Fadd Fsub Fmul a24
+ (F*F*F*F)
+ (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
+ (fun x f => dlet y := x in f y)
+ X1 P1 P2.
Definition uncurried_ladderstep
- := apply9_nd
- (@uncurry_unop_fe2213_32)
- ladderstep.
+ := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
+ => let a24 := fst (fst a24_x0_P1_P2) in
+ let x0 := snd (fst a24_x0_P1_P2) in
+ let '(P1, P2) := eta (snd a24_x0_P1_P2) in
+ let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
+ @ladderstep_other_assoc
+ _ add sub mul
+ a24 x0 (P10, P11) (P20, P21).
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
@@ -83,51 +115,14 @@ Local Ltac repeat_step_interpf :=
repeat (unfold interpf_step at 1; change k with k' at 1);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF2213_32.fe2213_32 -> GF2213_32.fe2213_32 -> GF2213_32.fe2213_32)
- (x' y' : GF2213_32.fe2213_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe2213_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
-Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe2213_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF2213_32.fe2213_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe2213_32 Apply' interp fst snd].
- let f := match goal with f : expr _ _ _ |- _ => f end in
- rewrite (invert_Abs_Some (e:=f) eq_refl).
- repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
- end.
- reflexivity.
-Qed.
-
-Lemma rladderstepZ_sigP : rexpr_sigP Expr9_4OpT uncurried_ladderstep rladderstepZ''.
+Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
cbv [rladderstepZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT rladderstepZ'' uncurried_ladderstep uncurry_unop_fe2213_32 SmartAbs rladderstepZ' exprArg MxDH.ladderstep_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe2213_32 ladderstep].
- intros; cbv beta zeta.
- unfold invert_Return at 14 15 16 17.
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type interp_base_type] in x.
+ destruct_head' prod.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -137,48 +132,58 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 14 15 16; unfold interpf_step.
- cbv beta iota delta [MxDH.ladderstep_other_assoc].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition rladderstepZ_sig : rexpr_9_4op_sig ladderstep.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv [ladderstep_other_assoc interp_flat_type GF2213_32.fe2213_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In Common.ExprBinOpT] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | _ => progress cbv [interp_base_type]
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
- eexists.
- apply rladderstepZ_sigP.
-Defined.
+ exact rladderstepZ_sigP'.
+Qed.
+Definition rladderstepZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
+
+Definition rladderstep_input_bounds
+ := (ExprUnOp_bounds, ExprUnOp_bounds,
+ ((ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW Expr9Op_bounds.
+Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition rladderstepW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- rladderstepW ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
+ rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW Expr9Op_bounds).
+Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v
index 60ab908b0..359a9c432 100644
--- a/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v
+++ b/src/SpecificGen/GF2213_32Reflective/Reified/PreFreeze.v
@@ -1,6 +1,6 @@
Require Import Crypto.SpecificGen.GF2213_32Reflective.CommonUnOp.
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF2213_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. 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.
diff --git a/src/SpecificGen/GF2213_32ReflectiveAddCoordinates.v b/src/SpecificGen/GF2213_32ReflectiveAddCoordinates.v
index aaf0c2833..dbe640f39 100644
--- a/src/SpecificGen/GF2213_32ReflectiveAddCoordinates.v
+++ b/src/SpecificGen/GF2213_32ReflectiveAddCoordinates.v
@@ -32,7 +32,7 @@ Import ListNotations.
Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Local Open Scope Z.
-Time Definition radd_coordinates : Expr9_4Op := Eval vm_compute in radd_coordinatesW.
+Time Definition radd_coordinates : Expr _ := Eval vm_compute in radd_coordinatesW.
Declare Reduction asm_interp_add_coordinates
:= cbv beta iota delta
@@ -57,7 +57,7 @@ Ltac asm_interp_add_coordinates
mapf_interp_flat_type WordW.interp_base_type word64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
-
+(*
Time Definition interp_radd_coordinates : SpecificGen.GF2213_32BoundedCommon.fe2213_32W
-> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
-> SpecificGen.GF2213_32BoundedCommon.fe2213_32W
@@ -74,3 +74,4 @@ Time Definition interp_radd_coordinates_correct : interp_radd_coordinates = inte
Lemma radd_coordinates_correct_and_bounded : op9_4_correct_and_bounded radd_coordinates add_coordinates.
Proof. exact_no_check radd_coordinatesW_correct_and_bounded. Time Qed.
+*)
diff --git a/src/SpecificGen/GF2519_32Bounded.v b/src/SpecificGen/GF2519_32Bounded.v
index 9cb723690..83dcec8a4 100644
--- a/src/SpecificGen/GF2519_32Bounded.v
+++ b/src/SpecificGen/GF2519_32Bounded.v
@@ -25,13 +25,23 @@ Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.
Local Open Scope Z.
+Local Ltac cbv_tuple_map :=
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'].
+
+Local Ltac post_bounded_t :=
+ (* much pain and hackery to work around [Defined] taking forever *)
+ cbv_tuple_map;
+ let blem' := fresh "blem'" in
+ let is_bounded_lem := fresh "is_bounded_lem" in
+ intros is_bounded_lem blem';
+ apply blem'; repeat apply conj; apply is_bounded_lem.
Local Ltac bounded_t opW blem :=
- apply blem; apply is_bounded_proj1_fe2519_32.
+ generalize blem; generalize is_bounded_proj1_fe2519_32; post_bounded_t.
Local Ltac bounded_wire_digits_t opW blem :=
- apply blem; apply is_bounded_proj1_wire_digits.
+ generalize blem; generalize is_bounded_proj1_wire_digits; post_bounded_t.
Local Ltac define_binop f g opW blem :=
- refine (exist_fe2519_32W (opW (proj1_fe2519_32W f) (proj1_fe2519_32W g)) _);
+ 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)) _);
@@ -47,17 +57,17 @@ Local Ltac define_unop_WireToFE f 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_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 addW (f : fe2519_32W * fe2519_32W) : fe2519_32W := Eval simpl in interp_radd f.
+Definition subW (f : fe2519_32W * fe2519_32W) : fe2519_32W := Eval simpl in interp_rsub f.
+Definition mulW (f : fe2519_32W * fe2519_32W) : fe2519_32W := Eval simpl in interp_rmul f.
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.
@@ -86,7 +96,7 @@ Definition freezeW (f : fe2519_32W) : fe2519_32W := Eval cbv beta delta [prefree
Local Transparent Let_In.
(* Wrapper to allow extracted code to not unfold [mulW] *)
Definition mulW_noinline := mulW.
-Definition powW (f : fe2519_32W) chain := fold_chain_opt (proj1_fe2519_32W one) mulW_noinline chain [f].
+Definition powW (f : fe2519_32W) chain := fold_chain_opt (proj1_fe2519_32W one) (fun f g => mulW_noinline (f, g)) chain [f].
Definition invW (f : fe2519_32W) : fe2519_32W
:= Eval cbv -[Let_In fe2519_32W mulW_noinline] in powW f (chain inv_ec).
@@ -95,11 +105,11 @@ Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
rewrite pre_rewrite;
intros; apply rop_cb; assumption.
-Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW (Curry.curry2 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.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW (Curry.curry2 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.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW (Curry.curry2 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.
@@ -142,6 +152,7 @@ Proof.
cbv [modulusW Tuple.map].
cbv [on_tuple List.map to_list to_list' from_list from_list'
+ HList.hlistP HList.hlistP'
Tuple.map2 on_tuple2 ListUtil.map2 fe2519_32WToZ length_fe2519_32].
cbv [postfreeze GF2519_32.postfreeze].
cbv [Let_In].
@@ -203,7 +214,8 @@ Proof.
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.
+ split; [ | assumption ].
+ rewrite H0', H0; reflexivity.
Qed.
Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
@@ -212,9 +224,11 @@ Proof.
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; pose proof (fun k0 k1 X Y => proj1 (mulW_correct_and_bounded (k0, k1) (conj X Y))) as H'.
+ cbv [Curry.curry2 Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list'] in H'.
+ rewrite <- H' by assumption.
+ apply mulW_correct_and_bounded; split; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded (_, _) (conj X Y))) by assumption; reflexivity. }
{ intros [|?]; autorewrite with simpl_nth_default;
(assumption || reflexivity). }
Qed.
@@ -269,8 +283,10 @@ Proof.
unfold GF2519_32.eqb.
simpl @fe2519_32WToZ in *; cbv beta iota.
intros.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe2519_32WToZ] in *.
rewrite <- frf, <- frg by assumption.
- rewrite <- fieldwisebW_correct.
+ etransitivity; [ eapply fieldwisebW_correct | ].
+ cbv [fe2519_32WToZ].
reflexivity.
Defined.
@@ -297,7 +313,7 @@ Proof.
lazymatch (eval cbv delta [GF2519_32.sqrt] in GF2519_32.sqrt) with
| (fun powf powf_squared f => dlet a := powf in _)
=> exact (dlet powx := powW (fe2519_32ZToW x) (chain GF2519_32.sqrt_ec) in
- GF2519_32.sqrt (fe2519_32WToZ powx) (fe2519_32WToZ (mulW_noinline powx powx)) x)
+ GF2519_32.sqrt (fe2519_32WToZ powx) (fe2519_32WToZ (mulW_noinline (powx, powx))) x)
| (fun f => pow f _)
=> exact (GF2519_32.sqrt x)
end.
@@ -324,21 +340,41 @@ Proof.
=> 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
- | mulW_noinline _ _ => apply mulW_correct_and_bounded
- | powW _ _ => apply powW_correct_and_bounded
- | sqrt_m1W => vm_compute; reflexivity
- | _ => assumption
- end
- end;
- subst_evars; reflexivity
+ pose proof (fun X Y => proj1 (mulW_correct_and_bounded (sqrt_m1W, a) (conj X Y))) as correctness;
+ let cbv_in_all _ := (cbv [fe2519_32WToZ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe2519_32WToZ Curry.curry2 HList.hlistP HList.hlistP'] in *; idtac) in
+ cbv_in_all ();
+ let solver _ := (repeat match goal with
+ | _ => progress subst
+ | _ => progress unfold fst, snd
+ | _ => progress cbv_in_all ()
+ | [ |- ?x /\ ?x ] => cut x; [ intro; split; assumption | ]
+ | [ |- is_bounded ?op = true ]
+ => let H := fresh in
+ lazymatch op with
+ | context[mulW (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[mulW_noinline (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[powW _ _] => pose proof powW_correct_and_bounded as H
+ | context[sqrt_m1W] => vm_compute; reflexivity
+ | _ => assumption
+ end;
+ cbv_in_all ();
+ apply H
+ end) in
+ rewrite <- correctness by solver (); clear correctness;
+ let lem := fresh in
+ pose proof eqbW_correct as lem; cbv_in_all (); rewrite <- lem by solver (); clear lem;
+ pose proof (pull_bool_if fe2519_32WToZ) as lem; cbv_in_all (); rewrite lem by solver (); clear lem;
+ subst_evars; reflexivity
end.
} Unfocus.
+ assert (Hfold : forall x, fe2519_32WToZ x = fe2519_32WToZ x) by reflexivity.
+ unfold fe2519_32WToZ at 2 in Hfold.
+ etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody; [ reflexivity | intro ].
+ apply Hfold.
+ } Unfocus.
+ clear Hfold.
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
@@ -346,15 +382,22 @@ Proof.
[ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
| _ => idtac
end;
- reflexivity. }
- { cbv [Let_In];
+ reflexivity.
+ }
+
+ { cbv [Let_In HList.hlistP HList.hlistP'];
try break_if;
repeat lazymatch goal with
| [ |- is_bounded (?WToZ (powW _ _)) = true ]
=> apply powW_correct_and_bounded; assumption
- | [ |- is_bounded (?WToZ (mulW _ _)) = true ]
- => apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ]
- end. }
+ | [ |- is_bounded (snd (?WToZ (_, powW _ _))) = true ]
+ => generalize powW_correct_and_bounded;
+ cbv [snd Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'];
+ let H := fresh in intro H; apply H; assumption
+ | [ |- is_bounded (?WToZ (mulW (_, _))) = true ]
+ => apply mulW_correct_and_bounded; split; [ vm_compute; reflexivity | ]
+ end.
+ }
Defined.
Definition sqrtW (f : fe2519_32W) : fe2519_32W :=
@@ -394,7 +437,7 @@ 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.
+Proof. define_unop f invW (fun x p => proj2 (invW_correct_and_bounded x p)). Defined.
Definition sqrt (f : fe2519_32) : fe2519_32.
Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
@@ -407,7 +450,12 @@ Local Ltac op_correct_t op opW_correct_and_bounded :=
=> rewrite proj1_wire_digits_exist_wire_digitsW
| _ => idtac
end;
- apply opW_correct_and_bounded;
+ generalize opW_correct_and_bounded;
+ cbv_tuple_map;
+ cbv [fst snd];
+ let H := fresh in
+ intro H; apply H;
+ repeat match goal with |- and _ _ => apply conj end;
lazymatch goal with
| [ |- is_bounded _ = true ]
=> apply is_bounded_proj1_fe2519_32
@@ -434,7 +482,7 @@ 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.
+Proof. op_correct_t inv (fun x p => proj1 (invW_correct_and_bounded x p)). 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.
diff --git a/src/SpecificGen/GF2519_32BoundedAddCoordinates.v b/src/SpecificGen/GF2519_32BoundedAddCoordinates.v
index 68b97f000..6ff1a8198 100644
--- a/src/SpecificGen/GF2519_32BoundedAddCoordinates.v
+++ b/src/SpecificGen/GF2519_32BoundedAddCoordinates.v
@@ -14,7 +14,7 @@ Local Ltac define_binop f g 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_coordinates / _ _ _ _ _ _ _ _ _.
+(*Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
Definition add_coordinatesW (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe2519_32W) : Tuple.tuple fe2519_32W 4
:= Eval simpl in interp_radd_coordinates x0 x1 x2 x3 x4 x5 x6 x7 x8.
@@ -75,3 +75,4 @@ Lemma add_coordinates_correct (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe2519_32)
(proj1_fe2519_32 x7)
(proj1_fe2519_32 x8).
Proof. op_correct_t add_coordinates add_coordinatesW_correct_and_bounded. Qed.
+*)
diff --git a/src/SpecificGen/GF2519_32BoundedCommon.v b/src/SpecificGen/GF2519_32BoundedCommon.v
index f6bf77cb0..b37b6bc4e 100644
--- a/src/SpecificGen/GF2519_32BoundedCommon.v
+++ b/src/SpecificGen/GF2519_32BoundedCommon.v
@@ -322,14 +322,14 @@ Ltac unfold_is_bounded_in' H :=
end.
Ltac preunfold_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.
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe2519_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
Ltac unfold_is_bounded_in H :=
preunfold_is_bounded_in H;
unfold_is_bounded_in' H.
Ltac preunfold_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].
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe2519_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
Ltac unfold_is_bounded :=
preunfold_is_bounded;
@@ -724,7 +724,7 @@ Notation in_op_correct_and_bounded k irop op
/\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe2519_32WToZ irop))%type)
(only parsing).
-Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
+(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
: forall (irop : Tower.tower_nd fe2519_32W (Tuple.tuple fe2519_32W count_out) count_in)
(op : Tower.tower_nd GF2519_32.fe2519_32 (Tuple.tuple GF2519_32.fe2519_32 count_out) count_in)
(cont : Prop -> Prop),
@@ -792,18 +792,14 @@ Qed.
Definition inm_op_correct_and_bounded1 count_in irop op
:= Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.
-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).
+ inm_op_correct_and_bounded count_in 1 irop op.*)
+Notation inm_op_correct_and_bounded n m irop op
+ := ((forall x,
+ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe2519_32WToZ x)
+ -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe2519_32WToZ x))))
+ (only parsing).
+Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
+Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe2519_32WToZ x) = true
@@ -818,20 +814,6 @@ Notation iunop_WireToFE_correct_and_bounded irop op
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).
-Notation i9top_correct_and_bounded k irop op
- := ((forall x0 x1 x2 x3 x4 x5 x6 x7 x8,
- is_bounded (fe2519_32WToZ x0) = true
- -> is_bounded (fe2519_32WToZ x1) = true
- -> is_bounded (fe2519_32WToZ x2) = true
- -> is_bounded (fe2519_32WToZ x3) = true
- -> is_bounded (fe2519_32WToZ x4) = true
- -> is_bounded (fe2519_32WToZ x5) = true
- -> is_bounded (fe2519_32WToZ x6) = true
- -> is_bounded (fe2519_32WToZ x7) = true
- -> is_bounded (fe2519_32WToZ x8) = true
- -> (Tuple.map (n:=k) fe2519_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)
- = op (fe2519_32WToZ x0) (fe2519_32WToZ x1) (fe2519_32WToZ x2) (fe2519_32WToZ x3) (fe2519_32WToZ x4) (fe2519_32WToZ x5) (fe2519_32WToZ x6) (fe2519_32WToZ x7) (fe2519_32WToZ x8))
- * HList.hlist (fun v => is_bounded v = true) (Tuple.map (n:=k) fe2519_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)))%type)
- (only parsing).
+Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-Definition prefreeze := GF2519_32.prefreeze.
+Notation prefreeze := GF2519_32.prefreeze.
diff --git a/src/SpecificGen/GF2519_32BoundedExtendedAddCoordinates.v b/src/SpecificGen/GF2519_32BoundedExtendedAddCoordinates.v
index 309a62324..eb6602760 100644
--- a/src/SpecificGen/GF2519_32BoundedExtendedAddCoordinates.v
+++ b/src/SpecificGen/GF2519_32BoundedExtendedAddCoordinates.v
@@ -5,7 +5,7 @@ Require Import Crypto.SpecificGen.GF2519_32ExtendedAddCoordinates.
Require Import Crypto.SpecificGen.GF2519_32BoundedAddCoordinates.
Require Import Crypto.Util.Tuple.
Require Import Crypto.Util.Tactics.
-
+(*
Lemma fieldwise_eq_extended_add_coordinates_full' twice_d P10 P11 P12 P13 P20 P21 P22 P23
: Tuple.fieldwise
(n:=4) GF2519_32BoundedCommon.eq
@@ -65,3 +65,4 @@ Proof.
destruct_head' prod.
rewrite <- fieldwise_eq_extended_add_coordinates_full'; reflexivity.
Qed.
+*)
diff --git a/src/SpecificGen/GF2519_32Reflective.v b/src/SpecificGen/GF2519_32Reflective.v
index 6b912c7ea..6c5759e0d 100644
--- a/src/SpecificGen/GF2519_32Reflective.v
+++ b/src/SpecificGen/GF2519_32Reflective.v
@@ -45,6 +45,7 @@ Declare Reduction asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2519_32W curry_unop_fe2519_32W curry_unop_wire_digitsW curry_9op_fe2519_32W
WordW.interp_op WordW.interp_base_type
@@ -56,6 +57,7 @@ Ltac asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe2519_32W curry_unop_fe2519_32W curry_unop_wire_digitsW curry_9op_fe2519_32W
WordW.interp_op WordW.interp_base_type
@@ -65,15 +67,15 @@ Ltac asm_interp
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
+Definition interp_radd : SpecificGen.GF2519_32BoundedCommon.fe2519_32W * SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rsub : SpecificGen.GF2519_32BoundedCommon.fe2519_32W * SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rmul : SpecificGen.GF2519_32BoundedCommon.fe2519_32W * SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
:= Eval asm_interp in interp_bexpr rmul.
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
diff --git a/src/SpecificGen/GF2519_32Reflective/Common.v b/src/SpecificGen/GF2519_32Reflective/Common.v
index db99360d7..ea27f9a6b 100644
--- a/src/SpecificGen/GF2519_32Reflective/Common.v
+++ b/src/SpecificGen/GF2519_32Reflective/Common.v
@@ -7,7 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Wf.
Require Import Crypto.Reflection.ExprInversion.
+Require Import Crypto.Reflection.Tuple.
Require Import Crypto.Reflection.Relations.
+Require Import Crypto.Reflection.Eta.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -15,13 +17,14 @@ Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Z.Syntax.Equality.
Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.Curry.
Require Import Crypto.Util.Tower.
Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Notations.
Notation Expr := (Expr base_type op).
@@ -43,30 +46,24 @@ Defined.
Definition Expr_n_OpT (count_out : nat) : flat_type base_type
:= Eval cbv [Syntax.tuple Syntax.tuple' fe2519_32T] in
Syntax.tuple fe2519_32T count_out.
-Fixpoint Expr_nm_OpT (count_in count_out : nat) : type base_type
- := match count_in with
- | 0 => Expr_n_OpT count_out
- | S n => SmartArrow fe2519_32T (Expr_nm_OpT n count_out)
- end.
+Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
+ := Eval cbv [Syntax.tuple Syntax.tuple' fe2519_32T Expr_n_OpT] in
+ Arrow (Syntax.tuple fe2519_32T count_in) (Expr_n_OpT count_out).
Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from (uncurry_unop_fe2519_32 ge_modulus). Defined.
+Proof. make_type_from ge_modulus. Defined.
Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Proof. make_type_from unpack. Defined.
Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from (uncurry_unop_fe2519_32 pack). Defined.
+Proof. make_type_from pack. Defined.
Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
Definition ExprArgT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpT.
+ := Eval compute in domain ExprUnOpT.
Definition ExprArgWireT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpFEToWireT.
-Definition ExprArgRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpT.
-Definition ExprArgWireRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Tbase TZ).
+ := Eval compute in domain ExprUnOpWireToFET.
+Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
@@ -75,99 +72,67 @@ Definition ExprBinOp : Type := Expr ExprBinOpT.
Definition ExprUnOp : Type := Expr ExprUnOpT.
Definition Expr4Op : Type := Expr Expr4OpT.
Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr ExprArgT.
-Definition ExprArgWire : Type := Expr ExprArgWireT.
-Definition ExprArgRev : Type := Expr ExprArgRevT.
-Definition ExprArgWireRev : Type := Expr ExprArgWireRevT.
+Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
+Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
Definition expr_nm_Op count_in count_out var : Type
:= expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Tbase TZ).
+Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) ExprArgT.
-Definition exprArgWire var : Type := expr base_type op (var:=var) ExprArgWireT.
-Definition exprArgRev var : Type := expr base_type op (var:=var) ExprArgRevT.
-Definition exprArgWireRev var : Type := expr base_type op (var:=var) ExprArgWireRevT.
-
-Local Ltac bounds_from_list_cps ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(fun T (extra : T) => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, extra))
- | (?x :: ?xs)%list => let bs := bounds_from_list_cps xs in
- constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
- end.
-
-Local Ltac make_bounds_cps ls extra :=
- let v := bounds_from_list_cps (List.rev ls) in
- let v := (eval compute in v) in
- exact (v _ extra).
-
-Local Ltac bounds_from_list ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
- | (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| Bounds.lower := fst x ; Bounds.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.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out : interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type.
-Proof.
- refine match count_in return interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type with
- | 0 => tt
- | S n => let v := interp_all_binders_for_to' (Expr_nm_Op_bounds n count_out) in
- interp_all_binders_for_of' _
- end; simpl.
- make_bounds_cps (Tuple.to_list _ bounds) v.
-Defined.
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
+Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
+
+Definition make_bound (x : Z * Z) : ZBounds.t
+ := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
+
+Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
+ := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
+ | 0 => tt
+ | S n
+ => let b := (Tuple.map make_bound bounds) in
+ let bs := Expr_nm_Op_bounds n count_out in
+ match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
+ | 0 => fun _ => b
+ | S n' => fun bs => (bs, b)
+ end bs
+ end.
+Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
:= Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
:= Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
:= Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
+Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
+ := Tuple.map make_bound wire_digit_bounds.
-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 (@WordW.interp_op) e).
+Definition interp_bexpr : ExprBinOp -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W * SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- := fun e => curry_unop_fe2519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.word64
- := fun e => curry_unop_fe2519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W -> SpecificGen.GF2519_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe2519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF2519_32BoundedCommon.wire_digitsW -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_9_4expr : Expr9_4Op
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
- -> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
+ -> Tuple.tuple SpecificGen.GF2519_32BoundedCommon.fe2519_32W 9
-> Tuple.tuple SpecificGen.GF2519_32BoundedCommon.fe2519_32W 4
- := fun e => curry_9op_fe2519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+ := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 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
@@ -181,40 +146,39 @@ Notation op9_4_correct_and_bounded rop op
Ltac rexpr_cbv :=
lazymatch goal with
- | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
=> let operf := head oper in
let uncurryf := head uncurry in
try cbv delta [T]; try cbv delta [oper];
try cbv beta iota delta [uncurryf]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
+ => let operf := head oper in
+ try cbv delta [T]; try cbv delta [oper]
end;
- cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+ cbv beta iota delta [interp_flat_type interp_base_type zero_ GF2519_32.fe2519_32 GF2519_32.wire_digits].
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 }
+ | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
(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 rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT (uncurry_9op_fe2519_32 op)) (only parsing).
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
+Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT 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 := ropW' in
- let ropZ := ropW' in
- let ropBounds := ropW' in
- let ropBoundedWordW := ropW' in
- ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
+ ropW' = proj1_sig ropZ_sig
+ /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
(only parsing).
Ltac app_tuples x y :=
@@ -227,7 +191,7 @@ Ltac app_tuples x y :=
Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
-
+(*
Fixpoint args_to_bounded_helperT {n}
(v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
@@ -299,14 +263,15 @@ Proof.
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
@@ -327,6 +292,7 @@ Ltac assoc_right_tuple x so_far :=
end
end.
+(*
Local Ltac make_args x :=
let x' := fresh "x'" in
compute in x |- *;
@@ -338,11 +304,6 @@ Local Ltac make_args x :=
let xv := assoc_right_tuple x'' (@None) in
refine (SmartVarf (xv : interp_flat_type _ t')).
-Definition unop_make_args {var} (x : exprArg var) : exprArgRev var.
-Proof. make_args x. Defined.
-Definition unop_wire_make_args {var} (x : exprArgWire var) : exprArgWireRev var.
-Proof. make_args x. Defined.
-
Local Ltac args_to_bounded x H :=
let x' := fresh in
set (x' := x);
@@ -351,29 +312,138 @@ Local Ltac args_to_bounded x H :=
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
+ let c := (eval cbv [domain 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) ]
).
+ *)
+
+Section gen.
+ Definition bounds_are_good_gen
+ {n : nat} (bounds : Tuple.tuple (Z * Z) n)
+ := let res :=
+ Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
+ in
+ List.fold_right andb true (Tuple.to_list n res).
+ Definition unop_args_to_bounded'
+ (bs : Z * Z)
+ (Hbs : bounds_are_good_gen (n:=1) bs = true)
+ (x : word64)
+ (H : is_bounded_gen (Tuple.map (fun v : word64 => (v : Z)) (n:=1) x) bs = true)
+ : BoundedWordW.BoundedWord.
+ Proof.
+ refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
+ unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
+ abstract (
+ destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
+ repeat apply conj; assumption
+ ).
+ Defined.
+ Fixpoint n_op_args_to_bounded'
+ n
+ : forall (bs : Tuple.tuple' (Z * Z) n)
+ (Hbs : bounds_are_good_gen (n:=S n) bs = true)
+ (x : Tuple.tuple' word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
+ Proof.
+ destruct n as [|n']; simpl in *.
+ { exact unop_args_to_bounded'. }
+ { refine (fun bs Hbs x H
+ => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
+ @unop_args_to_bounded' (snd bs) _ (snd x) _));
+ clear n_op_args_to_bounded';
+ simpl in *;
+ [ clear x H | clear Hbs | clear x H | clear Hbs ];
+ unfold bounds_are_good_gen, is_bounded_gen in *;
+ abstract (
+ repeat first [ progress simpl in *
+ | assumption
+ | reflexivity
+ | progress Bool.split_andb
+ | progress destruct_head prod
+ | match goal with
+ | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
+ end
+ | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
+ | progress break_match_hyps
+ | rewrite Bool.andb_true_iff; apply conj
+ | unfold Tuple.map, Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
+ ). }
+ Defined.
+
+ Definition n_op_args_to_bounded
+ n
+ : forall (bs : Tuple.tuple (Z * Z) n)
+ (Hbs : bounds_are_good_gen bs = true)
+ (x : Tuple.tuple word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
+ := match n with
+ | 0 => fun _ _ _ _ => tt
+ | S n' => @n_op_args_to_bounded' n'
+ end.
-Definition unop_args_to_bounded (x : fe2519_32W) (H : is_bounded (fe2519_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
+ Fixpoint nm_op_args_to_bounded' n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple' (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => @n_op_args_to_bounded m bs Hbs
+ | S n' => fun x H
+ => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
+ @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
+ end.
+ Definition nm_op_args_to_bounded n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => fun _ _ => tt
+ | S n' => @nm_op_args_to_bounded' n' m bs Hbs
+ end.
+End gen.
+
+Local Ltac get_inner_len T :=
+ lazymatch T with
+ | (?T * _)%type => get_inner_len T
+ | ?T => get_len T
+ end.
+Local Ltac get_outer_len T :=
+ lazymatch T with
+ | (?A * ?B)%type => let a := get_outer_len A in
+ let b := get_outer_len B in
+ (eval compute in (a + b)%nat)
+ | ?T => constr:(1%nat)
+ end.
+Local Ltac args_to_bounded x H :=
+ let t := type of x in
+ let m := get_inner_len t in
+ let n := get_outer_len t in
+ let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
+ let H := (eval cbv beta in (H eq_refl)) in
+ exact H.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe2519_32W * fe2519_32W)
(H : is_bounded (fe2519_32WToZ (fst x)) = true)
(H' : is_bounded (fe2519_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
-Proof.
- let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
- exact v.
-Defined.
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
+Proof. args_to_bounded x (conj H H'). Defined.
+Definition unop_args_to_bounded (x : fe2519_32W) (H : is_bounded (fe2519_32WToZ x) = true)
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain 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 _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
+Proof. args_to_bounded x H. Defined.
Definition op9_args_to_bounded (x : fe2519_32W * fe2519_32W * fe2519_32W * fe2519_32W * fe2519_32W * fe2519_32W * fe2519_32W * fe2519_32W * fe2519_32W)
(H0 : is_bounded (fe2519_32WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
(H1 : is_bounded (fe2519_32WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
@@ -384,58 +454,39 @@ Definition op9_args_to_bounded (x : fe2519_32W * fe2519_32W * fe2519_32W * fe251
(H6 : is_bounded (fe2519_32WToZ (snd (fst (fst x)))) = true)
(H7 : is_bounded (fe2519_32WToZ (snd (fst x))) = true)
(H8 : is_bounded (fe2519_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for Expr9_4OpT).
-Proof.
- let v := constr:(unop_args_to_bounded _ H8) in
- let v := app_tuples (unop_args_to_bounded _ H7) v in
- let v := app_tuples (unop_args_to_bounded _ H6) v in
- let v := app_tuples (unop_args_to_bounded _ H5) v in
- let v := app_tuples (unop_args_to_bounded _ H4) v in
- let v := app_tuples (unop_args_to_bounded _ H3) v in
- let v := app_tuples (unop_args_to_bounded _ H2) v in
- let v := app_tuples (unop_args_to_bounded _ H1) v in
- let v := app_tuples (unop_args_to_bounded _ H0) v in
- exact v.
-Defined.
-
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
+Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
+Local Ltac make_bounds_prop' bounds bounds' :=
+ first [ refine (andb _ _);
+ [ destruct bounds' as [bounds' _], bounds as [bounds _]
+ | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
+ try make_bounds_prop' bounds bounds'
+ | 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) ].
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.
+ make_bounds_prop' bounds bounds'.
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
Proof.
refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders Expr9_4OpT)) : bool.
+Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-
-Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
+(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe2519_32 f k)).
@@ -460,12 +511,11 @@ Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ v
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe2519_32 f k)).
-
+*)
(* 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
@@ -473,23 +523,25 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
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 interp_9_4expr
+ interp_flat_type_eta interp_flat_type_eta_gen
curry_binop_fe2519_32W curry_unop_fe2519_32W curry_unop_wire_digitsW curry_9op_fe2519_32W
curry_binop_fe2519_32 curry_unop_fe2519_32 curry_unop_wire_digits curry_9op_fe2519_32
uncurry_binop_fe2519_32W uncurry_unop_fe2519_32W uncurry_unop_wire_digitsW uncurry_9op_fe2519_32W
uncurry_binop_fe2519_32 uncurry_unop_fe2519_32 uncurry_unop_wire_digits uncurry_9op_fe2519_32
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT
- interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] 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 *;
+ change word64 with WordW.wordW in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
+ | progress unfold RelationClasses.Reflexive
| reflexivity
| assumption
| progress destruct_head' base_type
@@ -509,23 +561,33 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
+ Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
+ simpl @interp_flat_type in *;
+ (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
+ change (WordW.interp_base_type TZ) with v in *);
+ cbv beta iota zeta in *;
lazymatch goal with
| [ |- fe2519_32WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
| [ |- wire_digitsWToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
+ | [ |- (Tuple.map fe2519_32WToZ ?x = _) /\ _ ]
+ => generalize dependent x; intros
| [ |- ((Tuple.map fe2519_32WToZ ?x = _) * _)%type ]
=> generalize dependent x; intros
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type remove_all_binders] in *;
+ cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
destruct_head' prod;
change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
+ make_bound
+ Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
destruct_head' ZBounds.bounds;
@@ -534,12 +596,12 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
destruct_head' and;
Z.ltb_to_lt;
change WordW.wordWToZ with word64ToZ in *;
- cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd];
+ cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe2519_32WToZ HList.hlistP HList.hlistP'];
+ cbv [WordW.bit_width BitSize64.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
repeat split; unfold_is_bounded;
Z.ltb_to_lt;
try omega; try reflexivity.
-
Ltac rexpr_correct :=
let ropW' := fresh in
let ropZ_sig := fresh in
@@ -555,9 +617,13 @@ Ltac rexpr_correct :=
Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) opW) bounds)
+ := (Interp (@ZBounds.interp_op) opW bounds)
(only parsing).
+Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
+
+Ltac prove_rexpr_wfT
+ := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
Module Export PrettyPrinting.
(* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
@@ -569,23 +635,21 @@ Module Export PrettyPrinting.
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 {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end
+ := fun (t : base_type) (x : ZBounds.interp_base_type t)
+ => match x with
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
+ => in_range l u
+ | None
+ => overflow
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
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
+ := match t return interp_flat_type _ t -> result with
+ | Tbase _ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
| Unit => fun _ => no
| Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
| no, no => no
diff --git a/src/SpecificGen/GF2519_32Reflective/Common9_4Op.v b/src/SpecificGen/GF2519_32Reflective/Common9_4Op.v
index 990ad0e53..90befd75b 100644
--- a/src/SpecificGen/GF2519_32Reflective/Common9_4Op.v
+++ b/src/SpecificGen/GF2519_32Reflective/Common9_4Op.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -42,8 +41,8 @@ Lemma Expr9_4Op_correct_and_bounded
let (Hx7, Hx8) := (eta_and Hx78) in
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -80,29 +79,24 @@ Lemma Expr9_4Op_correct_and_bounded
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => op9_4_bounds_good bounds = true
| None => False
end)
: op9_4_correct_and_bounded ropW op.
Proof.
- intros x0 x1 x2 x3 x4 x5 x6 x7 x8.
- intros Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8.
- pose x0 as x0'.
- pose x1 as x1'.
- pose x2 as x2'.
- pose x3 as x3'.
- pose x4 as x4'.
- pose x5 as x5'.
- pose x6 as x6'.
- pose x7 as x7'.
- pose x8 as x8'.
- hnf in x0, x1, x2, x3, x4, x5, x6, x7, x8; destruct_head' prod.
- specialize (H0 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ intros xs Hxs.
+ pose xs as xs'.
+ compute in xs.
+ destruct_head' prod.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
+ pose Hxs as Hxs'.
+ destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
+ specialize (H0 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ specialize (H1 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded (x0', x1', x2', x3', x4', x5', x6', x7', x8') Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
+ Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.64s) (successful) *)*)
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v b/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
index c0c390d0b..bf8795d45 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonBinOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -18,8 +17,8 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -33,18 +32,19 @@ Lemma ExprBinOp_correct_and_bounded
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
end)
: binop_correct_and_bounded 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
+ intros xy HxHy.
+ pose xy as xy'.
+ compute in xy; destruct_head' prod.
+ specialize (H0 xy' HxHy).
+ specialize (H1 xy' HxHy).
+ destruct HxHy as [Hx Hy].
+ let args := constr:(binop_args_to_bounded xy' 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
index b2117295a..418dfd5e5 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOp_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
index 1b20ac0f5..6668a3b64 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToWire.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
index dc49c8ec8..bdf2a49c7 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpFEToZ.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe2519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
index b500e4cb0..c2f260694 100644
--- a/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF2519_32Reflective/CommonUnOpWireToFE.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF2519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
let args := unopWireToFE_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/AddCoordinates.v b/src/SpecificGen/GF2519_32Reflective/Reified/AddCoordinates.v
index 6c0ee65e7..4411362c0 100644
--- a/src/SpecificGen/GF2519_32Reflective/Reified/AddCoordinates.v
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/AddCoordinates.v
@@ -7,7 +7,6 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.ExprInversion.
Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
@@ -17,7 +16,10 @@ Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
+Require Import Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.Mul.
@@ -33,24 +35,28 @@ Require Import Bedrock.Word.
Definition radd_coordinatesZ' var twice_d P1 P2
:= @Extended.add_coordinates_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
twice_d _
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
P1 P2.
+Local Notation eta x := (fst x, snd x).
+
Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => radd_coordinatesZ'
- var (Return twice_d)
- (Return P10, Return P11, Return P12, Return P13)
- (Return P20, Return P21, Return P22, Return P23))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
+ => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
+ := twice_d_P1_P2 in
+ radd_coordinatesZ'
+ var (SmartVarf twice_d)
+ (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
+ (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
Definition add_coordinates
:= fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
@@ -59,70 +65,60 @@ Definition add_coordinates
twice_d (P10, P11, P12, P13) (P20, P21, P22, P23).
Definition uncurried_add_coordinates
- := apply9_nd
- (@uncurry_unop_fe2519_32)
- add_coordinates.
+ := fun twice_d_P1_P2
+ => let twice_d := fst twice_d_P1_P2 in
+ let (P1, P2) := eta (snd twice_d_P1_P2) in
+ @Extended.add_coordinates
+ _ add sub mul
+ twice_d P1 P2.
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
let H := fresh in
pose k as k';
assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
+ repeat (unfold k'; change k with k'; unfold interpf_step);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF2519_32.fe2519_32 -> GF2519_32.fe2519_32 -> GF2519_32.fe2519_32)
- (x' y' : GF2519_32.fe2519_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe2519_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
+Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe2519_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF2519_32.fe2519_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe2519_32 Apply' interp fst snd].
- rewrite (invert_Abs_Some (e:=f) eq_refl).
+ cbv [radd_coordinatesZ''].
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type] in x.
+ destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
+ | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
end.
- reflexivity.
-Qed.
-
-Lemma radd_coordinatesZ_sigP : rexpr_sigP Expr9_4OpT uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT radd_coordinatesZ'' uncurried_add_coordinates uncurry_unop_fe2519_32 SmartAbs radd_coordinatesZ' exprArg Extended.add_coordinates_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe2519_32 add_coordinates].
- intros.
- unfold invert_Return at 13 14 15 16.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [radd_coordinatesZ' add_coordinates Extended.add_coordinates_gen uncurried_add_coordinates SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -132,48 +128,56 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 13 14 15; unfold interpf_step.
- cbv beta iota delta [Extended.add_coordinates].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition radd_coordinatesZ_sig : rexpr_9_4op_sig add_coordinates.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv beta iota delta [Extended.add_coordinates interp_flat_type interp_base_type GF2519_32.fe2519_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- eexists.
- apply radd_coordinatesZ_sigP.
-Defined.
+ exact radd_coordinatesZ_sigP'.
+Qed.
+Definition radd_coordinatesZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
+
+Definition radd_coordinates_input_bounds
+ := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW Expr9Op_bounds.
+Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition radd_coordinatesW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- radd_coordinatesW add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
+ radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW Expr9Op_bounds).
+Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/LadderStep.v b/src/SpecificGen/GF2519_32Reflective/Reified/LadderStep.v
index b29f1822b..1a4204457 100644
--- a/src/SpecificGen/GF2519_32Reflective/Reified/LadderStep.v
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/LadderStep.v
@@ -7,8 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Relations.
Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -18,6 +19,7 @@ Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
Require Import Crypto.Spec.MxDH.
+Require Import Crypto.SpecificGen.GF2519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF2519_32Reflective.Reified.Mul.
@@ -30,50 +32,80 @@ Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
Require Import Bedrock.Word.
-(* XXX FIXME: Remove dummy code *)
-Definition rladderstepZ' var (T:=_) (dummy0 dummy1 dummy2 a24 x0 : T) P1 P2
+Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
:= @MxDH.ladderstep_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
a24
_
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
x0
P1 P2.
Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21
- => rladderstepZ'
- var (Return dummy0) (Return dummy1) (Return dummy2)
- (Return a24) (Return x0)
- (Return P10, Return P11)
- (Return P20, Return P21))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
+ => let '(a24, x0, ((P10, P11), (P20, P21)))
+ := a24_x0_P1_P2 in
+ rladderstepZ'
+ var (SmartVarf a24) (SmartVarf x0)
+ (SmartVarf P10, SmartVarf P11)
+ (SmartVarf P20, SmartVarf P21)))).
-Definition ladderstep (T := _)
- := fun (dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21 : T)
- => @MxDH.ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
+Local Notation eta x := (fst x, snd x).
+
+Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
+ Eval cbv beta delta [MxDH.ladderstep_gen] in
+ @MxDH.ladderstep_gen
+ F Fadd Fsub Fmul a24
+ (F*F*F*F)
+ (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
+ (fun x f => dlet y := x in f y)
+ X1 P1 P2.
Definition uncurried_ladderstep
- := apply9_nd
- (@uncurry_unop_fe2519_32)
- ladderstep.
+ := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
+ => let a24 := fst (fst a24_x0_P1_P2) in
+ let x0 := snd (fst a24_x0_P1_P2) in
+ let '(P1, P2) := eta (snd a24_x0_P1_P2) in
+ let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
+ @ladderstep_other_assoc
+ _ add sub mul
+ a24 x0 (P10, P11) (P20, P21).
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
@@ -83,51 +115,14 @@ Local Ltac repeat_step_interpf :=
repeat (unfold interpf_step at 1; change k with k' at 1);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF2519_32.fe2519_32 -> GF2519_32.fe2519_32 -> GF2519_32.fe2519_32)
- (x' y' : GF2519_32.fe2519_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe2519_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
-Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe2519_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF2519_32.fe2519_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe2519_32 Apply' interp fst snd].
- let f := match goal with f : expr _ _ _ |- _ => f end in
- rewrite (invert_Abs_Some (e:=f) eq_refl).
- repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
- end.
- reflexivity.
-Qed.
-
-Lemma rladderstepZ_sigP : rexpr_sigP Expr9_4OpT uncurried_ladderstep rladderstepZ''.
+Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
cbv [rladderstepZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT rladderstepZ'' uncurried_ladderstep uncurry_unop_fe2519_32 SmartAbs rladderstepZ' exprArg MxDH.ladderstep_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe2519_32 ladderstep].
- intros; cbv beta zeta.
- unfold invert_Return at 14 15 16 17.
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type interp_base_type] in x.
+ destruct_head' prod.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -137,48 +132,58 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 14 15 16; unfold interpf_step.
- cbv beta iota delta [MxDH.ladderstep_other_assoc].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition rladderstepZ_sig : rexpr_9_4op_sig ladderstep.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv [ladderstep_other_assoc interp_flat_type GF2519_32.fe2519_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In Common.ExprBinOpT] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | _ => progress cbv [interp_base_type]
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
- eexists.
- apply rladderstepZ_sigP.
-Defined.
+ exact rladderstepZ_sigP'.
+Qed.
+Definition rladderstepZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
+
+Definition rladderstep_input_bounds
+ := (ExprUnOp_bounds, ExprUnOp_bounds,
+ ((ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW Expr9Op_bounds.
+Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition rladderstepW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- rladderstepW ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
+ rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW Expr9Op_bounds).
+Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v
index dc9e2dcef..fd509d77c 100644
--- a/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v
+++ b/src/SpecificGen/GF2519_32Reflective/Reified/PreFreeze.v
@@ -1,6 +1,6 @@
Require Import Crypto.SpecificGen.GF2519_32Reflective.CommonUnOp.
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF2519_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. 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.
diff --git a/src/SpecificGen/GF2519_32ReflectiveAddCoordinates.v b/src/SpecificGen/GF2519_32ReflectiveAddCoordinates.v
index 3ceccc8dd..4010eac57 100644
--- a/src/SpecificGen/GF2519_32ReflectiveAddCoordinates.v
+++ b/src/SpecificGen/GF2519_32ReflectiveAddCoordinates.v
@@ -32,7 +32,7 @@ Import ListNotations.
Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Local Open Scope Z.
-Time Definition radd_coordinates : Expr9_4Op := Eval vm_compute in radd_coordinatesW.
+Time Definition radd_coordinates : Expr _ := Eval vm_compute in radd_coordinatesW.
Declare Reduction asm_interp_add_coordinates
:= cbv beta iota delta
@@ -57,7 +57,7 @@ Ltac asm_interp_add_coordinates
mapf_interp_flat_type WordW.interp_base_type word64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
-
+(*
Time Definition interp_radd_coordinates : SpecificGen.GF2519_32BoundedCommon.fe2519_32W
-> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
-> SpecificGen.GF2519_32BoundedCommon.fe2519_32W
@@ -74,3 +74,4 @@ Time Definition interp_radd_coordinates_correct : interp_radd_coordinates = inte
Lemma radd_coordinates_correct_and_bounded : op9_4_correct_and_bounded radd_coordinates add_coordinates.
Proof. exact_no_check radd_coordinatesW_correct_and_bounded. Time Qed.
+*)
diff --git a/src/SpecificGen/GF25519_32Bounded.v b/src/SpecificGen/GF25519_32Bounded.v
index 03a1bb125..025950bee 100644
--- a/src/SpecificGen/GF25519_32Bounded.v
+++ b/src/SpecificGen/GF25519_32Bounded.v
@@ -25,13 +25,23 @@ Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.
Local Open Scope Z.
+Local Ltac cbv_tuple_map :=
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'].
+
+Local Ltac post_bounded_t :=
+ (* much pain and hackery to work around [Defined] taking forever *)
+ cbv_tuple_map;
+ let blem' := fresh "blem'" in
+ let is_bounded_lem := fresh "is_bounded_lem" in
+ intros is_bounded_lem blem';
+ apply blem'; repeat apply conj; apply is_bounded_lem.
Local Ltac bounded_t opW blem :=
- apply blem; apply is_bounded_proj1_fe25519_32.
+ generalize blem; generalize is_bounded_proj1_fe25519_32; post_bounded_t.
Local Ltac bounded_wire_digits_t opW blem :=
- apply blem; apply is_bounded_proj1_wire_digits.
+ generalize blem; generalize is_bounded_proj1_wire_digits; post_bounded_t.
Local Ltac define_binop f g opW blem :=
- refine (exist_fe25519_32W (opW (proj1_fe25519_32W f) (proj1_fe25519_32W g)) _);
+ 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)) _);
@@ -47,17 +57,17 @@ Local Ltac define_unop_WireToFE f 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_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 addW (f : fe25519_32W * fe25519_32W) : fe25519_32W := Eval simpl in interp_radd f.
+Definition subW (f : fe25519_32W * fe25519_32W) : fe25519_32W := Eval simpl in interp_rsub f.
+Definition mulW (f : fe25519_32W * fe25519_32W) : fe25519_32W := Eval simpl in interp_rmul f.
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.
@@ -86,7 +96,7 @@ Definition freezeW (f : fe25519_32W) : fe25519_32W := Eval cbv beta delta [prefr
Local Transparent Let_In.
(* Wrapper to allow extracted code to not unfold [mulW] *)
Definition mulW_noinline := mulW.
-Definition powW (f : fe25519_32W) chain := fold_chain_opt (proj1_fe25519_32W one) mulW_noinline chain [f].
+Definition powW (f : fe25519_32W) chain := fold_chain_opt (proj1_fe25519_32W one) (fun f g => mulW_noinline (f, g)) chain [f].
Definition invW (f : fe25519_32W) : fe25519_32W
:= Eval cbv -[Let_In fe25519_32W mulW_noinline] in powW f (chain inv_ec).
@@ -95,11 +105,11 @@ Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
rewrite pre_rewrite;
intros; apply rop_cb; assumption.
-Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW (Curry.curry2 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.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW (Curry.curry2 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.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW (Curry.curry2 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.
@@ -142,6 +152,7 @@ Proof.
cbv [modulusW Tuple.map].
cbv [on_tuple List.map to_list to_list' from_list from_list'
+ HList.hlistP HList.hlistP'
Tuple.map2 on_tuple2 ListUtil.map2 fe25519_32WToZ length_fe25519_32].
cbv [postfreeze GF25519_32.postfreeze].
cbv [Let_In].
@@ -203,7 +214,8 @@ Proof.
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.
+ split; [ | assumption ].
+ rewrite H0', H0; reflexivity.
Qed.
Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
@@ -212,9 +224,11 @@ Proof.
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; pose proof (fun k0 k1 X Y => proj1 (mulW_correct_and_bounded (k0, k1) (conj X Y))) as H'.
+ cbv [Curry.curry2 Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list'] in H'.
+ rewrite <- H' by assumption.
+ apply mulW_correct_and_bounded; split; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded (_, _) (conj X Y))) by assumption; reflexivity. }
{ intros [|?]; autorewrite with simpl_nth_default;
(assumption || reflexivity). }
Qed.
@@ -269,8 +283,10 @@ Proof.
unfold GF25519_32.eqb.
simpl @fe25519_32WToZ in *; cbv beta iota.
intros.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe25519_32WToZ] in *.
rewrite <- frf, <- frg by assumption.
- rewrite <- fieldwisebW_correct.
+ etransitivity; [ eapply fieldwisebW_correct | ].
+ cbv [fe25519_32WToZ].
reflexivity.
Defined.
@@ -297,7 +313,7 @@ Proof.
lazymatch (eval cbv delta [GF25519_32.sqrt] in GF25519_32.sqrt) with
| (fun powf powf_squared f => dlet a := powf in _)
=> exact (dlet powx := powW (fe25519_32ZToW x) (chain GF25519_32.sqrt_ec) in
- GF25519_32.sqrt (fe25519_32WToZ powx) (fe25519_32WToZ (mulW_noinline powx powx)) x)
+ GF25519_32.sqrt (fe25519_32WToZ powx) (fe25519_32WToZ (mulW_noinline (powx, powx))) x)
| (fun f => pow f _)
=> exact (GF25519_32.sqrt x)
end.
@@ -324,21 +340,41 @@ Proof.
=> 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
- | mulW_noinline _ _ => apply mulW_correct_and_bounded
- | powW _ _ => apply powW_correct_and_bounded
- | sqrt_m1W => vm_compute; reflexivity
- | _ => assumption
- end
- end;
- subst_evars; reflexivity
+ pose proof (fun X Y => proj1 (mulW_correct_and_bounded (sqrt_m1W, a) (conj X Y))) as correctness;
+ let cbv_in_all _ := (cbv [fe25519_32WToZ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe25519_32WToZ Curry.curry2 HList.hlistP HList.hlistP'] in *; idtac) in
+ cbv_in_all ();
+ let solver _ := (repeat match goal with
+ | _ => progress subst
+ | _ => progress unfold fst, snd
+ | _ => progress cbv_in_all ()
+ | [ |- ?x /\ ?x ] => cut x; [ intro; split; assumption | ]
+ | [ |- is_bounded ?op = true ]
+ => let H := fresh in
+ lazymatch op with
+ | context[mulW (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[mulW_noinline (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[powW _ _] => pose proof powW_correct_and_bounded as H
+ | context[sqrt_m1W] => vm_compute; reflexivity
+ | _ => assumption
+ end;
+ cbv_in_all ();
+ apply H
+ end) in
+ rewrite <- correctness by solver (); clear correctness;
+ let lem := fresh in
+ pose proof eqbW_correct as lem; cbv_in_all (); rewrite <- lem by solver (); clear lem;
+ pose proof (pull_bool_if fe25519_32WToZ) as lem; cbv_in_all (); rewrite lem by solver (); clear lem;
+ subst_evars; reflexivity
end.
} Unfocus.
+ assert (Hfold : forall x, fe25519_32WToZ x = fe25519_32WToZ x) by reflexivity.
+ unfold fe25519_32WToZ at 2 in Hfold.
+ etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody; [ reflexivity | intro ].
+ apply Hfold.
+ } Unfocus.
+ clear Hfold.
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
@@ -346,15 +382,22 @@ Proof.
[ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
| _ => idtac
end;
- reflexivity. }
- { cbv [Let_In];
+ reflexivity.
+ }
+
+ { cbv [Let_In HList.hlistP HList.hlistP'];
try break_if;
repeat lazymatch goal with
| [ |- is_bounded (?WToZ (powW _ _)) = true ]
=> apply powW_correct_and_bounded; assumption
- | [ |- is_bounded (?WToZ (mulW _ _)) = true ]
- => apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ]
- end. }
+ | [ |- is_bounded (snd (?WToZ (_, powW _ _))) = true ]
+ => generalize powW_correct_and_bounded;
+ cbv [snd Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'];
+ let H := fresh in intro H; apply H; assumption
+ | [ |- is_bounded (?WToZ (mulW (_, _))) = true ]
+ => apply mulW_correct_and_bounded; split; [ vm_compute; reflexivity | ]
+ end.
+ }
Defined.
Definition sqrtW (f : fe25519_32W) : fe25519_32W :=
@@ -394,7 +437,7 @@ 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.
+Proof. define_unop f invW (fun x p => proj2 (invW_correct_and_bounded x p)). Defined.
Definition sqrt (f : fe25519_32) : fe25519_32.
Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
@@ -407,7 +450,12 @@ Local Ltac op_correct_t op opW_correct_and_bounded :=
=> rewrite proj1_wire_digits_exist_wire_digitsW
| _ => idtac
end;
- apply opW_correct_and_bounded;
+ generalize opW_correct_and_bounded;
+ cbv_tuple_map;
+ cbv [fst snd];
+ let H := fresh in
+ intro H; apply H;
+ repeat match goal with |- and _ _ => apply conj end;
lazymatch goal with
| [ |- is_bounded _ = true ]
=> apply is_bounded_proj1_fe25519_32
@@ -434,7 +482,7 @@ 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.
+Proof. op_correct_t inv (fun x p => proj1 (invW_correct_and_bounded x p)). 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.
diff --git a/src/SpecificGen/GF25519_32BoundedAddCoordinates.v b/src/SpecificGen/GF25519_32BoundedAddCoordinates.v
index 5a1fd12b3..6e5bdae4f 100644
--- a/src/SpecificGen/GF25519_32BoundedAddCoordinates.v
+++ b/src/SpecificGen/GF25519_32BoundedAddCoordinates.v
@@ -14,7 +14,7 @@ Local Ltac define_binop f g 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_coordinates / _ _ _ _ _ _ _ _ _.
+(*Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
Definition add_coordinatesW (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe25519_32W) : Tuple.tuple fe25519_32W 4
:= Eval simpl in interp_radd_coordinates x0 x1 x2 x3 x4 x5 x6 x7 x8.
@@ -75,3 +75,4 @@ Lemma add_coordinates_correct (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe25519_32)
(proj1_fe25519_32 x7)
(proj1_fe25519_32 x8).
Proof. op_correct_t add_coordinates add_coordinatesW_correct_and_bounded. Qed.
+*)
diff --git a/src/SpecificGen/GF25519_32BoundedCommon.v b/src/SpecificGen/GF25519_32BoundedCommon.v
index bf76e028b..f21a391b8 100644
--- a/src/SpecificGen/GF25519_32BoundedCommon.v
+++ b/src/SpecificGen/GF25519_32BoundedCommon.v
@@ -322,14 +322,14 @@ Ltac unfold_is_bounded_in' H :=
end.
Ltac preunfold_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.
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe25519_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
Ltac unfold_is_bounded_in H :=
preunfold_is_bounded_in H;
unfold_is_bounded_in' H.
Ltac preunfold_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].
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe25519_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
Ltac unfold_is_bounded :=
preunfold_is_bounded;
@@ -724,7 +724,7 @@ Notation in_op_correct_and_bounded k irop op
/\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe25519_32WToZ irop))%type)
(only parsing).
-Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
+(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
: forall (irop : Tower.tower_nd fe25519_32W (Tuple.tuple fe25519_32W count_out) count_in)
(op : Tower.tower_nd GF25519_32.fe25519_32 (Tuple.tuple GF25519_32.fe25519_32 count_out) count_in)
(cont : Prop -> Prop),
@@ -792,18 +792,14 @@ Qed.
Definition inm_op_correct_and_bounded1 count_in irop op
:= Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.
-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).
+ inm_op_correct_and_bounded count_in 1 irop op.*)
+Notation inm_op_correct_and_bounded n m irop op
+ := ((forall x,
+ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe25519_32WToZ x)
+ -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe25519_32WToZ x))))
+ (only parsing).
+Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
+Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe25519_32WToZ x) = true
@@ -818,20 +814,6 @@ Notation iunop_WireToFE_correct_and_bounded irop op
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).
-Notation i9top_correct_and_bounded k irop op
- := ((forall x0 x1 x2 x3 x4 x5 x6 x7 x8,
- is_bounded (fe25519_32WToZ x0) = true
- -> is_bounded (fe25519_32WToZ x1) = true
- -> is_bounded (fe25519_32WToZ x2) = true
- -> is_bounded (fe25519_32WToZ x3) = true
- -> is_bounded (fe25519_32WToZ x4) = true
- -> is_bounded (fe25519_32WToZ x5) = true
- -> is_bounded (fe25519_32WToZ x6) = true
- -> is_bounded (fe25519_32WToZ x7) = true
- -> is_bounded (fe25519_32WToZ x8) = true
- -> (Tuple.map (n:=k) fe25519_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)
- = op (fe25519_32WToZ x0) (fe25519_32WToZ x1) (fe25519_32WToZ x2) (fe25519_32WToZ x3) (fe25519_32WToZ x4) (fe25519_32WToZ x5) (fe25519_32WToZ x6) (fe25519_32WToZ x7) (fe25519_32WToZ x8))
- * HList.hlist (fun v => is_bounded v = true) (Tuple.map (n:=k) fe25519_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)))%type)
- (only parsing).
+Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-Definition prefreeze := GF25519_32.prefreeze.
+Notation prefreeze := GF25519_32.prefreeze.
diff --git a/src/SpecificGen/GF25519_32BoundedExtendedAddCoordinates.v b/src/SpecificGen/GF25519_32BoundedExtendedAddCoordinates.v
index c23bcea4a..4d28f5376 100644
--- a/src/SpecificGen/GF25519_32BoundedExtendedAddCoordinates.v
+++ b/src/SpecificGen/GF25519_32BoundedExtendedAddCoordinates.v
@@ -5,7 +5,7 @@ Require Import Crypto.SpecificGen.GF25519_32ExtendedAddCoordinates.
Require Import Crypto.SpecificGen.GF25519_32BoundedAddCoordinates.
Require Import Crypto.Util.Tuple.
Require Import Crypto.Util.Tactics.
-
+(*
Lemma fieldwise_eq_extended_add_coordinates_full' twice_d P10 P11 P12 P13 P20 P21 P22 P23
: Tuple.fieldwise
(n:=4) GF25519_32BoundedCommon.eq
@@ -65,3 +65,4 @@ Proof.
destruct_head' prod.
rewrite <- fieldwise_eq_extended_add_coordinates_full'; reflexivity.
Qed.
+*)
diff --git a/src/SpecificGen/GF25519_32Reflective.v b/src/SpecificGen/GF25519_32Reflective.v
index cfa76bf01..6e3f169ea 100644
--- a/src/SpecificGen/GF25519_32Reflective.v
+++ b/src/SpecificGen/GF25519_32Reflective.v
@@ -45,6 +45,7 @@ Declare Reduction asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_32W curry_unop_fe25519_32W curry_unop_wire_digitsW curry_9op_fe25519_32W
WordW.interp_op WordW.interp_base_type
@@ -56,6 +57,7 @@ Ltac asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_32W curry_unop_fe25519_32W curry_unop_wire_digitsW curry_9op_fe25519_32W
WordW.interp_op WordW.interp_base_type
@@ -65,15 +67,15 @@ Ltac asm_interp
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
+Definition interp_radd : SpecificGen.GF25519_32BoundedCommon.fe25519_32W * SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rsub : SpecificGen.GF25519_32BoundedCommon.fe25519_32W * SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rmul : SpecificGen.GF25519_32BoundedCommon.fe25519_32W * SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
:= Eval asm_interp in interp_bexpr rmul.
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
diff --git a/src/SpecificGen/GF25519_32Reflective/Common.v b/src/SpecificGen/GF25519_32Reflective/Common.v
index 3d3686a54..29d653cd6 100644
--- a/src/SpecificGen/GF25519_32Reflective/Common.v
+++ b/src/SpecificGen/GF25519_32Reflective/Common.v
@@ -7,7 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Wf.
Require Import Crypto.Reflection.ExprInversion.
+Require Import Crypto.Reflection.Tuple.
Require Import Crypto.Reflection.Relations.
+Require Import Crypto.Reflection.Eta.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -15,13 +17,14 @@ Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Z.Syntax.Equality.
Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.Curry.
Require Import Crypto.Util.Tower.
Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Notations.
Notation Expr := (Expr base_type op).
@@ -43,30 +46,24 @@ Defined.
Definition Expr_n_OpT (count_out : nat) : flat_type base_type
:= Eval cbv [Syntax.tuple Syntax.tuple' fe25519_32T] in
Syntax.tuple fe25519_32T count_out.
-Fixpoint Expr_nm_OpT (count_in count_out : nat) : type base_type
- := match count_in with
- | 0 => Expr_n_OpT count_out
- | S n => SmartArrow fe25519_32T (Expr_nm_OpT n count_out)
- end.
+Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
+ := Eval cbv [Syntax.tuple Syntax.tuple' fe25519_32T Expr_n_OpT] in
+ Arrow (Syntax.tuple fe25519_32T count_in) (Expr_n_OpT count_out).
Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from (uncurry_unop_fe25519_32 ge_modulus). Defined.
+Proof. make_type_from ge_modulus. Defined.
Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Proof. make_type_from unpack. Defined.
Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from (uncurry_unop_fe25519_32 pack). Defined.
+Proof. make_type_from pack. Defined.
Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
Definition ExprArgT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpT.
+ := Eval compute in domain ExprUnOpT.
Definition ExprArgWireT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpFEToWireT.
-Definition ExprArgRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpT.
-Definition ExprArgWireRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Tbase TZ).
+ := Eval compute in domain ExprUnOpWireToFET.
+Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
@@ -75,99 +72,67 @@ Definition ExprBinOp : Type := Expr ExprBinOpT.
Definition ExprUnOp : Type := Expr ExprUnOpT.
Definition Expr4Op : Type := Expr Expr4OpT.
Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr ExprArgT.
-Definition ExprArgWire : Type := Expr ExprArgWireT.
-Definition ExprArgRev : Type := Expr ExprArgRevT.
-Definition ExprArgWireRev : Type := Expr ExprArgWireRevT.
+Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
+Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
Definition expr_nm_Op count_in count_out var : Type
:= expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Tbase TZ).
+Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) ExprArgT.
-Definition exprArgWire var : Type := expr base_type op (var:=var) ExprArgWireT.
-Definition exprArgRev var : Type := expr base_type op (var:=var) ExprArgRevT.
-Definition exprArgWireRev var : Type := expr base_type op (var:=var) ExprArgWireRevT.
-
-Local Ltac bounds_from_list_cps ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(fun T (extra : T) => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, extra))
- | (?x :: ?xs)%list => let bs := bounds_from_list_cps xs in
- constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
- end.
-
-Local Ltac make_bounds_cps ls extra :=
- let v := bounds_from_list_cps (List.rev ls) in
- let v := (eval compute in v) in
- exact (v _ extra).
-
-Local Ltac bounds_from_list ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
- | (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| Bounds.lower := fst x ; Bounds.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.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out : interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type.
-Proof.
- refine match count_in return interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type with
- | 0 => tt
- | S n => let v := interp_all_binders_for_to' (Expr_nm_Op_bounds n count_out) in
- interp_all_binders_for_of' _
- end; simpl.
- make_bounds_cps (Tuple.to_list _ bounds) v.
-Defined.
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
+Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
+
+Definition make_bound (x : Z * Z) : ZBounds.t
+ := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
+
+Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
+ := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
+ | 0 => tt
+ | S n
+ => let b := (Tuple.map make_bound bounds) in
+ let bs := Expr_nm_Op_bounds n count_out in
+ match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
+ | 0 => fun _ => b
+ | S n' => fun bs => (bs, b)
+ end bs
+ end.
+Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
:= Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
:= Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
:= Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
+Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
+ := Tuple.map make_bound wire_digit_bounds.
-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 (@WordW.interp_op) e).
+Definition interp_bexpr : ExprBinOp -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W * SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- := fun e => curry_unop_fe25519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.word64
- := fun e => curry_unop_fe25519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe25519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF25519_32BoundedCommon.wire_digitsW -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_9_4expr : Expr9_4Op
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
- -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
+ -> Tuple.tuple SpecificGen.GF25519_32BoundedCommon.fe25519_32W 9
-> Tuple.tuple SpecificGen.GF25519_32BoundedCommon.fe25519_32W 4
- := fun e => curry_9op_fe25519_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+ := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 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
@@ -181,40 +146,39 @@ Notation op9_4_correct_and_bounded rop op
Ltac rexpr_cbv :=
lazymatch goal with
- | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
=> let operf := head oper in
let uncurryf := head uncurry in
try cbv delta [T]; try cbv delta [oper];
try cbv beta iota delta [uncurryf]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
+ => let operf := head oper in
+ try cbv delta [T]; try cbv delta [oper]
end;
- cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+ cbv beta iota delta [interp_flat_type interp_base_type zero_ GF25519_32.fe25519_32 GF25519_32.wire_digits].
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 }
+ | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
(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 rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT (uncurry_9op_fe25519_32 op)) (only parsing).
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
+Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT 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 := ropW' in
- let ropZ := ropW' in
- let ropBounds := ropW' in
- let ropBoundedWordW := ropW' in
- ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
+ ropW' = proj1_sig ropZ_sig
+ /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
(only parsing).
Ltac app_tuples x y :=
@@ -227,7 +191,7 @@ Ltac app_tuples x y :=
Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
-
+(*
Fixpoint args_to_bounded_helperT {n}
(v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
@@ -299,14 +263,15 @@ Proof.
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
@@ -327,6 +292,7 @@ Ltac assoc_right_tuple x so_far :=
end
end.
+(*
Local Ltac make_args x :=
let x' := fresh "x'" in
compute in x |- *;
@@ -338,11 +304,6 @@ Local Ltac make_args x :=
let xv := assoc_right_tuple x'' (@None) in
refine (SmartVarf (xv : interp_flat_type _ t')).
-Definition unop_make_args {var} (x : exprArg var) : exprArgRev var.
-Proof. make_args x. Defined.
-Definition unop_wire_make_args {var} (x : exprArgWire var) : exprArgWireRev var.
-Proof. make_args x. Defined.
-
Local Ltac args_to_bounded x H :=
let x' := fresh in
set (x' := x);
@@ -351,29 +312,138 @@ Local Ltac args_to_bounded x H :=
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
+ let c := (eval cbv [domain 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) ]
).
+ *)
+
+Section gen.
+ Definition bounds_are_good_gen
+ {n : nat} (bounds : Tuple.tuple (Z * Z) n)
+ := let res :=
+ Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
+ in
+ List.fold_right andb true (Tuple.to_list n res).
+ Definition unop_args_to_bounded'
+ (bs : Z * Z)
+ (Hbs : bounds_are_good_gen (n:=1) bs = true)
+ (x : word64)
+ (H : is_bounded_gen (Tuple.map (fun v : word64 => (v : Z)) (n:=1) x) bs = true)
+ : BoundedWordW.BoundedWord.
+ Proof.
+ refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
+ unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
+ abstract (
+ destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
+ repeat apply conj; assumption
+ ).
+ Defined.
+ Fixpoint n_op_args_to_bounded'
+ n
+ : forall (bs : Tuple.tuple' (Z * Z) n)
+ (Hbs : bounds_are_good_gen (n:=S n) bs = true)
+ (x : Tuple.tuple' word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
+ Proof.
+ destruct n as [|n']; simpl in *.
+ { exact unop_args_to_bounded'. }
+ { refine (fun bs Hbs x H
+ => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
+ @unop_args_to_bounded' (snd bs) _ (snd x) _));
+ clear n_op_args_to_bounded';
+ simpl in *;
+ [ clear x H | clear Hbs | clear x H | clear Hbs ];
+ unfold bounds_are_good_gen, is_bounded_gen in *;
+ abstract (
+ repeat first [ progress simpl in *
+ | assumption
+ | reflexivity
+ | progress Bool.split_andb
+ | progress destruct_head prod
+ | match goal with
+ | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
+ end
+ | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
+ | progress break_match_hyps
+ | rewrite Bool.andb_true_iff; apply conj
+ | unfold Tuple.map, Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
+ ). }
+ Defined.
+
+ Definition n_op_args_to_bounded
+ n
+ : forall (bs : Tuple.tuple (Z * Z) n)
+ (Hbs : bounds_are_good_gen bs = true)
+ (x : Tuple.tuple word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
+ := match n with
+ | 0 => fun _ _ _ _ => tt
+ | S n' => @n_op_args_to_bounded' n'
+ end.
-Definition unop_args_to_bounded (x : fe25519_32W) (H : is_bounded (fe25519_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
+ Fixpoint nm_op_args_to_bounded' n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple' (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => @n_op_args_to_bounded m bs Hbs
+ | S n' => fun x H
+ => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
+ @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
+ end.
+ Definition nm_op_args_to_bounded n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => fun _ _ => tt
+ | S n' => @nm_op_args_to_bounded' n' m bs Hbs
+ end.
+End gen.
+
+Local Ltac get_inner_len T :=
+ lazymatch T with
+ | (?T * _)%type => get_inner_len T
+ | ?T => get_len T
+ end.
+Local Ltac get_outer_len T :=
+ lazymatch T with
+ | (?A * ?B)%type => let a := get_outer_len A in
+ let b := get_outer_len B in
+ (eval compute in (a + b)%nat)
+ | ?T => constr:(1%nat)
+ end.
+Local Ltac args_to_bounded x H :=
+ let t := type of x in
+ let m := get_inner_len t in
+ let n := get_outer_len t in
+ let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
+ let H := (eval cbv beta in (H eq_refl)) in
+ exact H.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe25519_32W * fe25519_32W)
(H : is_bounded (fe25519_32WToZ (fst x)) = true)
(H' : is_bounded (fe25519_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
-Proof.
- let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
- exact v.
-Defined.
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
+Proof. args_to_bounded x (conj H H'). Defined.
+Definition unop_args_to_bounded (x : fe25519_32W) (H : is_bounded (fe25519_32WToZ x) = true)
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain 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 _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
+Proof. args_to_bounded x H. Defined.
Definition op9_args_to_bounded (x : fe25519_32W * fe25519_32W * fe25519_32W * fe25519_32W * fe25519_32W * fe25519_32W * fe25519_32W * fe25519_32W * fe25519_32W)
(H0 : is_bounded (fe25519_32WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
(H1 : is_bounded (fe25519_32WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
@@ -384,58 +454,39 @@ Definition op9_args_to_bounded (x : fe25519_32W * fe25519_32W * fe25519_32W * fe
(H6 : is_bounded (fe25519_32WToZ (snd (fst (fst x)))) = true)
(H7 : is_bounded (fe25519_32WToZ (snd (fst x))) = true)
(H8 : is_bounded (fe25519_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for Expr9_4OpT).
-Proof.
- let v := constr:(unop_args_to_bounded _ H8) in
- let v := app_tuples (unop_args_to_bounded _ H7) v in
- let v := app_tuples (unop_args_to_bounded _ H6) v in
- let v := app_tuples (unop_args_to_bounded _ H5) v in
- let v := app_tuples (unop_args_to_bounded _ H4) v in
- let v := app_tuples (unop_args_to_bounded _ H3) v in
- let v := app_tuples (unop_args_to_bounded _ H2) v in
- let v := app_tuples (unop_args_to_bounded _ H1) v in
- let v := app_tuples (unop_args_to_bounded _ H0) v in
- exact v.
-Defined.
-
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
+Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
+Local Ltac make_bounds_prop' bounds bounds' :=
+ first [ refine (andb _ _);
+ [ destruct bounds' as [bounds' _], bounds as [bounds _]
+ | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
+ try make_bounds_prop' bounds bounds'
+ | 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) ].
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.
+ make_bounds_prop' bounds bounds'.
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
Proof.
refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders Expr9_4OpT)) : bool.
+Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-
-Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
+(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe25519_32 f k)).
@@ -460,12 +511,11 @@ Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ v
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe25519_32 f k)).
-
+*)
(* 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
@@ -473,23 +523,25 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
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 interp_9_4expr
+ interp_flat_type_eta interp_flat_type_eta_gen
curry_binop_fe25519_32W curry_unop_fe25519_32W curry_unop_wire_digitsW curry_9op_fe25519_32W
curry_binop_fe25519_32 curry_unop_fe25519_32 curry_unop_wire_digits curry_9op_fe25519_32
uncurry_binop_fe25519_32W uncurry_unop_fe25519_32W uncurry_unop_wire_digitsW uncurry_9op_fe25519_32W
uncurry_binop_fe25519_32 uncurry_unop_fe25519_32 uncurry_unop_wire_digits uncurry_9op_fe25519_32
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT
- interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] 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 *;
+ change word64 with WordW.wordW in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
+ | progress unfold RelationClasses.Reflexive
| reflexivity
| assumption
| progress destruct_head' base_type
@@ -509,23 +561,33 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
+ Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
+ simpl @interp_flat_type in *;
+ (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
+ change (WordW.interp_base_type TZ) with v in *);
+ cbv beta iota zeta in *;
lazymatch goal with
| [ |- fe25519_32WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
| [ |- wire_digitsWToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
+ | [ |- (Tuple.map fe25519_32WToZ ?x = _) /\ _ ]
+ => generalize dependent x; intros
| [ |- ((Tuple.map fe25519_32WToZ ?x = _) * _)%type ]
=> generalize dependent x; intros
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type remove_all_binders] in *;
+ cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
destruct_head' prod;
change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
+ make_bound
+ Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
destruct_head' ZBounds.bounds;
@@ -534,12 +596,12 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
destruct_head' and;
Z.ltb_to_lt;
change WordW.wordWToZ with word64ToZ in *;
- cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd];
+ cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe25519_32WToZ HList.hlistP HList.hlistP'];
+ cbv [WordW.bit_width BitSize64.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
repeat split; unfold_is_bounded;
Z.ltb_to_lt;
try omega; try reflexivity.
-
Ltac rexpr_correct :=
let ropW' := fresh in
let ropZ_sig := fresh in
@@ -555,9 +617,13 @@ Ltac rexpr_correct :=
Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) opW) bounds)
+ := (Interp (@ZBounds.interp_op) opW bounds)
(only parsing).
+Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
+
+Ltac prove_rexpr_wfT
+ := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
Module Export PrettyPrinting.
(* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
@@ -569,23 +635,21 @@ Module Export PrettyPrinting.
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 {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end
+ := fun (t : base_type) (x : ZBounds.interp_base_type t)
+ => match x with
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
+ => in_range l u
+ | None
+ => overflow
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
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
+ := match t return interp_flat_type _ t -> result with
+ | Tbase _ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
| Unit => fun _ => no
| Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
| no, no => no
diff --git a/src/SpecificGen/GF25519_32Reflective/Common9_4Op.v b/src/SpecificGen/GF25519_32Reflective/Common9_4Op.v
index 286724558..e9b0209db 100644
--- a/src/SpecificGen/GF25519_32Reflective/Common9_4Op.v
+++ b/src/SpecificGen/GF25519_32Reflective/Common9_4Op.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -42,8 +41,8 @@ Lemma Expr9_4Op_correct_and_bounded
let (Hx7, Hx8) := (eta_and Hx78) in
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -80,29 +79,24 @@ Lemma Expr9_4Op_correct_and_bounded
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => op9_4_bounds_good bounds = true
| None => False
end)
: op9_4_correct_and_bounded ropW op.
Proof.
- intros x0 x1 x2 x3 x4 x5 x6 x7 x8.
- intros Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8.
- pose x0 as x0'.
- pose x1 as x1'.
- pose x2 as x2'.
- pose x3 as x3'.
- pose x4 as x4'.
- pose x5 as x5'.
- pose x6 as x6'.
- pose x7 as x7'.
- pose x8 as x8'.
- hnf in x0, x1, x2, x3, x4, x5, x6, x7, x8; destruct_head' prod.
- specialize (H0 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ intros xs Hxs.
+ pose xs as xs'.
+ compute in xs.
+ destruct_head' prod.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
+ pose Hxs as Hxs'.
+ destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
+ specialize (H0 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ specialize (H1 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded (x0', x1', x2', x3', x4', x5', x6', x7', x8') Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
+ Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.64s) (successful) *)*)
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v b/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
index 6a8a54fdd..48fc98b9d 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonBinOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -18,8 +17,8 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -33,18 +32,19 @@ Lemma ExprBinOp_correct_and_bounded
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
end)
: binop_correct_and_bounded 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
+ intros xy HxHy.
+ pose xy as xy'.
+ compute in xy; destruct_head' prod.
+ specialize (H0 xy' HxHy).
+ specialize (H1 xy' HxHy).
+ destruct HxHy as [Hx Hy].
+ let args := constr:(binop_args_to_bounded xy' 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
index 8df0c1685..8bfdba221 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOp_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
index d97e1e8c9..1bc40e3a9 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToWire.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
index 1a6f994e5..9d9e24690 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpFEToZ.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe25519_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
index 62572dc0c..3d6ba7d3b 100644
--- a/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF25519_32Reflective/CommonUnOpWireToFE.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
let args := unopWireToFE_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/AddCoordinates.v b/src/SpecificGen/GF25519_32Reflective/Reified/AddCoordinates.v
index e013a84e5..5d0b9deb7 100644
--- a/src/SpecificGen/GF25519_32Reflective/Reified/AddCoordinates.v
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/AddCoordinates.v
@@ -7,7 +7,6 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.ExprInversion.
Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
@@ -17,7 +16,10 @@ Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
+Require Import Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.Mul.
@@ -33,24 +35,28 @@ Require Import Bedrock.Word.
Definition radd_coordinatesZ' var twice_d P1 P2
:= @Extended.add_coordinates_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
twice_d _
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
P1 P2.
+Local Notation eta x := (fst x, snd x).
+
Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => radd_coordinatesZ'
- var (Return twice_d)
- (Return P10, Return P11, Return P12, Return P13)
- (Return P20, Return P21, Return P22, Return P23))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
+ => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
+ := twice_d_P1_P2 in
+ radd_coordinatesZ'
+ var (SmartVarf twice_d)
+ (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
+ (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
Definition add_coordinates
:= fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
@@ -59,70 +65,60 @@ Definition add_coordinates
twice_d (P10, P11, P12, P13) (P20, P21, P22, P23).
Definition uncurried_add_coordinates
- := apply9_nd
- (@uncurry_unop_fe25519_32)
- add_coordinates.
+ := fun twice_d_P1_P2
+ => let twice_d := fst twice_d_P1_P2 in
+ let (P1, P2) := eta (snd twice_d_P1_P2) in
+ @Extended.add_coordinates
+ _ add sub mul
+ twice_d P1 P2.
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
let H := fresh in
pose k as k';
assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
+ repeat (unfold k'; change k with k'; unfold interpf_step);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF25519_32.fe25519_32 -> GF25519_32.fe25519_32 -> GF25519_32.fe25519_32)
- (x' y' : GF25519_32.fe25519_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe25519_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
+Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe25519_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF25519_32.fe25519_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe25519_32 Apply' interp fst snd].
- rewrite (invert_Abs_Some (e:=f) eq_refl).
+ cbv [radd_coordinatesZ''].
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type] in x.
+ destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
+ | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
end.
- reflexivity.
-Qed.
-
-Lemma radd_coordinatesZ_sigP : rexpr_sigP Expr9_4OpT uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT radd_coordinatesZ'' uncurried_add_coordinates uncurry_unop_fe25519_32 SmartAbs radd_coordinatesZ' exprArg Extended.add_coordinates_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe25519_32 add_coordinates].
- intros.
- unfold invert_Return at 13 14 15 16.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [radd_coordinatesZ' add_coordinates Extended.add_coordinates_gen uncurried_add_coordinates SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -132,48 +128,56 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 13 14 15; unfold interpf_step.
- cbv beta iota delta [Extended.add_coordinates].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition radd_coordinatesZ_sig : rexpr_9_4op_sig add_coordinates.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv beta iota delta [Extended.add_coordinates interp_flat_type interp_base_type GF25519_32.fe25519_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- eexists.
- apply radd_coordinatesZ_sigP.
-Defined.
+ exact radd_coordinatesZ_sigP'.
+Qed.
+Definition radd_coordinatesZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
+
+Definition radd_coordinates_input_bounds
+ := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW Expr9Op_bounds.
+Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition radd_coordinatesW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- radd_coordinatesW add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
+ radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW Expr9Op_bounds).
+Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/LadderStep.v b/src/SpecificGen/GF25519_32Reflective/Reified/LadderStep.v
index ae71337fd..36919539d 100644
--- a/src/SpecificGen/GF25519_32Reflective/Reified/LadderStep.v
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/LadderStep.v
@@ -7,8 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Relations.
Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -18,6 +19,7 @@ Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
Require Import Crypto.Spec.MxDH.
+Require Import Crypto.SpecificGen.GF25519_32Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF25519_32Reflective.Reified.Mul.
@@ -30,50 +32,80 @@ Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
Require Import Bedrock.Word.
-(* XXX FIXME: Remove dummy code *)
-Definition rladderstepZ' var (T:=_) (dummy0 dummy1 dummy2 a24 x0 : T) P1 P2
+Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
:= @MxDH.ladderstep_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
a24
_
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
x0
P1 P2.
Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21
- => rladderstepZ'
- var (Return dummy0) (Return dummy1) (Return dummy2)
- (Return a24) (Return x0)
- (Return P10, Return P11)
- (Return P20, Return P21))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
+ => let '(a24, x0, ((P10, P11), (P20, P21)))
+ := a24_x0_P1_P2 in
+ rladderstepZ'
+ var (SmartVarf a24) (SmartVarf x0)
+ (SmartVarf P10, SmartVarf P11)
+ (SmartVarf P20, SmartVarf P21)))).
-Definition ladderstep (T := _)
- := fun (dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21 : T)
- => @MxDH.ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
+Local Notation eta x := (fst x, snd x).
+
+Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
+ Eval cbv beta delta [MxDH.ladderstep_gen] in
+ @MxDH.ladderstep_gen
+ F Fadd Fsub Fmul a24
+ (F*F*F*F)
+ (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
+ (fun x f => dlet y := x in f y)
+ X1 P1 P2.
Definition uncurried_ladderstep
- := apply9_nd
- (@uncurry_unop_fe25519_32)
- ladderstep.
+ := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
+ => let a24 := fst (fst a24_x0_P1_P2) in
+ let x0 := snd (fst a24_x0_P1_P2) in
+ let '(P1, P2) := eta (snd a24_x0_P1_P2) in
+ let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
+ @ladderstep_other_assoc
+ _ add sub mul
+ a24 x0 (P10, P11) (P20, P21).
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
@@ -83,51 +115,14 @@ Local Ltac repeat_step_interpf :=
repeat (unfold interpf_step at 1; change k with k' at 1);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF25519_32.fe25519_32 -> GF25519_32.fe25519_32 -> GF25519_32.fe25519_32)
- (x' y' : GF25519_32.fe25519_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe25519_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
-Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe25519_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF25519_32.fe25519_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe25519_32 Apply' interp fst snd].
- let f := match goal with f : expr _ _ _ |- _ => f end in
- rewrite (invert_Abs_Some (e:=f) eq_refl).
- repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
- end.
- reflexivity.
-Qed.
-
-Lemma rladderstepZ_sigP : rexpr_sigP Expr9_4OpT uncurried_ladderstep rladderstepZ''.
+Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
cbv [rladderstepZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT rladderstepZ'' uncurried_ladderstep uncurry_unop_fe25519_32 SmartAbs rladderstepZ' exprArg MxDH.ladderstep_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe25519_32 ladderstep].
- intros; cbv beta zeta.
- unfold invert_Return at 14 15 16 17.
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type interp_base_type] in x.
+ destruct_head' prod.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -137,48 +132,58 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 14 15 16; unfold interpf_step.
- cbv beta iota delta [MxDH.ladderstep_other_assoc].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition rladderstepZ_sig : rexpr_9_4op_sig ladderstep.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv [ladderstep_other_assoc interp_flat_type GF25519_32.fe25519_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In Common.ExprBinOpT] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | _ => progress cbv [interp_base_type]
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
- eexists.
- apply rladderstepZ_sigP.
-Defined.
+ exact rladderstepZ_sigP'.
+Qed.
+Definition rladderstepZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
+
+Definition rladderstep_input_bounds
+ := (ExprUnOp_bounds, ExprUnOp_bounds,
+ ((ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW Expr9Op_bounds.
+Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition rladderstepW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- rladderstepW ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
+ rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW Expr9Op_bounds).
+Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v
index 1ba921294..8cdacd63f 100644
--- a/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v
+++ b/src/SpecificGen/GF25519_32Reflective/Reified/PreFreeze.v
@@ -1,6 +1,6 @@
Require Import Crypto.SpecificGen.GF25519_32Reflective.CommonUnOp.
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF25519_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. 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.
diff --git a/src/SpecificGen/GF25519_32ReflectiveAddCoordinates.v b/src/SpecificGen/GF25519_32ReflectiveAddCoordinates.v
index 1392cb2f6..c2a9414f4 100644
--- a/src/SpecificGen/GF25519_32ReflectiveAddCoordinates.v
+++ b/src/SpecificGen/GF25519_32ReflectiveAddCoordinates.v
@@ -32,7 +32,7 @@ Import ListNotations.
Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Local Open Scope Z.
-Time Definition radd_coordinates : Expr9_4Op := Eval vm_compute in radd_coordinatesW.
+Time Definition radd_coordinates : Expr _ := Eval vm_compute in radd_coordinatesW.
Declare Reduction asm_interp_add_coordinates
:= cbv beta iota delta
@@ -57,7 +57,7 @@ Ltac asm_interp_add_coordinates
mapf_interp_flat_type WordW.interp_base_type word64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
-
+(*
Time Definition interp_radd_coordinates : SpecificGen.GF25519_32BoundedCommon.fe25519_32W
-> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
-> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
@@ -74,3 +74,4 @@ Time Definition interp_radd_coordinates_correct : interp_radd_coordinates = inte
Lemma radd_coordinates_correct_and_bounded : op9_4_correct_and_bounded radd_coordinates add_coordinates.
Proof. exact_no_check radd_coordinatesW_correct_and_bounded. Time Qed.
+*)
diff --git a/src/SpecificGen/GF25519_64Bounded.v b/src/SpecificGen/GF25519_64Bounded.v
index 8100182ed..e602af2e5 100644
--- a/src/SpecificGen/GF25519_64Bounded.v
+++ b/src/SpecificGen/GF25519_64Bounded.v
@@ -25,13 +25,23 @@ Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.
Local Open Scope Z.
+Local Ltac cbv_tuple_map :=
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'].
+
+Local Ltac post_bounded_t :=
+ (* much pain and hackery to work around [Defined] taking forever *)
+ cbv_tuple_map;
+ let blem' := fresh "blem'" in
+ let is_bounded_lem := fresh "is_bounded_lem" in
+ intros is_bounded_lem blem';
+ apply blem'; repeat apply conj; apply is_bounded_lem.
Local Ltac bounded_t opW blem :=
- apply blem; apply is_bounded_proj1_fe25519_64.
+ generalize blem; generalize is_bounded_proj1_fe25519_64; post_bounded_t.
Local Ltac bounded_wire_digits_t opW blem :=
- apply blem; apply is_bounded_proj1_wire_digits.
+ generalize blem; generalize is_bounded_proj1_wire_digits; post_bounded_t.
Local Ltac define_binop f g opW blem :=
- refine (exist_fe25519_64W (opW (proj1_fe25519_64W f) (proj1_fe25519_64W g)) _);
+ 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)) _);
@@ -47,17 +57,17 @@ Local Ltac define_unop_WireToFE f opW blem :=
Local Opaque Let_In.
Local Opaque Z.add Z.sub Z.mul Z.shiftl Z.shiftr Z.land Z.lor Z.eqb NToWord128.
-Local Arguments interp_radd / _ _.
-Local Arguments interp_rsub / _ _.
-Local Arguments interp_rmul / _ _.
+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 addW (f : fe25519_64W * fe25519_64W) : fe25519_64W := Eval simpl in interp_radd f.
+Definition subW (f : fe25519_64W * fe25519_64W) : fe25519_64W := Eval simpl in interp_rsub f.
+Definition mulW (f : fe25519_64W * fe25519_64W) : fe25519_64W := Eval simpl in interp_rmul f.
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) : word128 := Eval simpl in interp_rge_modulus f.
@@ -86,7 +96,7 @@ Definition freezeW (f : fe25519_64W) : fe25519_64W := Eval cbv beta delta [prefr
Local Transparent Let_In.
(* Wrapper to allow extracted code to not unfold [mulW] *)
Definition mulW_noinline := mulW.
-Definition powW (f : fe25519_64W) chain := fold_chain_opt (proj1_fe25519_64W one) mulW_noinline chain [f].
+Definition powW (f : fe25519_64W) chain := fold_chain_opt (proj1_fe25519_64W one) (fun f g => mulW_noinline (f, g)) chain [f].
Definition invW (f : fe25519_64W) : fe25519_64W
:= Eval cbv -[Let_In fe25519_64W mulW_noinline] in powW f (chain inv_ec).
@@ -95,11 +105,11 @@ Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
rewrite pre_rewrite;
intros; apply rop_cb; assumption.
-Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW (Curry.curry2 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.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW (Curry.curry2 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.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW (Curry.curry2 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.
@@ -142,6 +152,7 @@ Proof.
cbv [modulusW Tuple.map].
cbv [on_tuple List.map to_list to_list' from_list from_list'
+ HList.hlistP HList.hlistP'
Tuple.map2 on_tuple2 ListUtil.map2 fe25519_64WToZ length_fe25519_64].
cbv [postfreeze GF25519_64.postfreeze].
cbv [Let_In].
@@ -203,7 +214,8 @@ Proof.
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.
+ split; [ | assumption ].
+ rewrite H0', H0; reflexivity.
Qed.
Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
@@ -212,9 +224,11 @@ Proof.
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; pose proof (fun k0 k1 X Y => proj1 (mulW_correct_and_bounded (k0, k1) (conj X Y))) as H'.
+ cbv [Curry.curry2 Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list'] in H'.
+ rewrite <- H' by assumption.
+ apply mulW_correct_and_bounded; split; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded (_, _) (conj X Y))) by assumption; reflexivity. }
{ intros [|?]; autorewrite with simpl_nth_default;
(assumption || reflexivity). }
Qed.
@@ -269,8 +283,10 @@ Proof.
unfold GF25519_64.eqb.
simpl @fe25519_64WToZ in *; cbv beta iota.
intros.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe25519_64WToZ] in *.
rewrite <- frf, <- frg by assumption.
- rewrite <- fieldwisebW_correct.
+ etransitivity; [ eapply fieldwisebW_correct | ].
+ cbv [fe25519_64WToZ].
reflexivity.
Defined.
@@ -297,7 +313,7 @@ Proof.
lazymatch (eval cbv delta [GF25519_64.sqrt] in GF25519_64.sqrt) with
| (fun powf powf_squared f => dlet a := powf in _)
=> exact (dlet powx := powW (fe25519_64ZToW x) (chain GF25519_64.sqrt_ec) in
- GF25519_64.sqrt (fe25519_64WToZ powx) (fe25519_64WToZ (mulW_noinline powx powx)) x)
+ GF25519_64.sqrt (fe25519_64WToZ powx) (fe25519_64WToZ (mulW_noinline (powx, powx))) x)
| (fun f => pow f _)
=> exact (GF25519_64.sqrt x)
end.
@@ -324,21 +340,41 @@ Proof.
=> 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
- | mulW_noinline _ _ => apply mulW_correct_and_bounded
- | powW _ _ => apply powW_correct_and_bounded
- | sqrt_m1W => vm_compute; reflexivity
- | _ => assumption
- end
- end;
- subst_evars; reflexivity
+ pose proof (fun X Y => proj1 (mulW_correct_and_bounded (sqrt_m1W, a) (conj X Y))) as correctness;
+ let cbv_in_all _ := (cbv [fe25519_64WToZ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe25519_64WToZ Curry.curry2 HList.hlistP HList.hlistP'] in *; idtac) in
+ cbv_in_all ();
+ let solver _ := (repeat match goal with
+ | _ => progress subst
+ | _ => progress unfold fst, snd
+ | _ => progress cbv_in_all ()
+ | [ |- ?x /\ ?x ] => cut x; [ intro; split; assumption | ]
+ | [ |- is_bounded ?op = true ]
+ => let H := fresh in
+ lazymatch op with
+ | context[mulW (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[mulW_noinline (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[powW _ _] => pose proof powW_correct_and_bounded as H
+ | context[sqrt_m1W] => vm_compute; reflexivity
+ | _ => assumption
+ end;
+ cbv_in_all ();
+ apply H
+ end) in
+ rewrite <- correctness by solver (); clear correctness;
+ let lem := fresh in
+ pose proof eqbW_correct as lem; cbv_in_all (); rewrite <- lem by solver (); clear lem;
+ pose proof (pull_bool_if fe25519_64WToZ) as lem; cbv_in_all (); rewrite lem by solver (); clear lem;
+ subst_evars; reflexivity
end.
} Unfocus.
+ assert (Hfold : forall x, fe25519_64WToZ x = fe25519_64WToZ x) by reflexivity.
+ unfold fe25519_64WToZ at 2 in Hfold.
+ etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody; [ reflexivity | intro ].
+ apply Hfold.
+ } Unfocus.
+ clear Hfold.
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
@@ -346,15 +382,22 @@ Proof.
[ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
| _ => idtac
end;
- reflexivity. }
- { cbv [Let_In];
+ reflexivity.
+ }
+
+ { cbv [Let_In HList.hlistP HList.hlistP'];
try break_if;
repeat lazymatch goal with
| [ |- is_bounded (?WToZ (powW _ _)) = true ]
=> apply powW_correct_and_bounded; assumption
- | [ |- is_bounded (?WToZ (mulW _ _)) = true ]
- => apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ]
- end. }
+ | [ |- is_bounded (snd (?WToZ (_, powW _ _))) = true ]
+ => generalize powW_correct_and_bounded;
+ cbv [snd Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'];
+ let H := fresh in intro H; apply H; assumption
+ | [ |- is_bounded (?WToZ (mulW (_, _))) = true ]
+ => apply mulW_correct_and_bounded; split; [ vm_compute; reflexivity | ]
+ end.
+ }
Defined.
Definition sqrtW (f : fe25519_64W) : fe25519_64W :=
@@ -394,7 +437,7 @@ 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.
+Proof. define_unop f invW (fun x p => proj2 (invW_correct_and_bounded x p)). Defined.
Definition sqrt (f : fe25519_64) : fe25519_64.
Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
@@ -407,7 +450,12 @@ Local Ltac op_correct_t op opW_correct_and_bounded :=
=> rewrite proj1_wire_digits_exist_wire_digitsW
| _ => idtac
end;
- apply opW_correct_and_bounded;
+ generalize opW_correct_and_bounded;
+ cbv_tuple_map;
+ cbv [fst snd];
+ let H := fresh in
+ intro H; apply H;
+ repeat match goal with |- and _ _ => apply conj end;
lazymatch goal with
| [ |- is_bounded _ = true ]
=> apply is_bounded_proj1_fe25519_64
@@ -434,7 +482,7 @@ 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.
+Proof. op_correct_t inv (fun x p => proj1 (invW_correct_and_bounded x p)). 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.
diff --git a/src/SpecificGen/GF25519_64BoundedAddCoordinates.v b/src/SpecificGen/GF25519_64BoundedAddCoordinates.v
index a1ba335c8..711af3078 100644
--- a/src/SpecificGen/GF25519_64BoundedAddCoordinates.v
+++ b/src/SpecificGen/GF25519_64BoundedAddCoordinates.v
@@ -14,7 +14,7 @@ Local Ltac define_binop f g opW blem :=
Local Opaque Let_In.
Local Opaque Z.add Z.sub Z.mul Z.shiftl Z.shiftr Z.land Z.lor Z.eqb NToWord128.
-Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
+(*Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
Definition add_coordinatesW (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe25519_64W) : Tuple.tuple fe25519_64W 4
:= Eval simpl in interp_radd_coordinates x0 x1 x2 x3 x4 x5 x6 x7 x8.
@@ -75,3 +75,4 @@ Lemma add_coordinates_correct (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe25519_64)
(proj1_fe25519_64 x7)
(proj1_fe25519_64 x8).
Proof. op_correct_t add_coordinates add_coordinatesW_correct_and_bounded. Qed.
+*)
diff --git a/src/SpecificGen/GF25519_64BoundedCommon.v b/src/SpecificGen/GF25519_64BoundedCommon.v
index b4c8e47ca..1c91f4d88 100644
--- a/src/SpecificGen/GF25519_64BoundedCommon.v
+++ b/src/SpecificGen/GF25519_64BoundedCommon.v
@@ -322,14 +322,14 @@ Ltac unfold_is_bounded_in' H :=
end.
Ltac preunfold_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.
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe25519_64 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
Ltac unfold_is_bounded_in H :=
preunfold_is_bounded_in H;
unfold_is_bounded_in' H.
Ltac preunfold_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].
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe25519_64 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
Ltac unfold_is_bounded :=
preunfold_is_bounded;
@@ -724,7 +724,7 @@ Notation in_op_correct_and_bounded k irop op
/\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe25519_64WToZ irop))%type)
(only parsing).
-Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
+(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
: forall (irop : Tower.tower_nd fe25519_64W (Tuple.tuple fe25519_64W count_out) count_in)
(op : Tower.tower_nd GF25519_64.fe25519_64 (Tuple.tuple GF25519_64.fe25519_64 count_out) count_in)
(cont : Prop -> Prop),
@@ -792,18 +792,14 @@ Qed.
Definition inm_op_correct_and_bounded1 count_in irop op
:= Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.
-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).
+ inm_op_correct_and_bounded count_in 1 irop op.*)
+Notation inm_op_correct_and_bounded n m irop op
+ := ((forall x,
+ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe25519_64WToZ x)
+ -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe25519_64WToZ x))))
+ (only parsing).
+Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
+Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe25519_64WToZ x) = true
@@ -818,20 +814,6 @@ Notation iunop_WireToFE_correct_and_bounded irop op
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).
-Notation i9top_correct_and_bounded k irop op
- := ((forall x0 x1 x2 x3 x4 x5 x6 x7 x8,
- is_bounded (fe25519_64WToZ x0) = true
- -> is_bounded (fe25519_64WToZ x1) = true
- -> is_bounded (fe25519_64WToZ x2) = true
- -> is_bounded (fe25519_64WToZ x3) = true
- -> is_bounded (fe25519_64WToZ x4) = true
- -> is_bounded (fe25519_64WToZ x5) = true
- -> is_bounded (fe25519_64WToZ x6) = true
- -> is_bounded (fe25519_64WToZ x7) = true
- -> is_bounded (fe25519_64WToZ x8) = true
- -> (Tuple.map (n:=k) fe25519_64WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)
- = op (fe25519_64WToZ x0) (fe25519_64WToZ x1) (fe25519_64WToZ x2) (fe25519_64WToZ x3) (fe25519_64WToZ x4) (fe25519_64WToZ x5) (fe25519_64WToZ x6) (fe25519_64WToZ x7) (fe25519_64WToZ x8))
- * HList.hlist (fun v => is_bounded v = true) (Tuple.map (n:=k) fe25519_64WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)))%type)
- (only parsing).
+Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-Definition prefreeze := GF25519_64.prefreeze.
+Notation prefreeze := GF25519_64.prefreeze.
diff --git a/src/SpecificGen/GF25519_64BoundedExtendedAddCoordinates.v b/src/SpecificGen/GF25519_64BoundedExtendedAddCoordinates.v
index d5e92dfac..b7d07d9c2 100644
--- a/src/SpecificGen/GF25519_64BoundedExtendedAddCoordinates.v
+++ b/src/SpecificGen/GF25519_64BoundedExtendedAddCoordinates.v
@@ -5,7 +5,7 @@ Require Import Crypto.SpecificGen.GF25519_64ExtendedAddCoordinates.
Require Import Crypto.SpecificGen.GF25519_64BoundedAddCoordinates.
Require Import Crypto.Util.Tuple.
Require Import Crypto.Util.Tactics.
-
+(*
Lemma fieldwise_eq_extended_add_coordinates_full' twice_d P10 P11 P12 P13 P20 P21 P22 P23
: Tuple.fieldwise
(n:=4) GF25519_64BoundedCommon.eq
@@ -65,3 +65,4 @@ Proof.
destruct_head' prod.
rewrite <- fieldwise_eq_extended_add_coordinates_full'; reflexivity.
Qed.
+*)
diff --git a/src/SpecificGen/GF25519_64Reflective.v b/src/SpecificGen/GF25519_64Reflective.v
index b92c70202..142c1e6ae 100644
--- a/src/SpecificGen/GF25519_64Reflective.v
+++ b/src/SpecificGen/GF25519_64Reflective.v
@@ -45,6 +45,7 @@ Declare Reduction asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_64W curry_unop_fe25519_64W curry_unop_wire_digitsW curry_9op_fe25519_64W
WordW.interp_op WordW.interp_base_type
@@ -56,6 +57,7 @@ Ltac asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe25519_64W curry_unop_fe25519_64W curry_unop_wire_digitsW curry_9op_fe25519_64W
WordW.interp_op WordW.interp_base_type
@@ -65,15 +67,15 @@ Ltac asm_interp
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
+Definition interp_radd : SpecificGen.GF25519_64BoundedCommon.fe25519_64W * SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rsub : SpecificGen.GF25519_64BoundedCommon.fe25519_64W * SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rmul : SpecificGen.GF25519_64BoundedCommon.fe25519_64W * SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
:= Eval asm_interp in interp_bexpr rmul.
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
diff --git a/src/SpecificGen/GF25519_64Reflective/Common.v b/src/SpecificGen/GF25519_64Reflective/Common.v
index b592451eb..edb0c5fb1 100644
--- a/src/SpecificGen/GF25519_64Reflective/Common.v
+++ b/src/SpecificGen/GF25519_64Reflective/Common.v
@@ -7,7 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Wf.
Require Import Crypto.Reflection.ExprInversion.
+Require Import Crypto.Reflection.Tuple.
Require Import Crypto.Reflection.Relations.
+Require Import Crypto.Reflection.Eta.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Crypto.Reflection.Z.Interpretations128.Relations.
Require Import Crypto.Reflection.Z.Interpretations128.RelationsCombinations.
@@ -15,13 +17,14 @@ Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Z.Syntax.Equality.
Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.Curry.
Require Import Crypto.Util.Tower.
Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Notations.
Notation Expr := (Expr base_type op).
@@ -43,30 +46,24 @@ Defined.
Definition Expr_n_OpT (count_out : nat) : flat_type base_type
:= Eval cbv [Syntax.tuple Syntax.tuple' fe25519_64T] in
Syntax.tuple fe25519_64T count_out.
-Fixpoint Expr_nm_OpT (count_in count_out : nat) : type base_type
- := match count_in with
- | 0 => Expr_n_OpT count_out
- | S n => SmartArrow fe25519_64T (Expr_nm_OpT n count_out)
- end.
+Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
+ := Eval cbv [Syntax.tuple Syntax.tuple' fe25519_64T Expr_n_OpT] in
+ Arrow (Syntax.tuple fe25519_64T count_in) (Expr_n_OpT count_out).
Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from (uncurry_unop_fe25519_64 ge_modulus). Defined.
+Proof. make_type_from ge_modulus. Defined.
Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Proof. make_type_from unpack. Defined.
Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from (uncurry_unop_fe25519_64 pack). Defined.
+Proof. make_type_from pack. Defined.
Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
Definition ExprArgT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpT.
+ := Eval compute in domain ExprUnOpT.
Definition ExprArgWireT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpFEToWireT.
-Definition ExprArgRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpT.
-Definition ExprArgWireRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Tbase TZ).
+ := Eval compute in domain ExprUnOpWireToFET.
+Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
@@ -75,99 +72,67 @@ Definition ExprBinOp : Type := Expr ExprBinOpT.
Definition ExprUnOp : Type := Expr ExprUnOpT.
Definition Expr4Op : Type := Expr Expr4OpT.
Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr ExprArgT.
-Definition ExprArgWire : Type := Expr ExprArgWireT.
-Definition ExprArgRev : Type := Expr ExprArgRevT.
-Definition ExprArgWireRev : Type := Expr ExprArgWireRevT.
+Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
+Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
Definition expr_nm_Op count_in count_out var : Type
:= expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Tbase TZ).
+Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) ExprArgT.
-Definition exprArgWire var : Type := expr base_type op (var:=var) ExprArgWireT.
-Definition exprArgRev var : Type := expr base_type op (var:=var) ExprArgRevT.
-Definition exprArgWireRev var : Type := expr base_type op (var:=var) ExprArgWireRevT.
-
-Local Ltac bounds_from_list_cps ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(fun T (extra : T) => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, extra))
- | (?x :: ?xs)%list => let bs := bounds_from_list_cps xs in
- constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
- end.
-
-Local Ltac make_bounds_cps ls extra :=
- let v := bounds_from_list_cps (List.rev ls) in
- let v := (eval compute in v) in
- exact (v _ extra).
-
-Local Ltac bounds_from_list ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
- | (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| Bounds.lower := fst x ; Bounds.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.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out : interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type.
-Proof.
- refine match count_in return interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type with
- | 0 => tt
- | S n => let v := interp_all_binders_for_to' (Expr_nm_Op_bounds n count_out) in
- interp_all_binders_for_of' _
- end; simpl.
- make_bounds_cps (Tuple.to_list _ bounds) v.
-Defined.
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
+Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
+
+Definition make_bound (x : Z * Z) : ZBounds.t
+ := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
+
+Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
+ := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
+ | 0 => tt
+ | S n
+ => let b := (Tuple.map make_bound bounds) in
+ let bs := Expr_nm_Op_bounds n count_out in
+ match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
+ | 0 => fun _ => b
+ | S n' => fun bs => (bs, b)
+ end bs
+ end.
+Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
:= Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
:= Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
:= Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
+Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
+ := Tuple.map make_bound wire_digit_bounds.
-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 (@WordW.interp_op) e).
+Definition interp_bexpr : ExprBinOp -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W * SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := fun e => curry_unop_fe25519_64W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.word128
- := fun e => curry_unop_fe25519_64W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W -> SpecificGen.GF25519_64BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe25519_64W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF25519_64BoundedCommon.wire_digitsW -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_9_4expr : Expr9_4Op
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
- -> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
+ -> Tuple.tuple SpecificGen.GF25519_64BoundedCommon.fe25519_64W 9
-> Tuple.tuple SpecificGen.GF25519_64BoundedCommon.fe25519_64W 4
- := fun e => curry_9op_fe25519_64W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+ := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 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
@@ -181,40 +146,39 @@ Notation op9_4_correct_and_bounded rop op
Ltac rexpr_cbv :=
lazymatch goal with
- | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
=> let operf := head oper in
let uncurryf := head uncurry in
try cbv delta [T]; try cbv delta [oper];
try cbv beta iota delta [uncurryf]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
+ => let operf := head oper in
+ try cbv delta [T]; try cbv delta [oper]
end;
- cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+ cbv beta iota delta [interp_flat_type interp_base_type zero_ GF25519_64.fe25519_64 GF25519_64.wire_digits].
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 }
+ | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
(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 rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT (uncurry_9op_fe25519_64 op)) (only parsing).
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
+Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT 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 := ropW' in
- let ropZ := ropW' in
- let ropBounds := ropW' in
- let ropBoundedWordW := ropW' in
- ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
+ ropW' = proj1_sig ropZ_sig
+ /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
(only parsing).
Ltac app_tuples x y :=
@@ -227,7 +191,7 @@ Ltac app_tuples x y :=
Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
-
+(*
Fixpoint args_to_bounded_helperT {n}
(v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
@@ -299,14 +263,15 @@ Proof.
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
@@ -327,6 +292,7 @@ Ltac assoc_right_tuple x so_far :=
end
end.
+(*
Local Ltac make_args x :=
let x' := fresh "x'" in
compute in x |- *;
@@ -338,11 +304,6 @@ Local Ltac make_args x :=
let xv := assoc_right_tuple x'' (@None) in
refine (SmartVarf (xv : interp_flat_type _ t')).
-Definition unop_make_args {var} (x : exprArg var) : exprArgRev var.
-Proof. make_args x. Defined.
-Definition unop_wire_make_args {var} (x : exprArgWire var) : exprArgWireRev var.
-Proof. make_args x. Defined.
-
Local Ltac args_to_bounded x H :=
let x' := fresh in
set (x' := x);
@@ -351,29 +312,138 @@ Local Ltac args_to_bounded x H :=
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
+ let c := (eval cbv [domain 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) ]
).
+ *)
+
+Section gen.
+ Definition bounds_are_good_gen
+ {n : nat} (bounds : Tuple.tuple (Z * Z) n)
+ := let res :=
+ Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
+ in
+ List.fold_right andb true (Tuple.to_list n res).
+ Definition unop_args_to_bounded'
+ (bs : Z * Z)
+ (Hbs : bounds_are_good_gen (n:=1) bs = true)
+ (x : word128)
+ (H : is_bounded_gen (Tuple.map (fun v : word128 => (v : Z)) (n:=1) x) bs = true)
+ : BoundedWordW.BoundedWord.
+ Proof.
+ refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
+ unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
+ abstract (
+ destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
+ repeat apply conj; assumption
+ ).
+ Defined.
+ Fixpoint n_op_args_to_bounded'
+ n
+ : forall (bs : Tuple.tuple' (Z * Z) n)
+ (Hbs : bounds_are_good_gen (n:=S n) bs = true)
+ (x : Tuple.tuple' word128 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word128 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
+ Proof.
+ destruct n as [|n']; simpl in *.
+ { exact unop_args_to_bounded'. }
+ { refine (fun bs Hbs x H
+ => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
+ @unop_args_to_bounded' (snd bs) _ (snd x) _));
+ clear n_op_args_to_bounded';
+ simpl in *;
+ [ clear x H | clear Hbs | clear x H | clear Hbs ];
+ unfold bounds_are_good_gen, is_bounded_gen in *;
+ abstract (
+ repeat first [ progress simpl in *
+ | assumption
+ | reflexivity
+ | progress Bool.split_andb
+ | progress destruct_head prod
+ | match goal with
+ | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
+ end
+ | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
+ | progress break_match_hyps
+ | rewrite Bool.andb_true_iff; apply conj
+ | unfold Tuple.map, Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
+ ). }
+ Defined.
+
+ Definition n_op_args_to_bounded
+ n
+ : forall (bs : Tuple.tuple (Z * Z) n)
+ (Hbs : bounds_are_good_gen bs = true)
+ (x : Tuple.tuple word128 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word128 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
+ := match n with
+ | 0 => fun _ _ _ _ => tt
+ | S n' => @n_op_args_to_bounded' n'
+ end.
-Definition unop_args_to_bounded (x : fe25519_64W) (H : is_bounded (fe25519_64WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
+ Fixpoint nm_op_args_to_bounded' n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word128 m) (Syntax.tuple' (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word128 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => @n_op_args_to_bounded m bs Hbs
+ | S n' => fun x H
+ => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
+ @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
+ end.
+ Definition nm_op_args_to_bounded n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word128 m) (Syntax.tuple (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word128 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => fun _ _ => tt
+ | S n' => @nm_op_args_to_bounded' n' m bs Hbs
+ end.
+End gen.
+
+Local Ltac get_inner_len T :=
+ lazymatch T with
+ | (?T * _)%type => get_inner_len T
+ | ?T => get_len T
+ end.
+Local Ltac get_outer_len T :=
+ lazymatch T with
+ | (?A * ?B)%type => let a := get_outer_len A in
+ let b := get_outer_len B in
+ (eval compute in (a + b)%nat)
+ | ?T => constr:(1%nat)
+ end.
+Local Ltac args_to_bounded x H :=
+ let t := type of x in
+ let m := get_inner_len t in
+ let n := get_outer_len t in
+ let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
+ let H := (eval cbv beta in (H eq_refl)) in
+ exact H.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe25519_64W * fe25519_64W)
(H : is_bounded (fe25519_64WToZ (fst x)) = true)
(H' : is_bounded (fe25519_64WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
-Proof.
- let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
- exact v.
-Defined.
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
+Proof. args_to_bounded x (conj H H'). Defined.
+Definition unop_args_to_bounded (x : fe25519_64W) (H : is_bounded (fe25519_64WToZ x) = true)
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain 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 _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
+Proof. args_to_bounded x H. Defined.
Definition op9_args_to_bounded (x : fe25519_64W * fe25519_64W * fe25519_64W * fe25519_64W * fe25519_64W * fe25519_64W * fe25519_64W * fe25519_64W * fe25519_64W)
(H0 : is_bounded (fe25519_64WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
(H1 : is_bounded (fe25519_64WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
@@ -384,58 +454,39 @@ Definition op9_args_to_bounded (x : fe25519_64W * fe25519_64W * fe25519_64W * fe
(H6 : is_bounded (fe25519_64WToZ (snd (fst (fst x)))) = true)
(H7 : is_bounded (fe25519_64WToZ (snd (fst x))) = true)
(H8 : is_bounded (fe25519_64WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for Expr9_4OpT).
-Proof.
- let v := constr:(unop_args_to_bounded _ H8) in
- let v := app_tuples (unop_args_to_bounded _ H7) v in
- let v := app_tuples (unop_args_to_bounded _ H6) v in
- let v := app_tuples (unop_args_to_bounded _ H5) v in
- let v := app_tuples (unop_args_to_bounded _ H4) v in
- let v := app_tuples (unop_args_to_bounded _ H3) v in
- let v := app_tuples (unop_args_to_bounded _ H2) v in
- let v := app_tuples (unop_args_to_bounded _ H1) v in
- let v := app_tuples (unop_args_to_bounded _ H0) v in
- exact v.
-Defined.
-
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
+Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
+Local Ltac make_bounds_prop' bounds bounds' :=
+ first [ refine (andb _ _);
+ [ destruct bounds' as [bounds' _], bounds as [bounds _]
+ | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
+ try make_bounds_prop' bounds bounds'
+ | 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) ].
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.
+ make_bounds_prop' bounds bounds'.
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
Proof.
refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders Expr9_4OpT)) : bool.
+Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-
-Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
+(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe25519_64 f k)).
@@ -460,12 +511,11 @@ Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ v
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe25519_64 f k)).
-
+*)
(* 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
@@ -473,23 +523,25 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
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 interp_9_4expr
+ interp_flat_type_eta interp_flat_type_eta_gen
curry_binop_fe25519_64W curry_unop_fe25519_64W curry_unop_wire_digitsW curry_9op_fe25519_64W
curry_binop_fe25519_64 curry_unop_fe25519_64 curry_unop_wire_digits curry_9op_fe25519_64
uncurry_binop_fe25519_64W uncurry_unop_fe25519_64W uncurry_unop_wire_digitsW uncurry_9op_fe25519_64W
uncurry_binop_fe25519_64 uncurry_unop_fe25519_64 uncurry_unop_wire_digits uncurry_9op_fe25519_64
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT
- interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] 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 *;
+ change word128 with WordW.wordW in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
+ | progress unfold RelationClasses.Reflexive
| reflexivity
| assumption
| progress destruct_head' base_type
@@ -509,23 +561,33 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
+ Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
+ simpl @interp_flat_type in *;
+ (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
+ change (WordW.interp_base_type TZ) with v in *);
+ cbv beta iota zeta in *;
lazymatch goal with
| [ |- fe25519_64WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
| [ |- wire_digitsWToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
+ | [ |- (Tuple.map fe25519_64WToZ ?x = _) /\ _ ]
+ => generalize dependent x; intros
| [ |- ((Tuple.map fe25519_64WToZ ?x = _) * _)%type ]
=> generalize dependent x; intros
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type remove_all_binders] in *;
+ cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
destruct_head' prod;
change word128ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
+ make_bound
+ Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
destruct_head' ZBounds.bounds;
@@ -534,12 +596,12 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
destruct_head' and;
Z.ltb_to_lt;
change WordW.wordWToZ with word128ToZ in *;
- cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd];
+ cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe25519_64WToZ HList.hlistP HList.hlistP'];
+ cbv [WordW.bit_width BitSize128.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
repeat split; unfold_is_bounded;
Z.ltb_to_lt;
try omega; try reflexivity.
-
Ltac rexpr_correct :=
let ropW' := fresh in
let ropZ_sig := fresh in
@@ -555,9 +617,13 @@ Ltac rexpr_correct :=
Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) opW) bounds)
+ := (Interp (@ZBounds.interp_op) opW bounds)
(only parsing).
+Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
+
+Ltac prove_rexpr_wfT
+ := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
Module Export PrettyPrinting.
(* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
@@ -569,23 +635,21 @@ Module Export PrettyPrinting.
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 {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end
+ := fun (t : base_type) (x : ZBounds.interp_base_type t)
+ => match x with
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
+ => in_range l u
+ | None
+ => overflow
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
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
+ := match t return interp_flat_type _ t -> result with
+ | Tbase _ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
| Unit => fun _ => no
| Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
| no, no => no
diff --git a/src/SpecificGen/GF25519_64Reflective/Common9_4Op.v b/src/SpecificGen/GF25519_64Reflective/Common9_4Op.v
index 9435884a4..a1443cefa 100644
--- a/src/SpecificGen/GF25519_64Reflective/Common9_4Op.v
+++ b/src/SpecificGen/GF25519_64Reflective/Common9_4Op.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -42,8 +41,8 @@ Lemma Expr9_4Op_correct_and_bounded
let (Hx7, Hx8) := (eta_and Hx78) in
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -80,29 +79,24 @@ Lemma Expr9_4Op_correct_and_bounded
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => op9_4_bounds_good bounds = true
| None => False
end)
: op9_4_correct_and_bounded ropW op.
Proof.
- intros x0 x1 x2 x3 x4 x5 x6 x7 x8.
- intros Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8.
- pose x0 as x0'.
- pose x1 as x1'.
- pose x2 as x2'.
- pose x3 as x3'.
- pose x4 as x4'.
- pose x5 as x5'.
- pose x6 as x6'.
- pose x7 as x7'.
- pose x8 as x8'.
- hnf in x0, x1, x2, x3, x4, x5, x6, x7, x8; destruct_head' prod.
- specialize (H0 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ intros xs Hxs.
+ pose xs as xs'.
+ compute in xs.
+ destruct_head' prod.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
+ pose Hxs as Hxs'.
+ destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
+ specialize (H0 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ specialize (H1 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded (x0', x1', x2', x3', x4', x5', x6', x7', x8') Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
+ Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.128s) (successful) *)*)
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v b/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
index 1d174ae2b..9f26d8521 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonBinOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -18,8 +17,8 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -33,18 +32,19 @@ Lemma ExprBinOp_correct_and_bounded
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
end)
: binop_correct_and_bounded 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
+ intros xy HxHy.
+ pose xy as xy'.
+ compute in xy; destruct_head' prod.
+ specialize (H0 xy' HxHy).
+ specialize (H1 xy' HxHy).
+ destruct HxHy as [Hx Hy].
+ let args := constr:(binop_args_to_bounded xy' 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
index a5b57d394..aa26e495a 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOp_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
index cdda3c4eb..f1f30d64c 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToWire.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
index d082f70b0..48a5b9610 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpFEToZ.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe25519_64WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
index 327a6b825..3e7112480 100644
--- a/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF25519_64Reflective/CommonUnOpWireToFE.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF25519_64BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
let args := unopWireToFE_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/AddCoordinates.v b/src/SpecificGen/GF25519_64Reflective/Reified/AddCoordinates.v
index 0f41ed239..3fe4fbeef 100644
--- a/src/SpecificGen/GF25519_64Reflective/Reified/AddCoordinates.v
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/AddCoordinates.v
@@ -7,7 +7,6 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.ExprInversion.
Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Crypto.Reflection.Z.Interpretations128.Relations.
@@ -17,7 +16,10 @@ Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
+Require Import Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.Mul.
@@ -33,24 +35,28 @@ Require Import Bedrock.Word.
Definition radd_coordinatesZ' var twice_d P1 P2
:= @Extended.add_coordinates_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
twice_d _
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
P1 P2.
+Local Notation eta x := (fst x, snd x).
+
Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => radd_coordinatesZ'
- var (Return twice_d)
- (Return P10, Return P11, Return P12, Return P13)
- (Return P20, Return P21, Return P22, Return P23))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
+ => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
+ := twice_d_P1_P2 in
+ radd_coordinatesZ'
+ var (SmartVarf twice_d)
+ (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
+ (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
Definition add_coordinates
:= fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
@@ -59,70 +65,60 @@ Definition add_coordinates
twice_d (P10, P11, P12, P13) (P20, P21, P22, P23).
Definition uncurried_add_coordinates
- := apply9_nd
- (@uncurry_unop_fe25519_64)
- add_coordinates.
+ := fun twice_d_P1_P2
+ => let twice_d := fst twice_d_P1_P2 in
+ let (P1, P2) := eta (snd twice_d_P1_P2) in
+ @Extended.add_coordinates
+ _ add sub mul
+ twice_d P1 P2.
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
let H := fresh in
pose k as k';
assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
+ repeat (unfold k'; change k with k'; unfold interpf_step);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF25519_64.fe25519_64 -> GF25519_64.fe25519_64 -> GF25519_64.fe25519_64)
- (x' y' : GF25519_64.fe25519_64)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe25519_64 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
+Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe25519_64 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF25519_64.fe25519_64] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe25519_64 Apply' interp fst snd].
- rewrite (invert_Abs_Some (e:=f) eq_refl).
+ cbv [radd_coordinatesZ''].
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type] in x.
+ destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
+ | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
end.
- reflexivity.
-Qed.
-
-Lemma radd_coordinatesZ_sigP : rexpr_sigP Expr9_4OpT uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT radd_coordinatesZ'' uncurried_add_coordinates uncurry_unop_fe25519_64 SmartAbs radd_coordinatesZ' exprArg Extended.add_coordinates_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe25519_64 add_coordinates].
- intros.
- unfold invert_Return at 13 14 15 16.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [radd_coordinatesZ' add_coordinates Extended.add_coordinates_gen uncurried_add_coordinates SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -132,48 +128,56 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 13 14 15; unfold interpf_step.
- cbv beta iota delta [Extended.add_coordinates].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition radd_coordinatesZ_sig : rexpr_9_4op_sig add_coordinates.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv beta iota delta [Extended.add_coordinates interp_flat_type interp_base_type GF25519_64.fe25519_64].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- eexists.
- apply radd_coordinatesZ_sigP.
-Defined.
+ exact radd_coordinatesZ_sigP'.
+Qed.
+Definition radd_coordinatesZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
+
+Definition radd_coordinates_input_bounds
+ := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW Expr9Op_bounds.
+Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition radd_coordinatesW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- radd_coordinatesW add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
+ radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW Expr9Op_bounds).
+Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/LadderStep.v b/src/SpecificGen/GF25519_64Reflective/Reified/LadderStep.v
index 66d998d08..f3d2204b0 100644
--- a/src/SpecificGen/GF25519_64Reflective/Reified/LadderStep.v
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/LadderStep.v
@@ -7,8 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Relations.
Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.Reflection.Z.Interpretations128.
Require Crypto.Reflection.Z.Interpretations128.Relations.
Require Import Crypto.Reflection.Z.Interpretations128.RelationsCombinations.
@@ -18,6 +19,7 @@ Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
Require Import Crypto.Spec.MxDH.
+Require Import Crypto.SpecificGen.GF25519_64Reflective.Common.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF25519_64Reflective.Reified.Mul.
@@ -30,50 +32,80 @@ Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
Require Import Bedrock.Word.
-(* XXX FIXME: Remove dummy code *)
-Definition rladderstepZ' var (T:=_) (dummy0 dummy1 dummy2 a24 x0 : T) P1 P2
+Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
:= @MxDH.ladderstep_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
a24
_
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
x0
P1 P2.
Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21
- => rladderstepZ'
- var (Return dummy0) (Return dummy1) (Return dummy2)
- (Return a24) (Return x0)
- (Return P10, Return P11)
- (Return P20, Return P21))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
+ => let '(a24, x0, ((P10, P11), (P20, P21)))
+ := a24_x0_P1_P2 in
+ rladderstepZ'
+ var (SmartVarf a24) (SmartVarf x0)
+ (SmartVarf P10, SmartVarf P11)
+ (SmartVarf P20, SmartVarf P21)))).
-Definition ladderstep (T := _)
- := fun (dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21 : T)
- => @MxDH.ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
+Local Notation eta x := (fst x, snd x).
+
+Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
+ Eval cbv beta delta [MxDH.ladderstep_gen] in
+ @MxDH.ladderstep_gen
+ F Fadd Fsub Fmul a24
+ (F*F*F*F)
+ (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
+ (fun x f => dlet y := x in f y)
+ X1 P1 P2.
Definition uncurried_ladderstep
- := apply9_nd
- (@uncurry_unop_fe25519_64)
- ladderstep.
+ := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
+ => let a24 := fst (fst a24_x0_P1_P2) in
+ let x0 := snd (fst a24_x0_P1_P2) in
+ let '(P1, P2) := eta (snd a24_x0_P1_P2) in
+ let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
+ @ladderstep_other_assoc
+ _ add sub mul
+ a24 x0 (P10, P11) (P20, P21).
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
@@ -83,51 +115,14 @@ Local Ltac repeat_step_interpf :=
repeat (unfold interpf_step at 1; change k with k' at 1);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF25519_64.fe25519_64 -> GF25519_64.fe25519_64 -> GF25519_64.fe25519_64)
- (x' y' : GF25519_64.fe25519_64)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe25519_64 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
-Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe25519_64 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF25519_64.fe25519_64] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe25519_64 Apply' interp fst snd].
- let f := match goal with f : expr _ _ _ |- _ => f end in
- rewrite (invert_Abs_Some (e:=f) eq_refl).
- repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
- end.
- reflexivity.
-Qed.
-
-Lemma rladderstepZ_sigP : rexpr_sigP Expr9_4OpT uncurried_ladderstep rladderstepZ''.
+Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
cbv [rladderstepZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT rladderstepZ'' uncurried_ladderstep uncurry_unop_fe25519_64 SmartAbs rladderstepZ' exprArg MxDH.ladderstep_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe25519_64 ladderstep].
- intros; cbv beta zeta.
- unfold invert_Return at 14 15 16 17.
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type interp_base_type] in x.
+ destruct_head' prod.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -137,48 +132,58 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 14 15 16; unfold interpf_step.
- cbv beta iota delta [MxDH.ladderstep_other_assoc].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition rladderstepZ_sig : rexpr_9_4op_sig ladderstep.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv [ladderstep_other_assoc interp_flat_type GF25519_64.fe25519_64].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In Common.ExprBinOpT] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | _ => progress cbv [interp_base_type]
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
- eexists.
- apply rladderstepZ_sigP.
-Defined.
+ exact rladderstepZ_sigP'.
+Qed.
+Definition rladderstepZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
+
+Definition rladderstep_input_bounds
+ := (ExprUnOp_bounds, ExprUnOp_bounds,
+ ((ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW Expr9Op_bounds.
+Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition rladderstepW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- rladderstepW ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
+ rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW Expr9Op_bounds).
+Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v
index 457ffa715..610060dfc 100644
--- a/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v
+++ b/src/SpecificGen/GF25519_64Reflective/Reified/PreFreeze.v
@@ -1,6 +1,6 @@
Require Import Crypto.SpecificGen.GF25519_64Reflective.CommonUnOp.
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF25519_64.prefreeze]. reify_sig. Defined.
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. 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.
diff --git a/src/SpecificGen/GF25519_64ReflectiveAddCoordinates.v b/src/SpecificGen/GF25519_64ReflectiveAddCoordinates.v
index 2527f2c90..d4a05dec1 100644
--- a/src/SpecificGen/GF25519_64ReflectiveAddCoordinates.v
+++ b/src/SpecificGen/GF25519_64ReflectiveAddCoordinates.v
@@ -32,7 +32,7 @@ Import ListNotations.
Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Local Open Scope Z.
-Time Definition radd_coordinates : Expr9_4Op := Eval vm_compute in radd_coordinatesW.
+Time Definition radd_coordinates : Expr _ := Eval vm_compute in radd_coordinatesW.
Declare Reduction asm_interp_add_coordinates
:= cbv beta iota delta
@@ -57,7 +57,7 @@ Ltac asm_interp_add_coordinates
mapf_interp_flat_type WordW.interp_base_type word128ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
-
+(*
Time Definition interp_radd_coordinates : SpecificGen.GF25519_64BoundedCommon.fe25519_64W
-> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
-> SpecificGen.GF25519_64BoundedCommon.fe25519_64W
@@ -74,3 +74,4 @@ Time Definition interp_radd_coordinates_correct : interp_radd_coordinates = inte
Lemma radd_coordinates_correct_and_bounded : op9_4_correct_and_bounded radd_coordinates add_coordinates.
Proof. exact_no_check radd_coordinatesW_correct_and_bounded. Time Qed.
+*)
diff --git a/src/SpecificGen/GF41417_32Bounded.v b/src/SpecificGen/GF41417_32Bounded.v
index 17e8f70e7..4748ee6ff 100644
--- a/src/SpecificGen/GF41417_32Bounded.v
+++ b/src/SpecificGen/GF41417_32Bounded.v
@@ -25,13 +25,23 @@ Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.
Local Open Scope Z.
+Local Ltac cbv_tuple_map :=
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'].
+
+Local Ltac post_bounded_t :=
+ (* much pain and hackery to work around [Defined] taking forever *)
+ cbv_tuple_map;
+ let blem' := fresh "blem'" in
+ let is_bounded_lem := fresh "is_bounded_lem" in
+ intros is_bounded_lem blem';
+ apply blem'; repeat apply conj; apply is_bounded_lem.
Local Ltac bounded_t opW blem :=
- apply blem; apply is_bounded_proj1_fe41417_32.
+ generalize blem; generalize is_bounded_proj1_fe41417_32; post_bounded_t.
Local Ltac bounded_wire_digits_t opW blem :=
- apply blem; apply is_bounded_proj1_wire_digits.
+ generalize blem; generalize is_bounded_proj1_wire_digits; post_bounded_t.
Local Ltac define_binop f g opW blem :=
- refine (exist_fe41417_32W (opW (proj1_fe41417_32W f) (proj1_fe41417_32W g)) _);
+ 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)) _);
@@ -47,17 +57,17 @@ Local Ltac define_unop_WireToFE f 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_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 addW (f : fe41417_32W * fe41417_32W) : fe41417_32W := Eval simpl in interp_radd f.
+Definition subW (f : fe41417_32W * fe41417_32W) : fe41417_32W := Eval simpl in interp_rsub f.
+Definition mulW (f : fe41417_32W * fe41417_32W) : fe41417_32W := Eval simpl in interp_rmul f.
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.
@@ -86,7 +96,7 @@ Definition freezeW (f : fe41417_32W) : fe41417_32W := Eval cbv beta delta [prefr
Local Transparent Let_In.
(* Wrapper to allow extracted code to not unfold [mulW] *)
Definition mulW_noinline := mulW.
-Definition powW (f : fe41417_32W) chain := fold_chain_opt (proj1_fe41417_32W one) mulW_noinline chain [f].
+Definition powW (f : fe41417_32W) chain := fold_chain_opt (proj1_fe41417_32W one) (fun f g => mulW_noinline (f, g)) chain [f].
Definition invW (f : fe41417_32W) : fe41417_32W
:= Eval cbv -[Let_In fe41417_32W mulW_noinline] in powW f (chain inv_ec).
@@ -95,11 +105,11 @@ Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
rewrite pre_rewrite;
intros; apply rop_cb; assumption.
-Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW (Curry.curry2 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.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW (Curry.curry2 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.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW (Curry.curry2 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.
@@ -142,6 +152,7 @@ Proof.
cbv [modulusW Tuple.map].
cbv [on_tuple List.map to_list to_list' from_list from_list'
+ HList.hlistP HList.hlistP'
Tuple.map2 on_tuple2 ListUtil.map2 fe41417_32WToZ length_fe41417_32].
cbv [postfreeze GF41417_32.postfreeze].
cbv [Let_In].
@@ -203,7 +214,8 @@ Proof.
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.
+ split; [ | assumption ].
+ rewrite H0', H0; reflexivity.
Qed.
Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
@@ -212,9 +224,11 @@ Proof.
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; pose proof (fun k0 k1 X Y => proj1 (mulW_correct_and_bounded (k0, k1) (conj X Y))) as H'.
+ cbv [Curry.curry2 Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list'] in H'.
+ rewrite <- H' by assumption.
+ apply mulW_correct_and_bounded; split; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded (_, _) (conj X Y))) by assumption; reflexivity. }
{ intros [|?]; autorewrite with simpl_nth_default;
(assumption || reflexivity). }
Qed.
@@ -269,8 +283,10 @@ Proof.
unfold GF41417_32.eqb.
simpl @fe41417_32WToZ in *; cbv beta iota.
intros.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe41417_32WToZ] in *.
rewrite <- frf, <- frg by assumption.
- rewrite <- fieldwisebW_correct.
+ etransitivity; [ eapply fieldwisebW_correct | ].
+ cbv [fe41417_32WToZ].
reflexivity.
Defined.
@@ -297,7 +313,7 @@ Proof.
lazymatch (eval cbv delta [GF41417_32.sqrt] in GF41417_32.sqrt) with
| (fun powf powf_squared f => dlet a := powf in _)
=> exact (dlet powx := powW (fe41417_32ZToW x) (chain GF41417_32.sqrt_ec) in
- GF41417_32.sqrt (fe41417_32WToZ powx) (fe41417_32WToZ (mulW_noinline powx powx)) x)
+ GF41417_32.sqrt (fe41417_32WToZ powx) (fe41417_32WToZ (mulW_noinline (powx, powx))) x)
| (fun f => pow f _)
=> exact (GF41417_32.sqrt x)
end.
@@ -324,21 +340,41 @@ Proof.
=> 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
- | mulW_noinline _ _ => apply mulW_correct_and_bounded
- | powW _ _ => apply powW_correct_and_bounded
- | sqrt_m1W => vm_compute; reflexivity
- | _ => assumption
- end
- end;
- subst_evars; reflexivity
+ pose proof (fun X Y => proj1 (mulW_correct_and_bounded (sqrt_m1W, a) (conj X Y))) as correctness;
+ let cbv_in_all _ := (cbv [fe41417_32WToZ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe41417_32WToZ Curry.curry2 HList.hlistP HList.hlistP'] in *; idtac) in
+ cbv_in_all ();
+ let solver _ := (repeat match goal with
+ | _ => progress subst
+ | _ => progress unfold fst, snd
+ | _ => progress cbv_in_all ()
+ | [ |- ?x /\ ?x ] => cut x; [ intro; split; assumption | ]
+ | [ |- is_bounded ?op = true ]
+ => let H := fresh in
+ lazymatch op with
+ | context[mulW (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[mulW_noinline (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[powW _ _] => pose proof powW_correct_and_bounded as H
+ | context[sqrt_m1W] => vm_compute; reflexivity
+ | _ => assumption
+ end;
+ cbv_in_all ();
+ apply H
+ end) in
+ rewrite <- correctness by solver (); clear correctness;
+ let lem := fresh in
+ pose proof eqbW_correct as lem; cbv_in_all (); rewrite <- lem by solver (); clear lem;
+ pose proof (pull_bool_if fe41417_32WToZ) as lem; cbv_in_all (); rewrite lem by solver (); clear lem;
+ subst_evars; reflexivity
end.
} Unfocus.
+ assert (Hfold : forall x, fe41417_32WToZ x = fe41417_32WToZ x) by reflexivity.
+ unfold fe41417_32WToZ at 2 in Hfold.
+ etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody; [ reflexivity | intro ].
+ apply Hfold.
+ } Unfocus.
+ clear Hfold.
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
@@ -346,15 +382,22 @@ Proof.
[ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
| _ => idtac
end;
- reflexivity. }
- { cbv [Let_In];
+ reflexivity.
+ }
+
+ { cbv [Let_In HList.hlistP HList.hlistP'];
try break_if;
repeat lazymatch goal with
| [ |- is_bounded (?WToZ (powW _ _)) = true ]
=> apply powW_correct_and_bounded; assumption
- | [ |- is_bounded (?WToZ (mulW _ _)) = true ]
- => apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ]
- end. }
+ | [ |- is_bounded (snd (?WToZ (_, powW _ _))) = true ]
+ => generalize powW_correct_and_bounded;
+ cbv [snd Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'];
+ let H := fresh in intro H; apply H; assumption
+ | [ |- is_bounded (?WToZ (mulW (_, _))) = true ]
+ => apply mulW_correct_and_bounded; split; [ vm_compute; reflexivity | ]
+ end.
+ }
Defined.
Definition sqrtW (f : fe41417_32W) : fe41417_32W :=
@@ -394,7 +437,7 @@ 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.
+Proof. define_unop f invW (fun x p => proj2 (invW_correct_and_bounded x p)). Defined.
Definition sqrt (f : fe41417_32) : fe41417_32.
Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
@@ -407,7 +450,12 @@ Local Ltac op_correct_t op opW_correct_and_bounded :=
=> rewrite proj1_wire_digits_exist_wire_digitsW
| _ => idtac
end;
- apply opW_correct_and_bounded;
+ generalize opW_correct_and_bounded;
+ cbv_tuple_map;
+ cbv [fst snd];
+ let H := fresh in
+ intro H; apply H;
+ repeat match goal with |- and _ _ => apply conj end;
lazymatch goal with
| [ |- is_bounded _ = true ]
=> apply is_bounded_proj1_fe41417_32
@@ -434,7 +482,7 @@ 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.
+Proof. op_correct_t inv (fun x p => proj1 (invW_correct_and_bounded x p)). 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.
diff --git a/src/SpecificGen/GF41417_32BoundedAddCoordinates.v b/src/SpecificGen/GF41417_32BoundedAddCoordinates.v
index e7d5842ce..b4f6ba5a8 100644
--- a/src/SpecificGen/GF41417_32BoundedAddCoordinates.v
+++ b/src/SpecificGen/GF41417_32BoundedAddCoordinates.v
@@ -14,7 +14,7 @@ Local Ltac define_binop f g 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_coordinates / _ _ _ _ _ _ _ _ _.
+(*Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
Definition add_coordinatesW (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe41417_32W) : Tuple.tuple fe41417_32W 4
:= Eval simpl in interp_radd_coordinates x0 x1 x2 x3 x4 x5 x6 x7 x8.
@@ -75,3 +75,4 @@ Lemma add_coordinates_correct (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe41417_32)
(proj1_fe41417_32 x7)
(proj1_fe41417_32 x8).
Proof. op_correct_t add_coordinates add_coordinatesW_correct_and_bounded. Qed.
+*)
diff --git a/src/SpecificGen/GF41417_32BoundedCommon.v b/src/SpecificGen/GF41417_32BoundedCommon.v
index 2bf69d425..6705d924b 100644
--- a/src/SpecificGen/GF41417_32BoundedCommon.v
+++ b/src/SpecificGen/GF41417_32BoundedCommon.v
@@ -322,14 +322,14 @@ Ltac unfold_is_bounded_in' H :=
end.
Ltac preunfold_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.
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe41417_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
Ltac unfold_is_bounded_in H :=
preunfold_is_bounded_in H;
unfold_is_bounded_in' H.
Ltac preunfold_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].
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe41417_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
Ltac unfold_is_bounded :=
preunfold_is_bounded;
@@ -724,7 +724,7 @@ Notation in_op_correct_and_bounded k irop op
/\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe41417_32WToZ irop))%type)
(only parsing).
-Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
+(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
: forall (irop : Tower.tower_nd fe41417_32W (Tuple.tuple fe41417_32W count_out) count_in)
(op : Tower.tower_nd GF41417_32.fe41417_32 (Tuple.tuple GF41417_32.fe41417_32 count_out) count_in)
(cont : Prop -> Prop),
@@ -792,18 +792,14 @@ Qed.
Definition inm_op_correct_and_bounded1 count_in irop op
:= Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.
-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).
+ inm_op_correct_and_bounded count_in 1 irop op.*)
+Notation inm_op_correct_and_bounded n m irop op
+ := ((forall x,
+ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe41417_32WToZ x)
+ -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe41417_32WToZ x))))
+ (only parsing).
+Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
+Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe41417_32WToZ x) = true
@@ -818,20 +814,6 @@ Notation iunop_WireToFE_correct_and_bounded irop op
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).
-Notation i9top_correct_and_bounded k irop op
- := ((forall x0 x1 x2 x3 x4 x5 x6 x7 x8,
- is_bounded (fe41417_32WToZ x0) = true
- -> is_bounded (fe41417_32WToZ x1) = true
- -> is_bounded (fe41417_32WToZ x2) = true
- -> is_bounded (fe41417_32WToZ x3) = true
- -> is_bounded (fe41417_32WToZ x4) = true
- -> is_bounded (fe41417_32WToZ x5) = true
- -> is_bounded (fe41417_32WToZ x6) = true
- -> is_bounded (fe41417_32WToZ x7) = true
- -> is_bounded (fe41417_32WToZ x8) = true
- -> (Tuple.map (n:=k) fe41417_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)
- = op (fe41417_32WToZ x0) (fe41417_32WToZ x1) (fe41417_32WToZ x2) (fe41417_32WToZ x3) (fe41417_32WToZ x4) (fe41417_32WToZ x5) (fe41417_32WToZ x6) (fe41417_32WToZ x7) (fe41417_32WToZ x8))
- * HList.hlist (fun v => is_bounded v = true) (Tuple.map (n:=k) fe41417_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)))%type)
- (only parsing).
+Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-Definition prefreeze := GF41417_32.prefreeze.
+Notation prefreeze := GF41417_32.prefreeze.
diff --git a/src/SpecificGen/GF41417_32BoundedExtendedAddCoordinates.v b/src/SpecificGen/GF41417_32BoundedExtendedAddCoordinates.v
index dc99e0b11..94bc097cf 100644
--- a/src/SpecificGen/GF41417_32BoundedExtendedAddCoordinates.v
+++ b/src/SpecificGen/GF41417_32BoundedExtendedAddCoordinates.v
@@ -5,7 +5,7 @@ Require Import Crypto.SpecificGen.GF41417_32ExtendedAddCoordinates.
Require Import Crypto.SpecificGen.GF41417_32BoundedAddCoordinates.
Require Import Crypto.Util.Tuple.
Require Import Crypto.Util.Tactics.
-
+(*
Lemma fieldwise_eq_extended_add_coordinates_full' twice_d P10 P11 P12 P13 P20 P21 P22 P23
: Tuple.fieldwise
(n:=4) GF41417_32BoundedCommon.eq
@@ -65,3 +65,4 @@ Proof.
destruct_head' prod.
rewrite <- fieldwise_eq_extended_add_coordinates_full'; reflexivity.
Qed.
+*)
diff --git a/src/SpecificGen/GF41417_32Reflective.v b/src/SpecificGen/GF41417_32Reflective.v
index 4b9b012b1..404111062 100644
--- a/src/SpecificGen/GF41417_32Reflective.v
+++ b/src/SpecificGen/GF41417_32Reflective.v
@@ -45,6 +45,7 @@ Declare Reduction asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe41417_32W curry_unop_fe41417_32W curry_unop_wire_digitsW curry_9op_fe41417_32W
WordW.interp_op WordW.interp_base_type
@@ -56,6 +57,7 @@ Ltac asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe41417_32W curry_unop_fe41417_32W curry_unop_wire_digitsW curry_9op_fe41417_32W
WordW.interp_op WordW.interp_base_type
@@ -65,15 +67,15 @@ Ltac asm_interp
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
+Definition interp_radd : SpecificGen.GF41417_32BoundedCommon.fe41417_32W * SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rsub : SpecificGen.GF41417_32BoundedCommon.fe41417_32W * SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rmul : SpecificGen.GF41417_32BoundedCommon.fe41417_32W * SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
:= Eval asm_interp in interp_bexpr rmul.
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
diff --git a/src/SpecificGen/GF41417_32Reflective/Common.v b/src/SpecificGen/GF41417_32Reflective/Common.v
index d537c6018..75ce71bca 100644
--- a/src/SpecificGen/GF41417_32Reflective/Common.v
+++ b/src/SpecificGen/GF41417_32Reflective/Common.v
@@ -7,7 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Wf.
Require Import Crypto.Reflection.ExprInversion.
+Require Import Crypto.Reflection.Tuple.
Require Import Crypto.Reflection.Relations.
+Require Import Crypto.Reflection.Eta.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -15,13 +17,14 @@ Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Z.Syntax.Equality.
Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.Curry.
Require Import Crypto.Util.Tower.
Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Notations.
Notation Expr := (Expr base_type op).
@@ -43,30 +46,24 @@ Defined.
Definition Expr_n_OpT (count_out : nat) : flat_type base_type
:= Eval cbv [Syntax.tuple Syntax.tuple' fe41417_32T] in
Syntax.tuple fe41417_32T count_out.
-Fixpoint Expr_nm_OpT (count_in count_out : nat) : type base_type
- := match count_in with
- | 0 => Expr_n_OpT count_out
- | S n => SmartArrow fe41417_32T (Expr_nm_OpT n count_out)
- end.
+Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
+ := Eval cbv [Syntax.tuple Syntax.tuple' fe41417_32T Expr_n_OpT] in
+ Arrow (Syntax.tuple fe41417_32T count_in) (Expr_n_OpT count_out).
Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from (uncurry_unop_fe41417_32 ge_modulus). Defined.
+Proof. make_type_from ge_modulus. Defined.
Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Proof. make_type_from unpack. Defined.
Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from (uncurry_unop_fe41417_32 pack). Defined.
+Proof. make_type_from pack. Defined.
Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
Definition ExprArgT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpT.
+ := Eval compute in domain ExprUnOpT.
Definition ExprArgWireT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpFEToWireT.
-Definition ExprArgRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpT.
-Definition ExprArgWireRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Tbase TZ).
+ := Eval compute in domain ExprUnOpWireToFET.
+Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
@@ -75,99 +72,67 @@ Definition ExprBinOp : Type := Expr ExprBinOpT.
Definition ExprUnOp : Type := Expr ExprUnOpT.
Definition Expr4Op : Type := Expr Expr4OpT.
Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr ExprArgT.
-Definition ExprArgWire : Type := Expr ExprArgWireT.
-Definition ExprArgRev : Type := Expr ExprArgRevT.
-Definition ExprArgWireRev : Type := Expr ExprArgWireRevT.
+Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
+Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
Definition expr_nm_Op count_in count_out var : Type
:= expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Tbase TZ).
+Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) ExprArgT.
-Definition exprArgWire var : Type := expr base_type op (var:=var) ExprArgWireT.
-Definition exprArgRev var : Type := expr base_type op (var:=var) ExprArgRevT.
-Definition exprArgWireRev var : Type := expr base_type op (var:=var) ExprArgWireRevT.
-
-Local Ltac bounds_from_list_cps ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(fun T (extra : T) => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, extra))
- | (?x :: ?xs)%list => let bs := bounds_from_list_cps xs in
- constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
- end.
-
-Local Ltac make_bounds_cps ls extra :=
- let v := bounds_from_list_cps (List.rev ls) in
- let v := (eval compute in v) in
- exact (v _ extra).
-
-Local Ltac bounds_from_list ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
- | (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| Bounds.lower := fst x ; Bounds.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.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out : interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type.
-Proof.
- refine match count_in return interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type with
- | 0 => tt
- | S n => let v := interp_all_binders_for_to' (Expr_nm_Op_bounds n count_out) in
- interp_all_binders_for_of' _
- end; simpl.
- make_bounds_cps (Tuple.to_list _ bounds) v.
-Defined.
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
+Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
+
+Definition make_bound (x : Z * Z) : ZBounds.t
+ := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
+
+Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
+ := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
+ | 0 => tt
+ | S n
+ => let b := (Tuple.map make_bound bounds) in
+ let bs := Expr_nm_Op_bounds n count_out in
+ match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
+ | 0 => fun _ => b
+ | S n' => fun bs => (bs, b)
+ end bs
+ end.
+Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
:= Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
:= Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
:= Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
+Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
+ := Tuple.map make_bound wire_digit_bounds.
-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 (@WordW.interp_op) e).
+Definition interp_bexpr : ExprBinOp -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W * SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- := fun e => curry_unop_fe41417_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.word64
- := fun e => curry_unop_fe41417_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W -> SpecificGen.GF41417_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe41417_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF41417_32BoundedCommon.wire_digitsW -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_9_4expr : Expr9_4Op
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
- -> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
+ -> Tuple.tuple SpecificGen.GF41417_32BoundedCommon.fe41417_32W 9
-> Tuple.tuple SpecificGen.GF41417_32BoundedCommon.fe41417_32W 4
- := fun e => curry_9op_fe41417_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+ := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 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
@@ -181,40 +146,39 @@ Notation op9_4_correct_and_bounded rop op
Ltac rexpr_cbv :=
lazymatch goal with
- | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
=> let operf := head oper in
let uncurryf := head uncurry in
try cbv delta [T]; try cbv delta [oper];
try cbv beta iota delta [uncurryf]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
+ => let operf := head oper in
+ try cbv delta [T]; try cbv delta [oper]
end;
- cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+ cbv beta iota delta [interp_flat_type interp_base_type zero_ GF41417_32.fe41417_32 GF41417_32.wire_digits].
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 }
+ | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
(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 rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT (uncurry_9op_fe41417_32 op)) (only parsing).
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
+Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT 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 := ropW' in
- let ropZ := ropW' in
- let ropBounds := ropW' in
- let ropBoundedWordW := ropW' in
- ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
+ ropW' = proj1_sig ropZ_sig
+ /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
(only parsing).
Ltac app_tuples x y :=
@@ -227,7 +191,7 @@ Ltac app_tuples x y :=
Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
-
+(*
Fixpoint args_to_bounded_helperT {n}
(v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
@@ -299,14 +263,15 @@ Proof.
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
@@ -327,6 +292,7 @@ Ltac assoc_right_tuple x so_far :=
end
end.
+(*
Local Ltac make_args x :=
let x' := fresh "x'" in
compute in x |- *;
@@ -338,11 +304,6 @@ Local Ltac make_args x :=
let xv := assoc_right_tuple x'' (@None) in
refine (SmartVarf (xv : interp_flat_type _ t')).
-Definition unop_make_args {var} (x : exprArg var) : exprArgRev var.
-Proof. make_args x. Defined.
-Definition unop_wire_make_args {var} (x : exprArgWire var) : exprArgWireRev var.
-Proof. make_args x. Defined.
-
Local Ltac args_to_bounded x H :=
let x' := fresh in
set (x' := x);
@@ -351,29 +312,138 @@ Local Ltac args_to_bounded x H :=
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
+ let c := (eval cbv [domain 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) ]
).
+ *)
+
+Section gen.
+ Definition bounds_are_good_gen
+ {n : nat} (bounds : Tuple.tuple (Z * Z) n)
+ := let res :=
+ Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
+ in
+ List.fold_right andb true (Tuple.to_list n res).
+ Definition unop_args_to_bounded'
+ (bs : Z * Z)
+ (Hbs : bounds_are_good_gen (n:=1) bs = true)
+ (x : word64)
+ (H : is_bounded_gen (Tuple.map (fun v : word64 => (v : Z)) (n:=1) x) bs = true)
+ : BoundedWordW.BoundedWord.
+ Proof.
+ refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
+ unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
+ abstract (
+ destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
+ repeat apply conj; assumption
+ ).
+ Defined.
+ Fixpoint n_op_args_to_bounded'
+ n
+ : forall (bs : Tuple.tuple' (Z * Z) n)
+ (Hbs : bounds_are_good_gen (n:=S n) bs = true)
+ (x : Tuple.tuple' word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
+ Proof.
+ destruct n as [|n']; simpl in *.
+ { exact unop_args_to_bounded'. }
+ { refine (fun bs Hbs x H
+ => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
+ @unop_args_to_bounded' (snd bs) _ (snd x) _));
+ clear n_op_args_to_bounded';
+ simpl in *;
+ [ clear x H | clear Hbs | clear x H | clear Hbs ];
+ unfold bounds_are_good_gen, is_bounded_gen in *;
+ abstract (
+ repeat first [ progress simpl in *
+ | assumption
+ | reflexivity
+ | progress Bool.split_andb
+ | progress destruct_head prod
+ | match goal with
+ | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
+ end
+ | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
+ | progress break_match_hyps
+ | rewrite Bool.andb_true_iff; apply conj
+ | unfold Tuple.map, Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
+ ). }
+ Defined.
+
+ Definition n_op_args_to_bounded
+ n
+ : forall (bs : Tuple.tuple (Z * Z) n)
+ (Hbs : bounds_are_good_gen bs = true)
+ (x : Tuple.tuple word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
+ := match n with
+ | 0 => fun _ _ _ _ => tt
+ | S n' => @n_op_args_to_bounded' n'
+ end.
-Definition unop_args_to_bounded (x : fe41417_32W) (H : is_bounded (fe41417_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
+ Fixpoint nm_op_args_to_bounded' n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple' (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => @n_op_args_to_bounded m bs Hbs
+ | S n' => fun x H
+ => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
+ @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
+ end.
+ Definition nm_op_args_to_bounded n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => fun _ _ => tt
+ | S n' => @nm_op_args_to_bounded' n' m bs Hbs
+ end.
+End gen.
+
+Local Ltac get_inner_len T :=
+ lazymatch T with
+ | (?T * _)%type => get_inner_len T
+ | ?T => get_len T
+ end.
+Local Ltac get_outer_len T :=
+ lazymatch T with
+ | (?A * ?B)%type => let a := get_outer_len A in
+ let b := get_outer_len B in
+ (eval compute in (a + b)%nat)
+ | ?T => constr:(1%nat)
+ end.
+Local Ltac args_to_bounded x H :=
+ let t := type of x in
+ let m := get_inner_len t in
+ let n := get_outer_len t in
+ let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
+ let H := (eval cbv beta in (H eq_refl)) in
+ exact H.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe41417_32W * fe41417_32W)
(H : is_bounded (fe41417_32WToZ (fst x)) = true)
(H' : is_bounded (fe41417_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
-Proof.
- let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
- exact v.
-Defined.
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
+Proof. args_to_bounded x (conj H H'). Defined.
+Definition unop_args_to_bounded (x : fe41417_32W) (H : is_bounded (fe41417_32WToZ x) = true)
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain 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 _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
+Proof. args_to_bounded x H. Defined.
Definition op9_args_to_bounded (x : fe41417_32W * fe41417_32W * fe41417_32W * fe41417_32W * fe41417_32W * fe41417_32W * fe41417_32W * fe41417_32W * fe41417_32W)
(H0 : is_bounded (fe41417_32WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
(H1 : is_bounded (fe41417_32WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
@@ -384,58 +454,39 @@ Definition op9_args_to_bounded (x : fe41417_32W * fe41417_32W * fe41417_32W * fe
(H6 : is_bounded (fe41417_32WToZ (snd (fst (fst x)))) = true)
(H7 : is_bounded (fe41417_32WToZ (snd (fst x))) = true)
(H8 : is_bounded (fe41417_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for Expr9_4OpT).
-Proof.
- let v := constr:(unop_args_to_bounded _ H8) in
- let v := app_tuples (unop_args_to_bounded _ H7) v in
- let v := app_tuples (unop_args_to_bounded _ H6) v in
- let v := app_tuples (unop_args_to_bounded _ H5) v in
- let v := app_tuples (unop_args_to_bounded _ H4) v in
- let v := app_tuples (unop_args_to_bounded _ H3) v in
- let v := app_tuples (unop_args_to_bounded _ H2) v in
- let v := app_tuples (unop_args_to_bounded _ H1) v in
- let v := app_tuples (unop_args_to_bounded _ H0) v in
- exact v.
-Defined.
-
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
+Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
+Local Ltac make_bounds_prop' bounds bounds' :=
+ first [ refine (andb _ _);
+ [ destruct bounds' as [bounds' _], bounds as [bounds _]
+ | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
+ try make_bounds_prop' bounds bounds'
+ | 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) ].
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.
+ make_bounds_prop' bounds bounds'.
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
Proof.
refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders Expr9_4OpT)) : bool.
+Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-
-Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
+(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe41417_32 f k)).
@@ -460,12 +511,11 @@ Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ v
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe41417_32 f k)).
-
+*)
(* 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
@@ -473,23 +523,25 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
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 interp_9_4expr
+ interp_flat_type_eta interp_flat_type_eta_gen
curry_binop_fe41417_32W curry_unop_fe41417_32W curry_unop_wire_digitsW curry_9op_fe41417_32W
curry_binop_fe41417_32 curry_unop_fe41417_32 curry_unop_wire_digits curry_9op_fe41417_32
uncurry_binop_fe41417_32W uncurry_unop_fe41417_32W uncurry_unop_wire_digitsW uncurry_9op_fe41417_32W
uncurry_binop_fe41417_32 uncurry_unop_fe41417_32 uncurry_unop_wire_digits uncurry_9op_fe41417_32
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT
- interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] 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 *;
+ change word64 with WordW.wordW in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
+ | progress unfold RelationClasses.Reflexive
| reflexivity
| assumption
| progress destruct_head' base_type
@@ -509,23 +561,33 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
+ Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
+ simpl @interp_flat_type in *;
+ (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
+ change (WordW.interp_base_type TZ) with v in *);
+ cbv beta iota zeta in *;
lazymatch goal with
| [ |- fe41417_32WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
| [ |- wire_digitsWToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
+ | [ |- (Tuple.map fe41417_32WToZ ?x = _) /\ _ ]
+ => generalize dependent x; intros
| [ |- ((Tuple.map fe41417_32WToZ ?x = _) * _)%type ]
=> generalize dependent x; intros
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type remove_all_binders] in *;
+ cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
destruct_head' prod;
change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
+ make_bound
+ Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
destruct_head' ZBounds.bounds;
@@ -534,12 +596,12 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
destruct_head' and;
Z.ltb_to_lt;
change WordW.wordWToZ with word64ToZ in *;
- cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd];
+ cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe41417_32WToZ HList.hlistP HList.hlistP'];
+ cbv [WordW.bit_width BitSize64.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
repeat split; unfold_is_bounded;
Z.ltb_to_lt;
try omega; try reflexivity.
-
Ltac rexpr_correct :=
let ropW' := fresh in
let ropZ_sig := fresh in
@@ -555,9 +617,13 @@ Ltac rexpr_correct :=
Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) opW) bounds)
+ := (Interp (@ZBounds.interp_op) opW bounds)
(only parsing).
+Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
+
+Ltac prove_rexpr_wfT
+ := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
Module Export PrettyPrinting.
(* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
@@ -569,23 +635,21 @@ Module Export PrettyPrinting.
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 {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end
+ := fun (t : base_type) (x : ZBounds.interp_base_type t)
+ => match x with
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
+ => in_range l u
+ | None
+ => overflow
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
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
+ := match t return interp_flat_type _ t -> result with
+ | Tbase _ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
| Unit => fun _ => no
| Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
| no, no => no
diff --git a/src/SpecificGen/GF41417_32Reflective/Common9_4Op.v b/src/SpecificGen/GF41417_32Reflective/Common9_4Op.v
index 4590d5871..0e7336d2e 100644
--- a/src/SpecificGen/GF41417_32Reflective/Common9_4Op.v
+++ b/src/SpecificGen/GF41417_32Reflective/Common9_4Op.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -42,8 +41,8 @@ Lemma Expr9_4Op_correct_and_bounded
let (Hx7, Hx8) := (eta_and Hx78) in
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -80,29 +79,24 @@ Lemma Expr9_4Op_correct_and_bounded
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => op9_4_bounds_good bounds = true
| None => False
end)
: op9_4_correct_and_bounded ropW op.
Proof.
- intros x0 x1 x2 x3 x4 x5 x6 x7 x8.
- intros Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8.
- pose x0 as x0'.
- pose x1 as x1'.
- pose x2 as x2'.
- pose x3 as x3'.
- pose x4 as x4'.
- pose x5 as x5'.
- pose x6 as x6'.
- pose x7 as x7'.
- pose x8 as x8'.
- hnf in x0, x1, x2, x3, x4, x5, x6, x7, x8; destruct_head' prod.
- specialize (H0 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ intros xs Hxs.
+ pose xs as xs'.
+ compute in xs.
+ destruct_head' prod.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
+ pose Hxs as Hxs'.
+ destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
+ specialize (H0 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ specialize (H1 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded (x0', x1', x2', x3', x4', x5', x6', x7', x8') Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
+ Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.64s) (successful) *)*)
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v b/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
index d55d12f1d..9cf8f3bf6 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonBinOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -18,8 +17,8 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -33,18 +32,19 @@ Lemma ExprBinOp_correct_and_bounded
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
end)
: binop_correct_and_bounded 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
+ intros xy HxHy.
+ pose xy as xy'.
+ compute in xy; destruct_head' prod.
+ specialize (H0 xy' HxHy).
+ specialize (H1 xy' HxHy).
+ destruct HxHy as [Hx Hy].
+ let args := constr:(binop_args_to_bounded xy' 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
index 92a928bfc..5c369d730 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOp_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
index 4c8923dd9..312d4439e 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToWire.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
index 748e2fd10..76fcf7e88 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpFEToZ.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe41417_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
index d70798964..553a9583d 100644
--- a/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF41417_32Reflective/CommonUnOpWireToFE.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF41417_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
let args := unopWireToFE_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/AddCoordinates.v b/src/SpecificGen/GF41417_32Reflective/Reified/AddCoordinates.v
index 0a8480432..362d69158 100644
--- a/src/SpecificGen/GF41417_32Reflective/Reified/AddCoordinates.v
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/AddCoordinates.v
@@ -7,7 +7,6 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.ExprInversion.
Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
@@ -17,7 +16,10 @@ Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
+Require Import Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.Mul.
@@ -33,24 +35,28 @@ Require Import Bedrock.Word.
Definition radd_coordinatesZ' var twice_d P1 P2
:= @Extended.add_coordinates_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
twice_d _
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
P1 P2.
+Local Notation eta x := (fst x, snd x).
+
Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => radd_coordinatesZ'
- var (Return twice_d)
- (Return P10, Return P11, Return P12, Return P13)
- (Return P20, Return P21, Return P22, Return P23))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
+ => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
+ := twice_d_P1_P2 in
+ radd_coordinatesZ'
+ var (SmartVarf twice_d)
+ (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
+ (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
Definition add_coordinates
:= fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
@@ -59,70 +65,60 @@ Definition add_coordinates
twice_d (P10, P11, P12, P13) (P20, P21, P22, P23).
Definition uncurried_add_coordinates
- := apply9_nd
- (@uncurry_unop_fe41417_32)
- add_coordinates.
+ := fun twice_d_P1_P2
+ => let twice_d := fst twice_d_P1_P2 in
+ let (P1, P2) := eta (snd twice_d_P1_P2) in
+ @Extended.add_coordinates
+ _ add sub mul
+ twice_d P1 P2.
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
let H := fresh in
pose k as k';
assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
+ repeat (unfold k'; change k with k'; unfold interpf_step);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF41417_32.fe41417_32 -> GF41417_32.fe41417_32 -> GF41417_32.fe41417_32)
- (x' y' : GF41417_32.fe41417_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe41417_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
+Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe41417_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF41417_32.fe41417_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe41417_32 Apply' interp fst snd].
- rewrite (invert_Abs_Some (e:=f) eq_refl).
+ cbv [radd_coordinatesZ''].
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type] in x.
+ destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
+ | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
end.
- reflexivity.
-Qed.
-
-Lemma radd_coordinatesZ_sigP : rexpr_sigP Expr9_4OpT uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT radd_coordinatesZ'' uncurried_add_coordinates uncurry_unop_fe41417_32 SmartAbs radd_coordinatesZ' exprArg Extended.add_coordinates_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe41417_32 add_coordinates].
- intros.
- unfold invert_Return at 13 14 15 16.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [radd_coordinatesZ' add_coordinates Extended.add_coordinates_gen uncurried_add_coordinates SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -132,48 +128,56 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 13 14 15; unfold interpf_step.
- cbv beta iota delta [Extended.add_coordinates].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition radd_coordinatesZ_sig : rexpr_9_4op_sig add_coordinates.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv beta iota delta [Extended.add_coordinates interp_flat_type interp_base_type GF41417_32.fe41417_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- eexists.
- apply radd_coordinatesZ_sigP.
-Defined.
+ exact radd_coordinatesZ_sigP'.
+Qed.
+Definition radd_coordinatesZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
+
+Definition radd_coordinates_input_bounds
+ := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW Expr9Op_bounds.
+Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition radd_coordinatesW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- radd_coordinatesW add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
+ radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW Expr9Op_bounds).
+Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/LadderStep.v b/src/SpecificGen/GF41417_32Reflective/Reified/LadderStep.v
index 0d1cd59cb..80e929103 100644
--- a/src/SpecificGen/GF41417_32Reflective/Reified/LadderStep.v
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/LadderStep.v
@@ -7,8 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Relations.
Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -18,6 +19,7 @@ Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
Require Import Crypto.Spec.MxDH.
+Require Import Crypto.SpecificGen.GF41417_32Reflective.Common.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF41417_32Reflective.Reified.Mul.
@@ -30,50 +32,80 @@ Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
Require Import Bedrock.Word.
-(* XXX FIXME: Remove dummy code *)
-Definition rladderstepZ' var (T:=_) (dummy0 dummy1 dummy2 a24 x0 : T) P1 P2
+Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
:= @MxDH.ladderstep_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
a24
_
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
x0
P1 P2.
Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21
- => rladderstepZ'
- var (Return dummy0) (Return dummy1) (Return dummy2)
- (Return a24) (Return x0)
- (Return P10, Return P11)
- (Return P20, Return P21))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
+ => let '(a24, x0, ((P10, P11), (P20, P21)))
+ := a24_x0_P1_P2 in
+ rladderstepZ'
+ var (SmartVarf a24) (SmartVarf x0)
+ (SmartVarf P10, SmartVarf P11)
+ (SmartVarf P20, SmartVarf P21)))).
-Definition ladderstep (T := _)
- := fun (dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21 : T)
- => @MxDH.ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
+Local Notation eta x := (fst x, snd x).
+
+Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
+ Eval cbv beta delta [MxDH.ladderstep_gen] in
+ @MxDH.ladderstep_gen
+ F Fadd Fsub Fmul a24
+ (F*F*F*F)
+ (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
+ (fun x f => dlet y := x in f y)
+ X1 P1 P2.
Definition uncurried_ladderstep
- := apply9_nd
- (@uncurry_unop_fe41417_32)
- ladderstep.
+ := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
+ => let a24 := fst (fst a24_x0_P1_P2) in
+ let x0 := snd (fst a24_x0_P1_P2) in
+ let '(P1, P2) := eta (snd a24_x0_P1_P2) in
+ let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
+ @ladderstep_other_assoc
+ _ add sub mul
+ a24 x0 (P10, P11) (P20, P21).
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
@@ -83,51 +115,14 @@ Local Ltac repeat_step_interpf :=
repeat (unfold interpf_step at 1; change k with k' at 1);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF41417_32.fe41417_32 -> GF41417_32.fe41417_32 -> GF41417_32.fe41417_32)
- (x' y' : GF41417_32.fe41417_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe41417_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
-Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe41417_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF41417_32.fe41417_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe41417_32 Apply' interp fst snd].
- let f := match goal with f : expr _ _ _ |- _ => f end in
- rewrite (invert_Abs_Some (e:=f) eq_refl).
- repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
- end.
- reflexivity.
-Qed.
-
-Lemma rladderstepZ_sigP : rexpr_sigP Expr9_4OpT uncurried_ladderstep rladderstepZ''.
+Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
cbv [rladderstepZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT rladderstepZ'' uncurried_ladderstep uncurry_unop_fe41417_32 SmartAbs rladderstepZ' exprArg MxDH.ladderstep_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe41417_32 ladderstep].
- intros; cbv beta zeta.
- unfold invert_Return at 14 15 16 17.
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type interp_base_type] in x.
+ destruct_head' prod.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -137,48 +132,58 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 14 15 16; unfold interpf_step.
- cbv beta iota delta [MxDH.ladderstep_other_assoc].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition rladderstepZ_sig : rexpr_9_4op_sig ladderstep.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv [ladderstep_other_assoc interp_flat_type GF41417_32.fe41417_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In Common.ExprBinOpT] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | _ => progress cbv [interp_base_type]
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
- eexists.
- apply rladderstepZ_sigP.
-Defined.
+ exact rladderstepZ_sigP'.
+Qed.
+Definition rladderstepZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
+
+Definition rladderstep_input_bounds
+ := (ExprUnOp_bounds, ExprUnOp_bounds,
+ ((ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW Expr9Op_bounds.
+Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition rladderstepW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- rladderstepW ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
+ rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW Expr9Op_bounds).
+Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v
index 889c6ca49..ede666d32 100644
--- a/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v
+++ b/src/SpecificGen/GF41417_32Reflective/Reified/PreFreeze.v
@@ -1,6 +1,6 @@
Require Import Crypto.SpecificGen.GF41417_32Reflective.CommonUnOp.
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF41417_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. 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.
diff --git a/src/SpecificGen/GF41417_32ReflectiveAddCoordinates.v b/src/SpecificGen/GF41417_32ReflectiveAddCoordinates.v
index c08f28dfe..54c260408 100644
--- a/src/SpecificGen/GF41417_32ReflectiveAddCoordinates.v
+++ b/src/SpecificGen/GF41417_32ReflectiveAddCoordinates.v
@@ -32,7 +32,7 @@ Import ListNotations.
Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Local Open Scope Z.
-Time Definition radd_coordinates : Expr9_4Op := Eval vm_compute in radd_coordinatesW.
+Time Definition radd_coordinates : Expr _ := Eval vm_compute in radd_coordinatesW.
Declare Reduction asm_interp_add_coordinates
:= cbv beta iota delta
@@ -57,7 +57,7 @@ Ltac asm_interp_add_coordinates
mapf_interp_flat_type WordW.interp_base_type word64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
-
+(*
Time Definition interp_radd_coordinates : SpecificGen.GF41417_32BoundedCommon.fe41417_32W
-> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
-> SpecificGen.GF41417_32BoundedCommon.fe41417_32W
@@ -74,3 +74,4 @@ Time Definition interp_radd_coordinates_correct : interp_radd_coordinates = inte
Lemma radd_coordinates_correct_and_bounded : op9_4_correct_and_bounded radd_coordinates add_coordinates.
Proof. exact_no_check radd_coordinatesW_correct_and_bounded. Time Qed.
+*)
diff --git a/src/SpecificGen/GF5211_32Bounded.v b/src/SpecificGen/GF5211_32Bounded.v
index 5bd8b85c2..02ea3614f 100644
--- a/src/SpecificGen/GF5211_32Bounded.v
+++ b/src/SpecificGen/GF5211_32Bounded.v
@@ -25,13 +25,23 @@ Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.
Local Open Scope Z.
+Local Ltac cbv_tuple_map :=
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'].
+
+Local Ltac post_bounded_t :=
+ (* much pain and hackery to work around [Defined] taking forever *)
+ cbv_tuple_map;
+ let blem' := fresh "blem'" in
+ let is_bounded_lem := fresh "is_bounded_lem" in
+ intros is_bounded_lem blem';
+ apply blem'; repeat apply conj; apply is_bounded_lem.
Local Ltac bounded_t opW blem :=
- apply blem; apply is_bounded_proj1_fe5211_32.
+ generalize blem; generalize is_bounded_proj1_fe5211_32; post_bounded_t.
Local Ltac bounded_wire_digits_t opW blem :=
- apply blem; apply is_bounded_proj1_wire_digits.
+ generalize blem; generalize is_bounded_proj1_wire_digits; post_bounded_t.
Local Ltac define_binop f g opW blem :=
- refine (exist_fe5211_32W (opW (proj1_fe5211_32W f) (proj1_fe5211_32W g)) _);
+ 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)) _);
@@ -47,17 +57,17 @@ Local Ltac define_unop_WireToFE f 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_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 addW (f : fe5211_32W * fe5211_32W) : fe5211_32W := Eval simpl in interp_radd f.
+Definition subW (f : fe5211_32W * fe5211_32W) : fe5211_32W := Eval simpl in interp_rsub f.
+Definition mulW (f : fe5211_32W * fe5211_32W) : fe5211_32W := Eval simpl in interp_rmul f.
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.
@@ -86,7 +96,7 @@ Definition freezeW (f : fe5211_32W) : fe5211_32W := Eval cbv beta delta [prefree
Local Transparent Let_In.
(* Wrapper to allow extracted code to not unfold [mulW] *)
Definition mulW_noinline := mulW.
-Definition powW (f : fe5211_32W) chain := fold_chain_opt (proj1_fe5211_32W one) mulW_noinline chain [f].
+Definition powW (f : fe5211_32W) chain := fold_chain_opt (proj1_fe5211_32W one) (fun f g => mulW_noinline (f, g)) chain [f].
Definition invW (f : fe5211_32W) : fe5211_32W
:= Eval cbv -[Let_In fe5211_32W mulW_noinline] in powW f (chain inv_ec).
@@ -95,11 +105,11 @@ Local Ltac port_correct_and_bounded pre_rewrite opW interp_rop rop_cb :=
rewrite pre_rewrite;
intros; apply rop_cb; assumption.
-Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW carry_add.
+Lemma addW_correct_and_bounded : ibinop_correct_and_bounded addW (Curry.curry2 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.
+Lemma subW_correct_and_bounded : ibinop_correct_and_bounded subW (Curry.curry2 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.
+Lemma mulW_correct_and_bounded : ibinop_correct_and_bounded mulW (Curry.curry2 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.
@@ -142,6 +152,7 @@ Proof.
cbv [modulusW Tuple.map].
cbv [on_tuple List.map to_list to_list' from_list from_list'
+ HList.hlistP HList.hlistP'
Tuple.map2 on_tuple2 ListUtil.map2 fe5211_32WToZ length_fe5211_32].
cbv [postfreeze GF5211_32.postfreeze].
cbv [Let_In].
@@ -203,7 +214,8 @@ Proof.
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.
+ split; [ | assumption ].
+ rewrite H0', H0; reflexivity.
Qed.
Lemma powW_correct_and_bounded chain : iunop_correct_and_bounded (fun x => powW x chain) (fun x => pow x chain).
@@ -212,9 +224,11 @@ Proof.
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; pose proof (fun k0 k1 X Y => proj1 (mulW_correct_and_bounded (k0, k1) (conj X Y))) as H'.
+ cbv [Curry.curry2 Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list'] in H'.
+ rewrite <- H' by assumption.
+ apply mulW_correct_and_bounded; split; assumption. }
+ { intros; rewrite (fun X Y => proj1 (mulW_correct_and_bounded (_, _) (conj X Y))) by assumption; reflexivity. }
{ intros [|?]; autorewrite with simpl_nth_default;
(assumption || reflexivity). }
Qed.
@@ -269,8 +283,10 @@ Proof.
unfold GF5211_32.eqb.
simpl @fe5211_32WToZ in *; cbv beta iota.
intros.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe5211_32WToZ] in *.
rewrite <- frf, <- frg by assumption.
- rewrite <- fieldwisebW_correct.
+ etransitivity; [ eapply fieldwisebW_correct | ].
+ cbv [fe5211_32WToZ].
reflexivity.
Defined.
@@ -297,7 +313,7 @@ Proof.
lazymatch (eval cbv delta [GF5211_32.sqrt] in GF5211_32.sqrt) with
| (fun powf powf_squared f => dlet a := powf in _)
=> exact (dlet powx := powW (fe5211_32ZToW x) (chain GF5211_32.sqrt_ec) in
- GF5211_32.sqrt (fe5211_32WToZ powx) (fe5211_32WToZ (mulW_noinline powx powx)) x)
+ GF5211_32.sqrt (fe5211_32WToZ powx) (fe5211_32WToZ (mulW_noinline (powx, powx))) x)
| (fun f => pow f _)
=> exact (GF5211_32.sqrt x)
end.
@@ -324,21 +340,41 @@ Proof.
=> 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
- | mulW_noinline _ _ => apply mulW_correct_and_bounded
- | powW _ _ => apply powW_correct_and_bounded
- | sqrt_m1W => vm_compute; reflexivity
- | _ => assumption
- end
- end;
- subst_evars; reflexivity
+ pose proof (fun X Y => proj1 (mulW_correct_and_bounded (sqrt_m1W, a) (conj X Y))) as correctness;
+ let cbv_in_all _ := (cbv [fe5211_32WToZ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' fe5211_32WToZ Curry.curry2 HList.hlistP HList.hlistP'] in *; idtac) in
+ cbv_in_all ();
+ let solver _ := (repeat match goal with
+ | _ => progress subst
+ | _ => progress unfold fst, snd
+ | _ => progress cbv_in_all ()
+ | [ |- ?x /\ ?x ] => cut x; [ intro; split; assumption | ]
+ | [ |- is_bounded ?op = true ]
+ => let H := fresh in
+ lazymatch op with
+ | context[mulW (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[mulW_noinline (_, _)] => pose proof mulW_correct_and_bounded as H
+ | context[powW _ _] => pose proof powW_correct_and_bounded as H
+ | context[sqrt_m1W] => vm_compute; reflexivity
+ | _ => assumption
+ end;
+ cbv_in_all ();
+ apply H
+ end) in
+ rewrite <- correctness by solver (); clear correctness;
+ let lem := fresh in
+ pose proof eqbW_correct as lem; cbv_in_all (); rewrite <- lem by solver (); clear lem;
+ pose proof (pull_bool_if fe5211_32WToZ) as lem; cbv_in_all (); rewrite lem by solver (); clear lem;
+ subst_evars; reflexivity
end.
} Unfocus.
+ assert (Hfold : forall x, fe5211_32WToZ x = fe5211_32WToZ x) by reflexivity.
+ unfold fe5211_32WToZ at 2 in Hfold.
+ etransitivity.
+ Focus 2. {
+ apply Proper_Let_In_nd_changebody; [ reflexivity | intro ].
+ apply Hfold.
+ } Unfocus.
+ clear Hfold.
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
@@ -346,15 +382,22 @@ Proof.
[ cbv [Let_In]; exact (fun x => x) | apply f_equal ]
| _ => idtac
end;
- reflexivity. }
- { cbv [Let_In];
+ reflexivity.
+ }
+
+ { cbv [Let_In HList.hlistP HList.hlistP'];
try break_if;
repeat lazymatch goal with
| [ |- is_bounded (?WToZ (powW _ _)) = true ]
=> apply powW_correct_and_bounded; assumption
- | [ |- is_bounded (?WToZ (mulW _ _)) = true ]
- => apply mulW_correct_and_bounded; [ vm_compute; reflexivity | ]
- end. }
+ | [ |- is_bounded (snd (?WToZ (_, powW _ _))) = true ]
+ => generalize powW_correct_and_bounded;
+ cbv [snd Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'];
+ let H := fresh in intro H; apply H; assumption
+ | [ |- is_bounded (?WToZ (mulW (_, _))) = true ]
+ => apply mulW_correct_and_bounded; split; [ vm_compute; reflexivity | ]
+ end.
+ }
Defined.
Definition sqrtW (f : fe5211_32W) : fe5211_32W :=
@@ -394,7 +437,7 @@ 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.
+Proof. define_unop f invW (fun x p => proj2 (invW_correct_and_bounded x p)). Defined.
Definition sqrt (f : fe5211_32) : fe5211_32.
Proof. define_unop f sqrtW sqrtW_correct_and_bounded. Defined.
@@ -407,7 +450,12 @@ Local Ltac op_correct_t op opW_correct_and_bounded :=
=> rewrite proj1_wire_digits_exist_wire_digitsW
| _ => idtac
end;
- apply opW_correct_and_bounded;
+ generalize opW_correct_and_bounded;
+ cbv_tuple_map;
+ cbv [fst snd];
+ let H := fresh in
+ intro H; apply H;
+ repeat match goal with |- and _ _ => apply conj end;
lazymatch goal with
| [ |- is_bounded _ = true ]
=> apply is_bounded_proj1_fe5211_32
@@ -434,7 +482,7 @@ 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.
+Proof. op_correct_t inv (fun x p => proj1 (invW_correct_and_bounded x p)). 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.
diff --git a/src/SpecificGen/GF5211_32BoundedAddCoordinates.v b/src/SpecificGen/GF5211_32BoundedAddCoordinates.v
index 17189f4d0..96dbea3de 100644
--- a/src/SpecificGen/GF5211_32BoundedAddCoordinates.v
+++ b/src/SpecificGen/GF5211_32BoundedAddCoordinates.v
@@ -14,7 +14,7 @@ Local Ltac define_binop f g 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_coordinates / _ _ _ _ _ _ _ _ _.
+(*Local Arguments interp_radd_coordinates / _ _ _ _ _ _ _ _ _.
Definition add_coordinatesW (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe5211_32W) : Tuple.tuple fe5211_32W 4
:= Eval simpl in interp_radd_coordinates x0 x1 x2 x3 x4 x5 x6 x7 x8.
@@ -75,3 +75,4 @@ Lemma add_coordinates_correct (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe5211_32)
(proj1_fe5211_32 x7)
(proj1_fe5211_32 x8).
Proof. op_correct_t add_coordinates add_coordinatesW_correct_and_bounded. Qed.
+*)
diff --git a/src/SpecificGen/GF5211_32BoundedCommon.v b/src/SpecificGen/GF5211_32BoundedCommon.v
index 20e8c84fd..2ecc4ed1c 100644
--- a/src/SpecificGen/GF5211_32BoundedCommon.v
+++ b/src/SpecificGen/GF5211_32BoundedCommon.v
@@ -322,14 +322,14 @@ Ltac unfold_is_bounded_in' H :=
end.
Ltac preunfold_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.
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe5211_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
Ltac unfold_is_bounded_in H :=
preunfold_is_bounded_in H;
unfold_is_bounded_in' H.
Ltac preunfold_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].
+ cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map List.map fold_right List.rev List.app length_fe5211_32 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
Ltac unfold_is_bounded :=
preunfold_is_bounded;
@@ -724,7 +724,7 @@ Notation in_op_correct_and_bounded k irop op
/\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe5211_32WToZ irop))%type)
(only parsing).
-Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
+(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
: forall (irop : Tower.tower_nd fe5211_32W (Tuple.tuple fe5211_32W count_out) count_in)
(op : Tower.tower_nd GF5211_32.fe5211_32 (Tuple.tuple GF5211_32.fe5211_32 count_out) count_in)
(cont : Prop -> Prop),
@@ -792,18 +792,14 @@ Qed.
Definition inm_op_correct_and_bounded1 count_in irop op
:= Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.
-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).
+ inm_op_correct_and_bounded count_in 1 irop op.*)
+Notation inm_op_correct_and_bounded n m irop op
+ := ((forall x,
+ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe5211_32WToZ x)
+ -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe5211_32WToZ x))))
+ (only parsing).
+Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
+Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
Notation iunop_FEToZ_correct irop op
:= (forall x,
is_bounded (fe5211_32WToZ x) = true
@@ -818,20 +814,6 @@ Notation iunop_WireToFE_correct_and_bounded irop op
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).
-Notation i9top_correct_and_bounded k irop op
- := ((forall x0 x1 x2 x3 x4 x5 x6 x7 x8,
- is_bounded (fe5211_32WToZ x0) = true
- -> is_bounded (fe5211_32WToZ x1) = true
- -> is_bounded (fe5211_32WToZ x2) = true
- -> is_bounded (fe5211_32WToZ x3) = true
- -> is_bounded (fe5211_32WToZ x4) = true
- -> is_bounded (fe5211_32WToZ x5) = true
- -> is_bounded (fe5211_32WToZ x6) = true
- -> is_bounded (fe5211_32WToZ x7) = true
- -> is_bounded (fe5211_32WToZ x8) = true
- -> (Tuple.map (n:=k) fe5211_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)
- = op (fe5211_32WToZ x0) (fe5211_32WToZ x1) (fe5211_32WToZ x2) (fe5211_32WToZ x3) (fe5211_32WToZ x4) (fe5211_32WToZ x5) (fe5211_32WToZ x6) (fe5211_32WToZ x7) (fe5211_32WToZ x8))
- * HList.hlist (fun v => is_bounded v = true) (Tuple.map (n:=k) fe5211_32WToZ (irop x0 x1 x2 x3 x4 x5 x6 x7 x8)))%type)
- (only parsing).
+Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-Definition prefreeze := GF5211_32.prefreeze.
+Notation prefreeze := GF5211_32.prefreeze.
diff --git a/src/SpecificGen/GF5211_32BoundedExtendedAddCoordinates.v b/src/SpecificGen/GF5211_32BoundedExtendedAddCoordinates.v
index f6ba06ec6..4fafe3a10 100644
--- a/src/SpecificGen/GF5211_32BoundedExtendedAddCoordinates.v
+++ b/src/SpecificGen/GF5211_32BoundedExtendedAddCoordinates.v
@@ -5,7 +5,7 @@ Require Import Crypto.SpecificGen.GF5211_32ExtendedAddCoordinates.
Require Import Crypto.SpecificGen.GF5211_32BoundedAddCoordinates.
Require Import Crypto.Util.Tuple.
Require Import Crypto.Util.Tactics.
-
+(*
Lemma fieldwise_eq_extended_add_coordinates_full' twice_d P10 P11 P12 P13 P20 P21 P22 P23
: Tuple.fieldwise
(n:=4) GF5211_32BoundedCommon.eq
@@ -65,3 +65,4 @@ Proof.
destruct_head' prod.
rewrite <- fieldwise_eq_extended_add_coordinates_full'; reflexivity.
Qed.
+*)
diff --git a/src/SpecificGen/GF5211_32Reflective.v b/src/SpecificGen/GF5211_32Reflective.v
index fe29cae6a..f6f70ffec 100644
--- a/src/SpecificGen/GF5211_32Reflective.v
+++ b/src/SpecificGen/GF5211_32Reflective.v
@@ -45,6 +45,7 @@ Declare Reduction asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe5211_32W curry_unop_fe5211_32W curry_unop_wire_digitsW curry_9op_fe5211_32W
WordW.interp_op WordW.interp_base_type
@@ -56,6 +57,7 @@ Ltac asm_interp
:= cbv beta iota delta
[id
interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
+ Eta.interp_flat_type_eta Eta.interp_flat_type_eta_gen
radd rsub rmul ropp rprefreeze rge_modulus rpack runpack
curry_binop_fe5211_32W curry_unop_fe5211_32W curry_unop_wire_digitsW curry_9op_fe5211_32W
WordW.interp_op WordW.interp_base_type
@@ -65,15 +67,15 @@ Ltac asm_interp
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
+Definition interp_radd : SpecificGen.GF5211_32BoundedCommon.fe5211_32W * SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rsub : SpecificGen.GF5211_32BoundedCommon.fe5211_32W * SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
:= Eval asm_interp in interp_bexpr 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
+Definition interp_rmul : SpecificGen.GF5211_32BoundedCommon.fe5211_32W * SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
:= Eval asm_interp in interp_bexpr rmul.
(*Print interp_rmul.*)
Definition interp_rmul_correct : interp_rmul = interp_bexpr rmul := eq_refl.
diff --git a/src/SpecificGen/GF5211_32Reflective/Common.v b/src/SpecificGen/GF5211_32Reflective/Common.v
index 2ea85da7f..176e20f08 100644
--- a/src/SpecificGen/GF5211_32Reflective/Common.v
+++ b/src/SpecificGen/GF5211_32Reflective/Common.v
@@ -7,7 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Wf.
Require Import Crypto.Reflection.ExprInversion.
+Require Import Crypto.Reflection.Tuple.
Require Import Crypto.Reflection.Relations.
+Require Import Crypto.Reflection.Eta.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -15,13 +17,14 @@ Require Import Crypto.Reflection.Z.Reify.
Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.Z.Syntax.Equality.
Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Util.Curry.
Require Import Crypto.Util.Tower.
Require Import Crypto.Util.LetIn.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ZUtil.
Require Import Crypto.Util.Tactics.
+Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Notations.
Notation Expr := (Expr base_type op).
@@ -43,30 +46,24 @@ Defined.
Definition Expr_n_OpT (count_out : nat) : flat_type base_type
:= Eval cbv [Syntax.tuple Syntax.tuple' fe5211_32T] in
Syntax.tuple fe5211_32T count_out.
-Fixpoint Expr_nm_OpT (count_in count_out : nat) : type base_type
- := match count_in with
- | 0 => Expr_n_OpT count_out
- | S n => SmartArrow fe5211_32T (Expr_nm_OpT n count_out)
- end.
+Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
+ := Eval cbv [Syntax.tuple Syntax.tuple' fe5211_32T Expr_n_OpT] in
+ Arrow (Syntax.tuple fe5211_32T count_in) (Expr_n_OpT count_out).
Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from (uncurry_unop_fe5211_32 ge_modulus). Defined.
+Proof. make_type_from ge_modulus. Defined.
Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from (uncurry_unop_wire_digits unpack). Defined.
+Proof. make_type_from unpack. Defined.
Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from (uncurry_unop_fe5211_32 pack). Defined.
+Proof. make_type_from pack. Defined.
Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
Definition ExprArgT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpT.
+ := Eval compute in domain ExprUnOpT.
Definition ExprArgWireT : flat_type base_type
- := Eval compute in remove_all_binders ExprUnOpFEToWireT.
-Definition ExprArgRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpT.
-Definition ExprArgWireRevT : flat_type base_type
- := Eval compute in all_binders_for ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Tbase TZ).
+ := Eval compute in domain ExprUnOpWireToFET.
+Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
@@ -75,99 +72,67 @@ Definition ExprBinOp : Type := Expr ExprBinOpT.
Definition ExprUnOp : Type := Expr ExprUnOpT.
Definition Expr4Op : Type := Expr Expr4OpT.
Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr ExprArgT.
-Definition ExprArgWire : Type := Expr ExprArgWireT.
-Definition ExprArgRev : Type := Expr ExprArgRevT.
-Definition ExprArgWireRev : Type := Expr ExprArgWireRevT.
+Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
+Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
Definition expr_nm_Op count_in count_out var : Type
:= expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Tbase TZ).
+Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) ExprArgT.
-Definition exprArgWire var : Type := expr base_type op (var:=var) ExprArgWireT.
-Definition exprArgRev var : Type := expr base_type op (var:=var) ExprArgRevT.
-Definition exprArgWireRev var : Type := expr base_type op (var:=var) ExprArgWireRevT.
-
-Local Ltac bounds_from_list_cps ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(fun T (extra : T) => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, extra))
- | (?x :: ?xs)%list => let bs := bounds_from_list_cps xs in
- constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
- end.
-
-Local Ltac make_bounds_cps ls extra :=
- let v := bounds_from_list_cps (List.rev ls) in
- let v := (eval compute in v) in
- exact (v _ extra).
-
-Local Ltac bounds_from_list ls :=
- lazymatch (eval hnf in ls) with
- | (?x :: nil)%list => constr:(Some {| Bounds.lower := fst x ; Bounds.upper := snd x |})
- | (?x :: ?xs)%list => let bs := bounds_from_list xs in
- constr:((Some {| Bounds.lower := fst x ; Bounds.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.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out : interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type.
-Proof.
- refine match count_in return interp_all_binders_for (Expr_nm_OpT count_in count_out) ZBounds.interp_base_type with
- | 0 => tt
- | S n => let v := interp_all_binders_for_to' (Expr_nm_Op_bounds n count_out) in
- interp_all_binders_for_of' _
- end; simpl.
- make_bounds_cps (Tuple.to_list _ bounds) v.
-Defined.
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
+Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
+
+Definition make_bound (x : Z * Z) : ZBounds.t
+ := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
+
+Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
+ := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
+ | 0 => tt
+ | S n
+ => let b := (Tuple.map make_bound bounds) in
+ let bs := Expr_nm_Op_bounds n count_out in
+ match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
+ | 0 => fun _ => b
+ | S n' => fun bs => (bs, b)
+ end bs
+ end.
+Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
:= Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
:= Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
:= Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
:= Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_all_binders_for ExprUnOpWireToFET ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ wire_digit_bounds). Defined.
+Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
+ := Tuple.map make_bound wire_digit_bounds.
-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 (@WordW.interp_op) e).
+Definition interp_bexpr : ExprBinOp -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W * SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- := fun e => curry_unop_fe5211_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.word64
- := fun e => curry_unop_fe5211_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W -> SpecificGen.GF5211_32BoundedCommon.wire_digitsW
- := fun e => curry_unop_fe5211_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> SpecificGen.GF5211_32BoundedCommon.wire_digitsW -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- := fun e => curry_unop_wire_digitsW (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Definition interp_9_4expr : Expr9_4Op
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
- -> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
+ -> Tuple.tuple SpecificGen.GF5211_32BoundedCommon.fe5211_32W 9
-> Tuple.tuple SpecificGen.GF5211_32BoundedCommon.fe5211_32W 4
- := fun e => curry_9op_fe5211_32W (Interp (@WordW.interp_op) e).
+ := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) op) (only parsing).
+ := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 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
@@ -181,40 +146,39 @@ Notation op9_4_correct_and_bounded rop op
Ltac rexpr_cbv :=
lazymatch goal with
- | [ |- { rexpr | interp_type_gen_rel_pointwise _ (Interp _ (t:=?T) rexpr) (?uncurry ?oper) } ]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
=> let operf := head oper in
let uncurryf := head uncurry in
try cbv delta [T]; try cbv delta [oper];
try cbv beta iota delta [uncurryf]
+ | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
+ => let operf := head oper in
+ try cbv delta [T]; try cbv delta [oper]
end;
- cbv beta iota delta [interp_flat_type Z.interp_base_type interp_base_type zero_].
+ cbv beta iota delta [interp_flat_type interp_base_type zero_ GF5211_32.fe5211_32 GF5211_32.wire_digits].
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 }
+ | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
(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 rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT (uncurry_9op_fe5211_32 op)) (only parsing).
+Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
+Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
+Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
+Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
+Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
+Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT 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 := ropW' in
- let ropZ := ropW' in
- let ropBounds := ropW' in
- let ropBoundedWordW := ropW' in
- ropZ = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise2 Relations.related_Z (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@Z.interp_op) ropZ)
- /\ interp_type_rel_pointwise2 Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@ZBounds.interp_op) ropBounds)
- /\ interp_type_rel_pointwise2 Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropBoundedWordW) (Interp (@WordW.interp_op) ropW))
+ ropW' = proj1_sig ropZ_sig
+ /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
+ /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
(only parsing).
Ltac app_tuples x y :=
@@ -227,7 +191,7 @@ Ltac app_tuples x y :=
Local Arguments Tuple.map2 : simpl never.
Local Arguments Tuple.map : simpl never.
-
+(*
Fixpoint args_to_bounded_helperT {n}
(v : Tuple.tuple' WordW.wordW n)
(bounds : Tuple.tuple' (Z * Z) n)
@@ -299,14 +263,15 @@ Proof.
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
@@ -327,6 +292,7 @@ Ltac assoc_right_tuple x so_far :=
end
end.
+(*
Local Ltac make_args x :=
let x' := fresh "x'" in
compute in x |- *;
@@ -338,11 +304,6 @@ Local Ltac make_args x :=
let xv := assoc_right_tuple x'' (@None) in
refine (SmartVarf (xv : interp_flat_type _ t')).
-Definition unop_make_args {var} (x : exprArg var) : exprArgRev var.
-Proof. make_args x. Defined.
-Definition unop_wire_make_args {var} (x : exprArgWire var) : exprArgWireRev var.
-Proof. make_args x. Defined.
-
Local Ltac args_to_bounded x H :=
let x' := fresh in
set (x' := x);
@@ -351,29 +312,138 @@ Local Ltac args_to_bounded x H :=
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
+ let c := (eval cbv [domain 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) ]
).
+ *)
+
+Section gen.
+ Definition bounds_are_good_gen
+ {n : nat} (bounds : Tuple.tuple (Z * Z) n)
+ := let res :=
+ Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
+ in
+ List.fold_right andb true (Tuple.to_list n res).
+ Definition unop_args_to_bounded'
+ (bs : Z * Z)
+ (Hbs : bounds_are_good_gen (n:=1) bs = true)
+ (x : word64)
+ (H : is_bounded_gen (Tuple.map (fun v : word64 => (v : Z)) (n:=1) x) bs = true)
+ : BoundedWordW.BoundedWord.
+ Proof.
+ refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
+ unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
+ abstract (
+ destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
+ repeat apply conj; assumption
+ ).
+ Defined.
+ Fixpoint n_op_args_to_bounded'
+ n
+ : forall (bs : Tuple.tuple' (Z * Z) n)
+ (Hbs : bounds_are_good_gen (n:=S n) bs = true)
+ (x : Tuple.tuple' word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
+ Proof.
+ destruct n as [|n']; simpl in *.
+ { exact unop_args_to_bounded'. }
+ { refine (fun bs Hbs x H
+ => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
+ @unop_args_to_bounded' (snd bs) _ (snd x) _));
+ clear n_op_args_to_bounded';
+ simpl in *;
+ [ clear x H | clear Hbs | clear x H | clear Hbs ];
+ unfold bounds_are_good_gen, is_bounded_gen in *;
+ abstract (
+ repeat first [ progress simpl in *
+ | assumption
+ | reflexivity
+ | progress Bool.split_andb
+ | progress destruct_head prod
+ | match goal with
+ | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
+ end
+ | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
+ | progress break_match_hyps
+ | rewrite Bool.andb_true_iff; apply conj
+ | unfold Tuple.map, Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
+ ). }
+ Defined.
+
+ Definition n_op_args_to_bounded
+ n
+ : forall (bs : Tuple.tuple (Z * Z) n)
+ (Hbs : bounds_are_good_gen bs = true)
+ (x : Tuple.tuple word64 n)
+ (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
+ := match n with
+ | 0 => fun _ _ _ _ => tt
+ | S n' => @n_op_args_to_bounded' n'
+ end.
-Definition unop_args_to_bounded (x : fe5211_32W) (H : is_bounded (fe5211_32WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
+ Fixpoint nm_op_args_to_bounded' n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple' (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => @n_op_args_to_bounded m bs Hbs
+ | S n' => fun x H
+ => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
+ @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
+ end.
+ Definition nm_op_args_to_bounded n m
+ (bs : Tuple.tuple (Z * Z) m)
+ (Hbs : bounds_are_good_gen bs = true)
+ : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple (Tbase TZ) n))
+ (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
+ x),
+ interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
+ := match n with
+ | 0 => fun _ _ => tt
+ | S n' => @nm_op_args_to_bounded' n' m bs Hbs
+ end.
+End gen.
+
+Local Ltac get_inner_len T :=
+ lazymatch T with
+ | (?T * _)%type => get_inner_len T
+ | ?T => get_len T
+ end.
+Local Ltac get_outer_len T :=
+ lazymatch T with
+ | (?A * ?B)%type => let a := get_outer_len A in
+ let b := get_outer_len B in
+ (eval compute in (a + b)%nat)
+ | ?T => constr:(1%nat)
+ end.
+Local Ltac args_to_bounded x H :=
+ let t := type of x in
+ let m := get_inner_len t in
+ let n := get_outer_len t in
+ let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
+ let H := (eval cbv beta in (H eq_refl)) in
+ exact H.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
Definition binop_args_to_bounded (x : fe5211_32W * fe5211_32W)
(H : is_bounded (fe5211_32WToZ (fst x)) = true)
(H' : is_bounded (fe5211_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for ExprBinOpT).
-Proof.
- let v := app_tuples (unop_args_to_bounded (fst x) H) (unop_args_to_bounded (snd x) H') in
- exact v.
-Defined.
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
+Proof. args_to_bounded x (conj H H'). Defined.
+Definition unop_args_to_bounded (x : fe5211_32W) (H : is_bounded (fe5211_32WToZ x) = true)
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain 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 _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
+Proof. args_to_bounded x H. Defined.
Definition op9_args_to_bounded (x : fe5211_32W * fe5211_32W * fe5211_32W * fe5211_32W * fe5211_32W * fe5211_32W * fe5211_32W * fe5211_32W * fe5211_32W)
(H0 : is_bounded (fe5211_32WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
(H1 : is_bounded (fe5211_32WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
@@ -384,58 +454,39 @@ Definition op9_args_to_bounded (x : fe5211_32W * fe5211_32W * fe5211_32W * fe521
(H6 : is_bounded (fe5211_32WToZ (snd (fst (fst x)))) = true)
(H7 : is_bounded (fe5211_32WToZ (snd (fst x))) = true)
(H8 : is_bounded (fe5211_32WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (all_binders_for Expr9_4OpT).
-Proof.
- let v := constr:(unop_args_to_bounded _ H8) in
- let v := app_tuples (unop_args_to_bounded _ H7) v in
- let v := app_tuples (unop_args_to_bounded _ H6) v in
- let v := app_tuples (unop_args_to_bounded _ H5) v in
- let v := app_tuples (unop_args_to_bounded _ H4) v in
- let v := app_tuples (unop_args_to_bounded _ H3) v in
- let v := app_tuples (unop_args_to_bounded _ H2) v in
- let v := app_tuples (unop_args_to_bounded _ H1) v in
- let v := app_tuples (unop_args_to_bounded _ H0) v in
- exact v.
-Defined.
-
+ : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
+Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
+Local Ltac make_bounds_prop' bounds bounds' :=
+ first [ refine (andb _ _);
+ [ destruct bounds' as [bounds' _], bounds as [bounds _]
+ | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
+ try make_bounds_prop' bounds bounds'
+ | 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) ].
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.
+ make_bounds_prop' bounds bounds'.
+Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain 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.
+Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
Proof.
refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (remove_all_binders Expr9_4OpT)) : bool.
+Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-
-Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
+(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe5211_32 f k)).
@@ -460,12 +511,11 @@ Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ v
:= fun x
=> LetIn (invert_Return (unop_make_args x))
(fun k => invert_Return (Apply length_fe5211_32 f k)).
-
+*)
(* 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
@@ -473,23 +523,25 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
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 interp_9_4expr
+ interp_flat_type_eta interp_flat_type_eta_gen
curry_binop_fe5211_32W curry_unop_fe5211_32W curry_unop_wire_digitsW curry_9op_fe5211_32W
curry_binop_fe5211_32 curry_unop_fe5211_32 curry_unop_wire_digits curry_9op_fe5211_32
uncurry_binop_fe5211_32W uncurry_unop_fe5211_32W uncurry_unop_wire_digitsW uncurry_9op_fe5211_32W
uncurry_binop_fe5211_32 uncurry_unop_fe5211_32 uncurry_unop_wire_digits uncurry_9op_fe5211_32
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT
- interp_type_gen_rel_pointwise interp_type_gen_rel_pointwise] in *;
+ ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] 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 *;
+ change word64 with WordW.wordW in *;
rewrite <- Heq; clear Heq;
destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise2_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
+ pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
specialize_by repeat first [ progress intros
+ | progress unfold RelationClasses.Reflexive
| reflexivity
| assumption
| progress destruct_head' base_type
@@ -509,23 +561,33 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
+ Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
+ simpl @interp_flat_type in *;
+ (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
+ change (WordW.interp_base_type TZ) with v in *);
+ cbv beta iota zeta in *;
lazymatch goal with
| [ |- fe5211_32WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
| [ |- wire_digitsWToZ ?x = _ /\ _ ]
=> generalize dependent x; intros
+ | [ |- (Tuple.map fe5211_32WToZ ?x = _) /\ _ ]
+ => generalize dependent x; intros
| [ |- ((Tuple.map fe5211_32WToZ ?x = _) * _)%type ]
=> generalize dependent x; intros
| [ |- _ = _ ]
=> exact Hbounds_left
end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type remove_all_binders] in *;
+ cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
destruct_head' prod;
change word64ToZ with WordW.wordWToZ in *;
(split; [ exact Hbounds_left | ]);
cbv [interp_flat_type] in *;
cbv [fst snd
+ Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
+ make_bound
+ Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
destruct_head' ZBounds.bounds;
@@ -534,12 +596,12 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
destruct_head' and;
Z.ltb_to_lt;
change WordW.wordWToZ with word64ToZ in *;
- cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd];
+ cbv [Tuple.map HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe5211_32WToZ HList.hlistP HList.hlistP'];
+ cbv [WordW.bit_width BitSize64.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
repeat split; unfold_is_bounded;
Z.ltb_to_lt;
try omega; try reflexivity.
-
Ltac rexpr_correct :=
let ropW' := fresh in
let ropZ_sig := fresh in
@@ -555,9 +617,13 @@ Ltac rexpr_correct :=
Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
Notation compute_bounds opW bounds
- := (ApplyInterpedAll (Interp (@ZBounds.interp_op) opW) bounds)
+ := (Interp (@ZBounds.interp_op) opW bounds)
(only parsing).
+Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
+
+Ltac prove_rexpr_wfT
+ := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
Module Export PrettyPrinting.
(* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
@@ -569,23 +635,21 @@ Module Export PrettyPrinting.
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 {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end
+ := fun (t : base_type) (x : ZBounds.interp_base_type t)
+ => match x with
+ | Some {| Bounds.lower := l ; Bounds.upper := u |}
+ => in_range l u
+ | None
+ => overflow
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
+ Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
+ := match t return interp_flat_type _ t -> result with
+ | Tbase _ => fun v => match v with
+ | overflow => yes
+ | in_range _ _ => no
+ | enlargen _ => borked
+ end
| Unit => fun _ => no
| Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
| no, no => no
diff --git a/src/SpecificGen/GF5211_32Reflective/Common9_4Op.v b/src/SpecificGen/GF5211_32Reflective/Common9_4Op.v
index da61302b4..c3878caee 100644
--- a/src/SpecificGen/GF5211_32Reflective/Common9_4Op.v
+++ b/src/SpecificGen/GF5211_32Reflective/Common9_4Op.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -42,8 +41,8 @@ Lemma Expr9_4Op_correct_and_bounded
let (Hx7, Hx8) := (eta_and Hx78) in
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -80,29 +79,24 @@ Lemma Expr9_4Op_correct_and_bounded
let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => op9_4_bounds_good bounds = true
| None => False
end)
: op9_4_correct_and_bounded ropW op.
Proof.
- intros x0 x1 x2 x3 x4 x5 x6 x7 x8.
- intros Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8.
- pose x0 as x0'.
- pose x1 as x1'.
- pose x2 as x2'.
- pose x3 as x3'.
- pose x4 as x4'.
- pose x5 as x5'.
- pose x6 as x6'.
- pose x7 as x7'.
- pose x8 as x8'.
- hnf in x0, x1, x2, x3, x4, x5, x6, x7, x8; destruct_head' prod.
- specialize (H0 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ intros xs Hxs.
+ pose xs as xs'.
+ compute in xs.
+ destruct_head' prod.
+ cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
+ pose Hxs as Hxs'.
+ destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
+ specialize (H0 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 (x0', x1', x2', x3', x4', x5', x6', x7', x8')
+ specialize (H1 xs'
(conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded (x0', x1', x2', x3', x4', x5', x6', x7', x8') Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
+ Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.64s) (successful) *)*)
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v b/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
index 6d4f73920..f54be4d6a 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonBinOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -18,8 +17,8 @@ Lemma ExprBinOp_correct_and_bounded
let Hy := let (Hx, Hy) := Hxy in Hy in
let args := binop_args_to_bounded xy Hx Hy in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -33,18 +32,19 @@ Lemma ExprBinOp_correct_and_bounded
let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => binop_bounds_good bounds = true
| None => False
end)
: binop_correct_and_bounded 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
+ intros xy HxHy.
+ pose xy as xy'.
+ compute in xy; destruct_head' prod.
+ specialize (H0 xy' HxHy).
+ specialize (H1 xy' HxHy).
+ destruct HxHy as [Hx Hy].
+ let args := constr:(binop_args_to_bounded xy' 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
index d8bc7dcfa..c4fa848f6 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOp.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOp_correct_and_bounded
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOp_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unop_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
index 91469dc14..433609622 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToWire.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToWire_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToWire_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
index 68389da47..6b48722c2 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpFEToZ.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
(Hx : is_bounded (fe5211_32WToZ x) = true),
let args := unop_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpFEToZ_correct_and_bounded
let args := unop_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopFEToZ_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v b/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
index 48b8c853a..324729d2c 100644
--- a/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
+++ b/src/SpecificGen/GF5211_32Reflective/CommonUnOpWireToFE.v
@@ -3,7 +3,6 @@ Require Import Crypto.SpecificGen.GF5211_32BoundedCommon.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Util.Tactics.
Local Opaque Interp.
@@ -15,8 +14,8 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
(Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
let args := unopWireToFE_args_to_bounded x Hx in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@BoundedWordW.interp_op) ropW)
- (LiftOption.to' (Some args)))
+ (Interp (@BoundedWordW.interp_op) ropW
+ (LiftOption.to' (Some args)))
with
| Some _ => True
| None => False
@@ -27,7 +26,7 @@ Lemma ExprUnOpWireToFE_correct_and_bounded
let args := unopWireToFE_args_to_bounded x Hx in
let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
match LiftOption.of'
- (ApplyInterpedAll (Interp (@ZBounds.interp_op) ropW) (LiftOption.to' (Some x')))
+ (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
with
| Some bounds => unopWireToFE_bounds_good bounds = true
| None => False
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/AddCoordinates.v b/src/SpecificGen/GF5211_32Reflective/Reified/AddCoordinates.v
index 90f391ba5..e85a6f2cd 100644
--- a/src/SpecificGen/GF5211_32Reflective/Reified/AddCoordinates.v
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/AddCoordinates.v
@@ -7,7 +7,6 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.ExprInversion.
Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
@@ -17,7 +16,10 @@ Require Export Crypto.Reflection.Z.Syntax.
Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
+Require Import Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.Mul.
@@ -33,24 +35,28 @@ Require Import Bedrock.Word.
Definition radd_coordinatesZ' var twice_d P1 P2
:= @Extended.add_coordinates_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
twice_d _
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
P1 P2.
+Local Notation eta x := (fst x, snd x).
+
Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => radd_coordinatesZ'
- var (Return twice_d)
- (Return P10, Return P11, Return P12, Return P13)
- (Return P20, Return P21, Return P22, Return P23))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
+ => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
+ := twice_d_P1_P2 in
+ radd_coordinatesZ'
+ var (SmartVarf twice_d)
+ (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
+ (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
Definition add_coordinates
:= fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
@@ -59,70 +65,60 @@ Definition add_coordinates
twice_d (P10, P11, P12, P13) (P20, P21, P22, P23).
Definition uncurried_add_coordinates
- := apply9_nd
- (@uncurry_unop_fe5211_32)
- add_coordinates.
+ := fun twice_d_P1_P2
+ => let twice_d := fst twice_d_P1_P2 in
+ let (P1, P2) := eta (snd twice_d_P1_P2) in
+ @Extended.add_coordinates
+ _ add sub mul
+ twice_d P1 P2.
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
let H := fresh in
pose k as k';
assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
+ repeat (unfold k'; change k with k'; unfold interpf_step);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF5211_32.fe5211_32 -> GF5211_32.fe5211_32 -> GF5211_32.fe5211_32)
- (x' y' : GF5211_32.fe5211_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe5211_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
+Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe5211_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF5211_32.fe5211_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe5211_32 Apply' interp fst snd].
- rewrite (invert_Abs_Some (e:=f) eq_refl).
+ cbv [radd_coordinatesZ''].
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type] in x.
+ destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
+ | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
end.
- reflexivity.
-Qed.
-
-Lemma radd_coordinatesZ_sigP : rexpr_sigP Expr9_4OpT uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT radd_coordinatesZ'' uncurried_add_coordinates uncurry_unop_fe5211_32 SmartAbs radd_coordinatesZ' exprArg Extended.add_coordinates_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe5211_32 add_coordinates].
- intros.
- unfold invert_Return at 13 14 15 16.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [radd_coordinatesZ' add_coordinates Extended.add_coordinates_gen uncurried_add_coordinates SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -132,48 +128,56 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 13 14 15; unfold interpf_step.
- cbv beta iota delta [Extended.add_coordinates].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition radd_coordinatesZ_sig : rexpr_9_4op_sig add_coordinates.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv beta iota delta [Extended.add_coordinates interp_flat_type interp_base_type GF5211_32.fe5211_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
Proof.
- eexists.
- apply radd_coordinatesZ_sigP.
-Defined.
+ exact radd_coordinatesZ_sigP'.
+Qed.
+Definition radd_coordinatesZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
+
+Definition radd_coordinates_input_bounds
+ := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW Expr9Op_bounds.
+Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition radd_coordinatesW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- radd_coordinatesW add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
+ radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW Expr9Op_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW Expr9Op_bounds).
+Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
+Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/LadderStep.v b/src/SpecificGen/GF5211_32Reflective/Reified/LadderStep.v
index 16c0173a0..4e7a5f9b3 100644
--- a/src/SpecificGen/GF5211_32Reflective/Reified/LadderStep.v
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/LadderStep.v
@@ -7,8 +7,9 @@ Require Import Crypto.Reflection.Syntax.
Require Import Crypto.Reflection.SmartMap.
Require Import Crypto.Reflection.Relations.
Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Application.
Require Import Crypto.Reflection.Linearize.
+Require Import Crypto.Reflection.Eta.
+Require Import Crypto.Reflection.EtaInterp.
Require Import Crypto.Reflection.Z.Interpretations64.
Require Crypto.Reflection.Z.Interpretations64.Relations.
Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
@@ -18,6 +19,7 @@ Require Import Crypto.Reflection.InterpWfRel.
Require Import Crypto.Reflection.LinearizeInterp.
Require Import Crypto.Reflection.WfReflective.
Require Import Crypto.Spec.MxDH.
+Require Import Crypto.SpecificGen.GF5211_32Reflective.Common.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.Add.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.Sub.
Require Import Crypto.SpecificGen.GF5211_32Reflective.Reified.Mul.
@@ -30,50 +32,80 @@ Require Import Crypto.Util.Tactics.
Require Import Crypto.Util.Notations.
Require Import Bedrock.Word.
-(* XXX FIXME: Remove dummy code *)
-Definition rladderstepZ' var (T:=_) (dummy0 dummy1 dummy2 a24 x0 : T) P1 P2
+Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
:= @MxDH.ladderstep_gen
_
- (fun x y => ApplyBinOp (proj1_sig raddZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rsubZ_sig var) x y)
- (fun x y => ApplyBinOp (proj1_sig rmulZ_sig var) x y)
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
+ (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
a24
_
- (fun x y z w => (invert_Return x, invert_Return y, invert_Return z, invert_Return w)%expr)
- (fun v f => LetIn (invert_Return v)
- (fun k => f (Return (SmartVarf k))))
+ (fun x y z w => (x, y, z, w)%expr)
+ (fun v f => LetIn v
+ (fun k => f (SmartVarf k)))
x0
P1 P2.
Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize (fun var
- => apply9
- (fun A B => SmartAbs (A := A) (B := B))
- (fun dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21
- => rladderstepZ'
- var (Return dummy0) (Return dummy1) (Return dummy2)
- (Return a24) (Return x0)
- (Return P10, Return P11)
- (Return P20, Return P21))).
+ := Linearize
+ (ExprEta
+ (fun var
+ => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
+ => let '(a24, x0, ((P10, P11), (P20, P21)))
+ := a24_x0_P1_P2 in
+ rladderstepZ'
+ var (SmartVarf a24) (SmartVarf x0)
+ (SmartVarf P10, SmartVarf P11)
+ (SmartVarf P20, SmartVarf P21)))).
-Definition ladderstep (T := _)
- := fun (dummy0 dummy1 dummy2 a24 x0 P10 P11 P20 P21 : T)
- => @MxDH.ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
+Local Notation eta x := (fst x, snd x).
+
+Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
+ Eval cbv beta delta [MxDH.ladderstep_gen] in
+ @MxDH.ladderstep_gen
+ F Fadd Fsub Fmul a24
+ (F*F*F*F)
+ (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
+ (fun x f => dlet y := x in f y)
+ X1 P1 P2.
Definition uncurried_ladderstep
- := apply9_nd
- (@uncurry_unop_fe5211_32)
- ladderstep.
+ := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
+ => let a24 := fst (fst a24_x0_P1_P2) in
+ let x0 := snd (fst a24_x0_P1_P2) in
+ let '(P1, P2) := eta (snd a24_x0_P1_P2) in
+ let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
+ @ladderstep_other_assoc
+ _ add sub mul
+ a24 x0 (P10, P11) (P20, P21).
+Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
+ (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
+ (only parsing).
Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (interp_type_gen_rel_pointwise (fun _ => Logic.eq) (Interp interp_op (t:=T) rexprZ) uncurried_op)
+ (forall x, rexpr_sigPf T uncurried_op rexprZ x)
(only parsing).
Local Notation rexpr_sig T uncurried_op :=
{ rexprZ | rexpr_sigP T uncurried_op rexprZ }
(only parsing).
+Local Ltac fold_interpf' :=
+ let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k'.
+
+Local Ltac fold_interpf :=
+ let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
+ let k' := fresh in
+ let H := fresh in
+ pose k as k';
+ assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
+ change k with k'; clearbody k'; subst k';
+ fold_interpf'.
+
Local Ltac repeat_step_interpf :=
let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
let k' := fresh in
@@ -83,51 +115,14 @@ Local Ltac repeat_step_interpf :=
repeat (unfold interpf_step at 1; change k with k' at 1);
clearbody k'; subst k'.
-Lemma interp_helper
- (f : Syntax.Expr base_type op ExprBinOpT)
- (x y : exprArg interp_base_type)
- (f' : GF5211_32.fe5211_32 -> GF5211_32.fe5211_32 -> GF5211_32.fe5211_32)
- (x' y' : GF5211_32.fe5211_32)
- (H : interp_type_gen_rel_pointwise
- (fun _ => Logic.eq)
- (Interp interp_op f) (uncurry_binop_fe5211_32 f'))
- (Hx : interpf interp_op (invert_Return x) = x')
- (Hy : interpf interp_op (invert_Return y) = y')
- : interpf interp_op (invert_Return (ApplyBinOp (f interp_base_type) x y)) = f' x' y'.
-Proof.
- cbv [interp_type_gen_rel_pointwise ExprBinOpT uncurry_binop_fe5211_32 interp_flat_type] in H.
- simpl @interp_base_type in *.
- cbv [GF5211_32.fe5211_32] in x', y'.
- destruct_head' prod.
- rewrite <- H; clear H.
- cbv [ExprArgT] in *; simpl in *.
- rewrite Hx, Hy; clear Hx Hy.
- unfold Let_In; simpl.
- cbv [Interp].
- simpl @interp_type.
- change (fun t => interp_base_type t) with interp_base_type in *.
- generalize (f interp_base_type); clear f; intro f.
- cbv [Apply length_fe5211_32 Apply' interp fst snd].
- let f := match goal with f : expr _ _ _ |- _ => f end in
- rewrite (invert_Abs_Some (e:=f) eq_refl).
- repeat match goal with
- | [ |- appcontext[invert_Abs ?f ?x] ]
- => generalize (invert_Abs f x); clear f;
- let f' := fresh "f" in
- intro f';
- first [ rewrite (invert_Abs_Some (e:=f') eq_refl)
- | rewrite (invert_Return_Some (e:=f') eq_refl) at 2 ]
- end.
- reflexivity.
-Qed.
-
-Lemma rladderstepZ_sigP : rexpr_sigP Expr9_4OpT uncurried_ladderstep rladderstepZ''.
+Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
cbv [rladderstepZ''].
- etransitivity; [ apply InterpLinearize | ].
- cbv beta iota delta [apply9 apply9_nd interp_type_gen_rel_pointwise Expr9_4OpT SmartArrow ExprArgT rladderstepZ'' uncurried_ladderstep uncurry_unop_fe5211_32 SmartAbs rladderstepZ' exprArg MxDH.ladderstep_gen Interp interp unop_make_args SmartVarf smart_interp_flat_map length_fe5211_32 ladderstep].
- intros; cbv beta zeta.
- unfold invert_Return at 14 15 16 17.
+ intro x; rewrite InterpLinearize, InterpExprEta.
+ cbv [domain interp_flat_type interp_base_type] in x.
+ destruct_head' prod.
+ cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
+ cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
repeat match goal with
| [ |- appcontext[@proj1_sig ?A ?B ?v] ]
=> let k := fresh "f" in
@@ -137,48 +132,58 @@ Proof.
set (k' := @proj1_sig A B k);
pose proof (proj2_sig k) as H;
change (proj1_sig k) with k' in H;
- clearbody k'; clear k
+ clearbody k'; clear k;
+ cbv beta in *
end.
- unfold interpf; repeat_step_interpf.
- unfold interpf at 14 15 16; unfold interpf_step.
- cbv beta iota delta [MxDH.ladderstep_other_assoc].
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (let x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros
- end;
- repeat match goal with
- | [ |- interpf interp_op (invert_Return (ApplyBinOp _ _ _)) = _ ]
- => apply interp_helper; [ assumption | | ]
- | [ |- interpf interp_op (invert_Return (Return (_, _)%expr)) = (_, _) ]
- => vm_compute; reflexivity
- | [ |- (_, _) = (_, _) ]
- => reflexivity
- | _ => simpl; rewrite <- !surjective_pairing; reflexivity
- end.
-Time Qed.
-
-Definition rladderstepZ_sig : rexpr_9_4op_sig ladderstep.
+ cbv [Interp Curry.curry2] in *.
+ unfold interpf, interpf_step; fold_interpf.
+ cbv [ladderstep_other_assoc interp_flat_type GF5211_32.fe5211_32].
+ Time
+ abstract (
+ repeat match goal with
+ | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
+ => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
+ (_ : y = y')
+ (_ : forall x, z x = z' x));
+ cbv beta; intros;
+ [ cbv [Let_In Common.ExprBinOpT] | ]
+ end;
+ repeat match goal with
+ | _ => rewrite !interpf_invert_Abs
+ | _ => rewrite_hyp !*
+ | _ => progress cbv [interp_base_type]
+ | [ |- ?x = ?x ] => reflexivity
+ | _ => rewrite <- !surjective_pairing
+ end
+ ).
+Time Defined.
+Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
Proof.
- eexists.
- apply rladderstepZ_sigP.
-Defined.
+ exact rladderstepZ_sigP'.
+Qed.
+Definition rladderstepZ_sig
+ := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
+
+Definition rladderstep_input_bounds
+ := (ExprUnOp_bounds, ExprUnOp_bounds,
+ ((ExprUnOp_bounds, ExprUnOp_bounds),
+ (ExprUnOp_bounds, ExprUnOp_bounds))).
Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW Expr9Op_bounds.
+Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
Local Obligation Tactic := intros; vm_compute; constructor.
+(*
Program Definition rladderstepW_correct_and_bounded
:= Expr9_4Op_correct_and_bounded
- rladderstepW ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
+ rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
_ _.
+ *)
Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW Expr9Op_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW Expr9Op_bounds).
+Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
+Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v b/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v
index 4788f9f1a..5cf7ea98a 100644
--- a/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v
+++ b/src/SpecificGen/GF5211_32Reflective/Reified/PreFreeze.v
@@ -1,6 +1,6 @@
Require Import Crypto.SpecificGen.GF5211_32Reflective.CommonUnOp.
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. cbv [prefreeze GF5211_32.prefreeze]. reify_sig. Defined.
+Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. 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.
diff --git a/src/SpecificGen/GF5211_32ReflectiveAddCoordinates.v b/src/SpecificGen/GF5211_32ReflectiveAddCoordinates.v
index ee7bbb2e2..8105fd507 100644
--- a/src/SpecificGen/GF5211_32ReflectiveAddCoordinates.v
+++ b/src/SpecificGen/GF5211_32ReflectiveAddCoordinates.v
@@ -32,7 +32,7 @@ Import ListNotations.
Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Local Open Scope Z.
-Time Definition radd_coordinates : Expr9_4Op := Eval vm_compute in radd_coordinatesW.
+Time Definition radd_coordinates : Expr _ := Eval vm_compute in radd_coordinatesW.
Declare Reduction asm_interp_add_coordinates
:= cbv beta iota delta
@@ -57,7 +57,7 @@ Ltac asm_interp_add_coordinates
mapf_interp_flat_type WordW.interp_base_type word64ize
Interp interp interp_flat_type interpf interp_flat_type fst snd].
-
+(*
Time Definition interp_radd_coordinates : SpecificGen.GF5211_32BoundedCommon.fe5211_32W
-> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
-> SpecificGen.GF5211_32BoundedCommon.fe5211_32W
@@ -74,3 +74,4 @@ Time Definition interp_radd_coordinates_correct : interp_radd_coordinates = inte
Lemma radd_coordinates_correct_and_bounded : op9_4_correct_and_bounded radd_coordinates add_coordinates.
Proof. exact_no_check radd_coordinatesW_correct_and_bounded. Time Qed.
+*)