From de0a4ce7d93437aa8229308cd06cd95f27f58809 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Thu, 17 Nov 2016 16:54:26 -0500 Subject: Update AddCoordinates Now the _correct_and_bounded lemma goes through --- src/SpecificGen/GF2213_32.v | 36 ++++++++++++++++++++++++++++++++++++ 1 file changed, 36 insertions(+) (limited to 'src/SpecificGen/GF2213_32.v') diff --git a/src/SpecificGen/GF2213_32.v b/src/SpecificGen/GF2213_32.v index 87da8dbcc..cd5ff40c5 100644 --- a/src/SpecificGen/GF2213_32.v +++ b/src/SpecificGen/GF2213_32.v @@ -176,6 +176,25 @@ Proof. etransitivity; apply app_n_correct. Qed. +Definition appify9 {T} (op : fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> T) (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe2213_32) := + app_n x0 (fun x0' => + app_n x1 (fun x1' => + app_n x2 (fun x2' => + app_n x3 (fun x3' => + app_n x4 (fun x4' => + app_n x5 (fun x5' => + app_n x6 (fun x6' => + app_n x7 (fun x7' => + app_n x8 (fun x8' => + op x0' x1' x2' x3' x4' x5' x6' x7' x8'))))))))). + +Lemma appify9_correct : forall {T} op x0 x1 x2 x3 x4 x5 x6 x7 x8, + @appify9 T op x0 x1 x2 x3 x4 x5 x6 x7 x8 = op x0 x1 x2 x3 x4 x5 x6 x7 x8. +Proof. + intros. cbv [appify9]. + repeat (etransitivity; [ apply app_n_correct | ]); reflexivity. +Qed. + Definition uncurry_unop_fe2213_32 {T} (op : fe2213_32 -> T) := Eval compute in Tuple.uncurry (n:=length_fe2213_32) op. Definition curry_unop_fe2213_32 {T} op : fe2213_32 -> T @@ -190,6 +209,23 @@ Definition uncurry_unop_wire_digits {T} (op : wire_digits -> T) Definition curry_unop_wire_digits {T} op : wire_digits -> T := Eval compute in fun f => app_n2 f (Tuple.curry (n:=length wire_widths) op). +Definition uncurry_9op_fe2213_32 {T} (op : fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> T) + := Eval compute in + uncurry_unop_fe2213_32 (fun x0 => + uncurry_unop_fe2213_32 (fun x1 => + uncurry_unop_fe2213_32 (fun x2 => + uncurry_unop_fe2213_32 (fun x3 => + uncurry_unop_fe2213_32 (fun x4 => + uncurry_unop_fe2213_32 (fun x5 => + uncurry_unop_fe2213_32 (fun x6 => + uncurry_unop_fe2213_32 (fun x7 => + uncurry_unop_fe2213_32 (fun x8 => + op x0 x1 x2 x3 x4 x5 x6 x7 x8))))))))). +Definition curry_9op_fe2213_32 {T} op : fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> fe2213_32 -> T + := Eval compute in + appify9 (fun x0 x1 x2 x3 x4 x5 x6 x7 x8 + => curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 (curry_unop_fe2213_32 op x0) x1) x2) x3) x4) x5) x6) x7) x8). + Definition add_sig (f g : fe2213_32) : { fg : fe2213_32 | fg = add_opt f g}. Proof. -- cgit v1.2.3