aboutsummaryrefslogtreecommitdiff
path: root/src/SpecificGen/GF25519_32Reflective/Common.v
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2016-11-25 13:02:06 -0500
committerGravatar Jason Gross <jgross@mit.edu>2016-11-25 13:02:06 -0500
commitfbdd082f2b7fab1fb5b074e193f1072d2ecdb2f7 (patch)
tree12d81977e0c454de17f63eb357144ca1fe31d9cb /src/SpecificGen/GF25519_32Reflective/Common.v
parentc51d5c4af858fb38946546b8d47f9d9dbb481cc0 (diff)
uncurry_n_op_fe25519
Diffstat (limited to 'src/SpecificGen/GF25519_32Reflective/Common.v')
-rw-r--r--src/SpecificGen/GF25519_32Reflective/Common.v48
1 files changed, 21 insertions, 27 deletions
diff --git a/src/SpecificGen/GF25519_32Reflective/Common.v b/src/SpecificGen/GF25519_32Reflective/Common.v
index 3389f908a..aad67f032 100644
--- a/src/SpecificGen/GF25519_32Reflective/Common.v
+++ b/src/SpecificGen/GF25519_32Reflective/Common.v
@@ -98,10 +98,10 @@ Local Ltac bounds_from_list_cps ls :=
constr:(fun T extra => (Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}, bs T extra))
end.
-Local Ltac make_bounds_cps ls :=
+Local Ltac make_bounds_cps ls extra :=
let v := bounds_from_list_cps (List.rev ls) in
let v := (eval compute in v) in
- pose v.
+ exact (v _ extra).
Local Ltac bounds_from_list ls :=
lazymatch (eval hnf in ls) with
@@ -116,36 +116,30 @@ Local Ltac make_bounds ls :=
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.
+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 := Expr_nm_Op_bounds n count_out in
- _
+ | 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).
- pose (p _ v).
- repeat (split; [ exact (fst p0) | ]; destruct p0 as [_ p0]).
- exact p0.
-*)
-Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ bounds ++ Tuple.to_list _ bounds)%list. Defined.
-Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ bounds). Defined.
-Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ bounds). Defined.
-Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type.
-Proof. make_bounds (Tuple.to_list _ bounds). Defined.
+ make_bounds_cps (Tuple.to_list _ bounds) v.
+Defined.
+Definition ExprBinOp_bounds : interp_all_binders_for ExprBinOpT ZBounds.interp_base_type
+ := Eval compute in Expr_nm_Op_bounds 2 1.
+Definition ExprUnOp_bounds : interp_all_binders_for ExprUnOpT ZBounds.interp_base_type
+ := Eval compute in Expr_nm_Op_bounds 1 1.
+Definition ExprUnOpFEToZ_bounds : interp_all_binders_for ExprUnOpFEToZT ZBounds.interp_base_type
+ := Eval compute in Expr_nm_Op_bounds 1 1.
+Definition ExprUnOpFEToWire_bounds : interp_all_binders_for ExprUnOpFEToWireT ZBounds.interp_base_type
+ := Eval compute in Expr_nm_Op_bounds 1 1.
+Definition Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type
+ := Eval compute in Expr_nm_Op_bounds 4 1.
+Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type
+ := 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 Expr4Op_bounds : interp_all_binders_for Expr4OpT ZBounds.interp_base_type.
-Proof.
- make_bounds (List.fold_right (@List.app _) nil (List.repeat (Tuple.to_list _ bounds) 4)).
-Defined.
-Definition Expr9Op_bounds : interp_all_binders_for Expr9_4OpT ZBounds.interp_base_type.
-Proof.
- make_bounds (List.fold_right (@List.app _) nil (List.repeat (Tuple.to_list _ bounds) 9)).
-Defined.
+
Definition interp_bexpr : ExprBinOp -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
:= fun e => curry_binop_fe25519_32W (Interp (@WordW.interp_op) e).
Definition interp_uexpr : ExprUnOp -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W -> SpecificGen.GF25519_32BoundedCommon.fe25519_32W
@@ -512,7 +506,7 @@ Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
repeat match goal with x := _ |- _ => subst x end;
cbv [id
binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
+ Relations.proj_eq_rel interp_flat_type_rel_pointwise2 SmartVarfMap interp_flat_type smart_interp_flat_map Application.all_binders_for fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Application.ApplyInterpedAll Application.ApplyInterpedAll' Application.interp_all_binders_for_of' Application.interp_all_binders_for_to' Application.fst_binder Application.snd_binder interp_flat_type_rel_pointwise2_gen_Prop Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower Application.remove_all_binders WordW.to_Z] in Hbounds_left, Hbounds_right;
lazymatch goal with
| [ |- fe25519_32WToZ ?x = _ /\ _ ]
=> generalize dependent x; intros