From 074dd72defb9df304175adf6e7d167ae7caea7bd Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Thu, 21 Jun 2018 16:49:53 -0400 Subject: WIP --- .../NewPipeline/AbstractInterpretation.v | 120 +- src/Experiments/NewPipeline/Arithmetic.v | 147 +- src/Experiments/NewPipeline/CLI.v | 74 +- src/Experiments/NewPipeline/CStringification.v | 1791 ++++++++++++-------- .../NewPipeline/GENERATEDIdentifiersWithoutTypes.v | 1257 ++++++++------ src/Experiments/NewPipeline/Language.v | 84 +- src/Experiments/NewPipeline/Rewriter.v | 232 ++- .../NewPipeline/SlowPrimeSynthesisExamples.v | 60 +- src/Experiments/NewPipeline/Toplevel1.v | 970 +++++++---- src/Experiments/NewPipeline/Toplevel2.v | 46 +- 10 files changed, 2985 insertions(+), 1796 deletions(-) (limited to 'src/Experiments/NewPipeline') diff --git a/src/Experiments/NewPipeline/AbstractInterpretation.v b/src/Experiments/NewPipeline/AbstractInterpretation.v index eada46692..55aa92ddc 100644 --- a/src/Experiments/NewPipeline/AbstractInterpretation.v +++ b/src/Experiments/NewPipeline/AbstractInterpretation.v @@ -3,6 +3,7 @@ Require Import Crypto.Util.ListUtil Coq.Lists.List Crypto.Util.ListUtil.FoldBool Require Import Crypto.Util.ZRange. Require Import Crypto.Util.ZRange.Operations. Require Import Crypto.Util.Option. +Require Import Crypto.Util.OptionList. Require Import Crypto.Util.ZUtil.Tactics.LtbToLt. Require Import Crypto.Util.LetIn. Require Import Crypto.Experiments.NewPipeline.Language. @@ -100,6 +101,19 @@ Module Compilers. => (@Some A a, @Some B b) | _ => fun _ => tt end. + Fixpoint lift_Some {t} : interp t -> option (base.interp t) + := match t with + | base.type.Z + | base.type.nat + | base.type.bool + => fun x => x + | base.type.unit + => fun x => Datatypes.Some tt + | base.type.list A + => fun ls => ls <- ls; ls <-- List.map (@lift_Some A) ls; Datatypes.Some ls + | base.type.prod A B + => fun '(a, b) => a <- @lift_Some A a; b <- @lift_Some B b; Datatypes.Some (a, b) + end%option. (** Keep data about list length and nat value, but not zrange *) Fixpoint strip_ranges {t} : interp t -> interp t := match t with @@ -257,12 +271,12 @@ Module Compilers. end. Fixpoint Some {t : type base.type} : type.interp t -> interp t := match t with - | type.base x =>@base.option.Some x + | type.base x => @base.option.Some x | type.arrow s d => fun _ _ => @None d end. Fixpoint strip_ranges {t : type base.type} : interp t -> interp t := match t with - | type.base x =>@base.option.strip_ranges x + | type.base x => @base.option.strip_ranges x | type.arrow s d => fun f x => @strip_ranges d (f x) end. Fixpoint is_tighter_than {t} : interp t -> interp t -> bool @@ -347,6 +361,8 @@ Module Compilers. => option_map (ident.interp idc) | ident.Z_of_nat as idc => option_map (fun n => r[Z.of_nat n~>Z.of_nat n]%zrange) + | ident.Z_to_nat as idc + => fun v => v <- to_literal v; Some (ident.interp idc v) | ident.List_length _ => option_map (@List.length _) | ident.Nat_max as idc @@ -385,13 +401,25 @@ Module Compilers. end | ident.Z_eqb as idc | ident.Z_leb as idc - | ident.Z_cc_m as idc + | ident.Z_geb as idc | ident.Z_pow as idc | ident.Z_modulo as idc => fun x y => match to_literal x, to_literal y with | Some x, Some y => of_literal (ident.interp idc x y) | _, _ => ZRange.type.base.option.None end + | ident.Z_bneg as idc + => fun x => match to_literal x with + | Some x => of_literal (ident.interp idc x) + | None => Datatypes.Some r[0~>1] + end + | ident.Z_lnot_modulo as idc + => fun v m + => match to_literal m, to_literal v with + | Some m, Some v => of_literal (ident.interp idc v m) + | Some m, None => Some r[0 ~> m] + | _, _ => None + end | ident.bool_rect _ => fun t f b => match b with @@ -447,20 +475,6 @@ Module Compilers. end | ident.List_update_nth _ => fun n f ls => ls <- ls; n <- n; Some (update_nth n f ls) - | ident.Z_mul_split as idc - | ident.Z_add_get_carry as idc - | ident.Z_sub_get_borrow as idc - => fun x y z => match to_literal x, to_literal y, to_literal z with - | Some x, Some y, Some z => of_literal (ident.interp idc x y z) - | _, _, _ => ZRange.type.base.option.None - end - | ident.Z_add_with_get_carry as idc - | ident.Z_sub_with_get_borrow as idc - | ident.Z_rshi as idc - => fun x y z w => match to_literal x, to_literal y, to_literal z, to_literal w with - | Some x, Some y, Some z, Some w => of_literal (ident.interp idc x y z w) - | _, _, _, _ => ZRange.type.base.option.None - end | ident.nil t => Some nil | ident.cons t => fun x => option_map (cons x) | ident.pair A B => pair @@ -473,70 +487,78 @@ Module Compilers. => fun ls1 ls2 => ls1 <- ls1; ls2 <- ls2; Some (List.app ls1 ls2) | ident.List_rev _ => option_map (@List.rev _) | ident.Z_opp as idc - | ident.Z_shiftr _ as idc - | ident.Z_shiftl _ as idc - | ident.Z_cc_m_concrete _ as idc + | ident.Z_log2 as idc + | ident.Z_log2_up as idc => fun x => x <- x; Some (ZRange.two_corners (ident.interp idc) x) | ident.Z_add as idc | ident.Z_mul as idc | ident.Z_sub as idc | ident.Z_div as idc - | ident.Z_rshi_concrete _ _ as idc + | ident.Z_shiftr as idc + | ident.Z_shiftl as idc => fun x y => x <- x; y <- y; Some (ZRange.four_corners (ident.interp idc) x y) | ident.Z_add_with_carry as idc => fun x y z => x <- x; y <- y; z <- z; Some (ZRange.eight_corners (ident.interp idc) x y z) - | ident.Z_land mask - => option_map (ZRange.land_bounds r[mask~>mask]) - | ident.Z_mul_split_concrete split_at - => fun x y - => match x, y with - | Some x, Some y + | ident.Z_cc_m as idc + => fun s x => s <- to_literal s; x <- x; Some (ZRange.two_corners (ident.interp idc s) x) + | ident.Z_rshi as idc + => fun s x y offset + => s <- to_literal s; x <- x; y <- y; offset <- to_literal offset; + Some (ZRange.four_corners (fun x' y' => ident.interp idc s x' y' offset) x y) + | ident.Z_land + => fun x y => x <- x; y <- y; Some (ZRange.land_bounds x y) + | ident.Z_lor + => fun x y => x <- x; y <- y; Some (ZRange.lor_bounds x y) + | ident.Z_mul_split + => fun split_at x y + => match to_literal split_at, x, y with + | Some split_at, Some x, Some y => ZRange.type.base.option.Some (t:=base.type.Z*base.type.Z) (ZRange.split_bounds (ZRange.four_corners Z.mul x y) split_at) - | _, _ => ZRange.type.base.option.None + | _, _, _ => ZRange.type.base.option.None end - | ident.Z_add_get_carry_concrete split_at - => fun x y - => match x, y with - | Some x, Some y + | ident.Z_add_get_carry + => fun split_at x y + => match to_literal split_at, x, y with + | Some split_at, Some x, Some y => ZRange.type.base.option.Some (t:=base.type.Z*base.type.Z) (ZRange.split_bounds (ZRange.four_corners Z.add x y) split_at) - | _, _ => ZRange.type.base.option.None + | _, _, _ => ZRange.type.base.option.None end - | ident.Z_add_with_get_carry_concrete split_at - => fun x y z - => match x, y, z with - | Some x, Some y, Some z + | ident.Z_add_with_get_carry + => fun split_at x y z + => match to_literal split_at, x, y, z with + | Some split_at, Some x, Some y, Some z => ZRange.type.base.option.Some (t:=base.type.Z*base.type.Z) (ZRange.split_bounds (ZRange.eight_corners (fun x y z => (x + y + z)%Z) x y z) split_at) - | _, _, _ => ZRange.type.base.option.None + | _, _, _, _ => ZRange.type.base.option.None end - | ident.Z_sub_get_borrow_concrete split_at - => fun x y - => match x, y with - | Some x, Some y + | ident.Z_sub_get_borrow + => fun split_at x y + => match to_literal split_at, x, y with + | Some split_at, Some x, Some y => ZRange.type.base.option.Some (t:=base.type.Z*base.type.Z) (let b := ZRange.split_bounds (ZRange.four_corners BinInt.Z.sub x y) split_at in (* N.B. sub_get_borrow returns - ((x - y) / split_at) as the borrow, so we need to negate *) (fst b, ZRange.opp (snd b))) - | _, _ => ZRange.type.base.option.None + | _, _, _ => ZRange.type.base.option.None end - | ident.Z_sub_with_get_borrow_concrete split_at - => fun x y z - => match x, y, z with - | Some x, Some y, Some z + | ident.Z_sub_with_get_borrow + => fun split_at x y z + => match to_literal split_at, x, y, z with + | Some split_at, Some x, Some y, Some z => ZRange.type.base.option.Some (t:=base.type.Z*base.type.Z) (let b := ZRange.split_bounds (ZRange.eight_corners (fun x y z => (y - z - x)%Z) x y z) split_at in (* N.B. sub_get_borrow returns - ((x - y) / split_at) as the borrow, so we need to negate *) (fst b, ZRange.opp (snd b))) - | _, _, _ => ZRange.type.base.option.None + | _, _, _, _ => ZRange.type.base.option.None end | ident.Z_zselect => fun _ y z => y <- y; z <- z; Some (ZRange.union y z) @@ -549,8 +571,6 @@ Module Compilers. (ZRange.four_corners Z.add x y) (ZRange.eight_corners (fun x y m => Z.max 0 (x + y - m)) x y m))) - | ident.Z_neg_snd (** TODO(jadep): This is only here for demonstration purposes; remove it once you no longer need it as a template *) - => fun '(a, b) => (a, option_map ZRange.opp b) | ident.Z_cast range => fun r : option zrange => Some match r with diff --git a/src/Experiments/NewPipeline/Arithmetic.v b/src/Experiments/NewPipeline/Arithmetic.v index d263ca151..04f4bdd4d 100644 --- a/src/Experiments/NewPipeline/Arithmetic.v +++ b/src/Experiments/NewPipeline/Arithmetic.v @@ -2079,7 +2079,7 @@ Module Freeze. Definition freeze n mask (m p:list Z) : list Z := let '(p, carry) := Rows.sub weight n p m in - let '(r, carry) := Rows.conditional_add weight n mask carry p m in + let '(r, carry) := Rows.conditional_add weight n mask (-carry) p m in r. Lemma freezeZ m s c y : @@ -2087,13 +2087,14 @@ Module Freeze. 0 < c < s -> s <> 0 -> 0 <= y < 2*m -> - ((y - m) + (if (dec ((y - m) / s = 0)) then 0 else m)) mod s + ((y - m) + (if (dec (-((y - m) / s) = 0)) then 0 else m)) mod s = y mod m. Proof using Type. clear; intros. transitivity ((y - m) mod m); repeat first [ progress intros | progress subst + | rewrite Z.opp_eq_0_iff in * | break_innermost_match_step | progress autorewrite with zsimplify_fast | rewrite Z.div_small_iff in * by auto @@ -2117,7 +2118,7 @@ Module Freeze. (Hmlen : length m = n) : Positional.eval weight n (@freeze n mask m p) = (Positional.eval weight n p - Positional.eval weight n m + - (if dec ((Positional.eval weight n p - Positional.eval weight n m) / weight n = 0) then 0 else Positional.eval weight n m)) + (if dec (-((Positional.eval weight n p - Positional.eval weight n m) / weight n) = 0) then 0 else Positional.eval weight n m)) mod weight n. (*if dec ((Positional.eval weight n p - Positional.eval weight n m) / weight n = 0) then Positional.eval weight n p - Positional.eval weight n m @@ -2144,13 +2145,13 @@ Module Freeze. (n_nonzero:n<>0%nat) (Hc : 0 < Associational.eval c < weight n) (Hmask : List.map (Z.land mask) m = m) - modulus (Hm : Positional.eval weight n m = Z.pos modulus) - (Hp : 0 <= Positional.eval weight n p < 2*(Z.pos modulus)) - (Hsc : Z.pos modulus = weight n - Associational.eval c) + (modulus:=weight n - Associational.eval c) + (Hm : Positional.eval weight n m = modulus) + (Hp : 0 <= Positional.eval weight n p < 2*modulus) (Hplen : length p = n) (Hmlen : length m = n) : Positional.eval weight n (@freeze n mask m p) - = Positional.eval weight n p mod (Z.pos modulus). + = Positional.eval weight n p mod modulus. Proof using wprops. pose proof (@weight_positive weight wprops n). rewrite eval_freeze_eq by assumption. @@ -2162,16 +2163,17 @@ Module Freeze. (n_nonzero:n<>0%nat) (Hc : 0 < Associational.eval c < weight n) (Hmask : List.map (Z.land mask) m = m) - modulus (Hm : Positional.eval weight n m = Z.pos modulus) - (Hp : 0 <= Positional.eval weight n p < 2*(Z.pos modulus)) - (Hsc : Z.pos modulus = weight n - Associational.eval c) + (modulus:=weight n - Associational.eval c) + (Hm : Positional.eval weight n m = modulus) + (Hp : 0 <= Positional.eval weight n p < 2*modulus) (Hplen : length p = n) (Hmlen : length m = n) - : @freeze n mask m p = Rows.partition weight n (Positional.eval weight n p mod (Z.pos modulus)). + : @freeze n mask m p = Rows.partition weight n (Positional.eval weight n p mod modulus). Proof using wprops. pose proof (@weight_positive weight wprops n). pose proof (fun v => Z.mod_pos_bound v (weight n) ltac:(lia)). - pose proof (Z.mod_pos_bound (Positional.eval weight n p) (Z.pos modulus) ltac:(lia)). + pose proof (Z.mod_pos_bound (Positional.eval weight n p) modulus ltac:(lia)). + subst modulus. erewrite <- eval_freeze by eassumption. cbv [freeze]; eta_expand. rewrite Rows.conditional_add_partitions by (auto; rewrite Rows.sub_partitions; auto; distr_length). @@ -2182,3 +2184,124 @@ Module Freeze. Qed. End Freeze. End Freeze. +Hint Rewrite Freeze.length_freeze : distr_length. + +Section freeze_mod_ops. + Import Positional. + Import Freeze. + Local Coercion Z.of_nat : nat >-> Z. + Local Coercion QArith_base.inject_Z : Z >-> Q. + (* Design constraints: + - inputs must be [Z] (b/c reification does not support Q) + - internal structure must not match on the arguments (b/c reification does not support [positive]) *) + Context (limbwidth_num limbwidth_den : Z) + (limbwidth_good : 0 < limbwidth_den <= limbwidth_num) + (s : Z) + (c : list (Z*Z)) + (n : nat) + (bitwidth : Z) + (m_enc : list Z) + (m_nz:s - Associational.eval c <> 0) (s_nz:s <> 0) + (Hn_nz : n <> 0%nat). + Local Notation bytes_weight := (@weight 8 1). + Local Notation weight := (@weight limbwidth_num limbwidth_den). + Let m := (s - Associational.eval c). + + Context (Hs : s = weight n). + Context (c_small : 0 < Associational.eval c < weight n) + (m_enc_bounded : List.map (BinInt.Z.land (Z.ones bitwidth)) m_enc = m_enc) + (m_enc_correct : Positional.eval weight n m_enc = m) + (Hm_enc_len : length m_enc = n). + + Definition wprops_bytes := (@wprops 8 1 ltac:(lia)). + Local Notation wprops := (@wprops limbwidth_num limbwidth_den limbwidth_good). + + Local Hint Immediate (weight_0 wprops). + Local Hint Immediate (weight_positive wprops). + Local Hint Immediate (weight_multiples wprops). + Local Hint Immediate (weight_divides wprops). + Local Hint Immediate (weight_0 wprops_bytes). + Local Hint Immediate (weight_positive wprops_bytes). + Local Hint Immediate (weight_multiples wprops_bytes). + Local Hint Immediate (weight_divides wprops_bytes). + Local Hint Resolve Z.positive_is_nonzero Z.lt_gt. + + Definition bytes_n := (1 + (Z.to_nat (Z.log2_up (weight n) / 8)))%nat. + + Definition to_bytes' (v : list Z) + := BaseConversion.convert_bases weight bytes_weight n bytes_n v. + + Definition from_bytes (v : list Z) + := BaseConversion.convert_bases bytes_weight weight bytes_n n v. + + Definition to_bytesmod (f : list Z) : list Z + := to_bytes' (freeze weight n (Z.ones bitwidth) m_enc f). + + Definition from_bytesmod (f : list Z) : list Z + := from_bytes f. + + Lemma eval_to_bytesmod + : forall (f : list Z) + (Hf : length f = n) + (Hf_bounded : 0 <= eval weight n f < 2 * m), + (eval bytes_weight bytes_n (to_bytesmod f)) = (eval weight n f) mod m + /\ to_bytesmod f = to_bytes' (Rows.partition weight n (Positional.eval weight n f mod m)). + Proof. + intros; subst m s; split. + { erewrite <- eval_freeze with (weight := weight) (n:=n) (mask:=Z.ones bitwidth) (m:=m_enc) ; auto using wprops. + erewrite <- BaseConversion.eval_convert_bases with (sw:=weight) (dw:=bytes_weight) (sn:=n) (dn:=bytes_n) (p:=freeze _ _ _ _ _) + by (cbv [bytes_n]; auto using wprops_bytes; distr_length; auto using wprops). + reflexivity. } + { cbv [to_bytesmod]. + erewrite freeze_partitions by eauto using wprops. + reflexivity. } + Qed. + + Lemma eval_from_bytesmod + : forall (f : list Z) + (Hf : length f = bytes_n), + eval weight n (from_bytesmod f) = eval bytes_weight bytes_n f. + Proof. + cbv [from_bytesmod from_bytes]; intros. + rewrite BaseConversion.eval_convert_bases by eauto using wprops. + reflexivity. + Qed. +End freeze_mod_ops. + +Section primitives. + Definition mulx (bitwidth : Z) := Eval cbv [Z.mul_split_at_bitwidth] in Z.mul_split_at_bitwidth bitwidth. + Definition addcarryx (bitwidth : Z) := Eval cbv [Z.add_with_get_carry Z.add_with_carry Z.get_carry] in Z.add_with_get_carry bitwidth. + Definition subborrowx (bitwidth : Z) := Eval cbv [Z.sub_with_get_borrow Z.sub_with_borrow Z.get_borrow Z.get_carry Z.add_with_carry] in Z.sub_with_get_borrow bitwidth. + Definition cmovznz (bitwidth : Z) (cond : Z) (z nz : Z) + := dlet t := (0 - Z.bneg (Z.bneg cond)) mod 2^bitwidth in Z.lor (Z.land t nz) (Z.land (Z.lnot_modulo t (2^bitwidth)) z). + + Lemma cmovznz_correct bitwidth cond z nz + : 0 <= z < 2^bitwidth + -> 0 <= nz < 2^bitwidth + -> cmovznz bitwidth cond z nz = Z.zselect cond z nz. + Proof. + intros. + assert (0 < 2^bitwidth) by omega. + assert (0 <= bitwidth) by auto with zarith. + assert (0 < bitwidth -> 1 < 2^bitwidth) by auto with zarith. + pose proof Z.log2_lt_pow2_alt. + assert (bitwidth = 0 \/ 0 < bitwidth) by omega. + repeat first [ progress cbv [cmovznz Z.zselect Z.bneg Let_In Z.lnot_modulo] + | progress split_iff + | progress subst + | progress Z.ltb_to_lt + | progress destruct_head'_or + | congruence + | omega + | progress break_innermost_match_step + | progress break_innermost_match_hyps_step + | progress autorewrite with zsimplify_const in * + | progress pull_Zmod + | progress intros + | rewrite !Z.sub_1_r, <- Z.ones_equiv, <- ?Z.sub_1_r + | rewrite Z_mod_nz_opp_full by (Z.rewrite_mod_small; omega) + | rewrite (Z.land_comm (Z.ones _)) + | rewrite Z.land_ones_low by auto with omega + | progress Z.rewrite_mod_small ]. + Qed. +End primitives. diff --git a/src/Experiments/NewPipeline/CLI.v b/src/Experiments/NewPipeline/CLI.v index ba240d093..ebe1fcca9 100644 --- a/src/Experiments/NewPipeline/CLI.v +++ b/src/Experiments/NewPipeline/CLI.v @@ -106,12 +106,14 @@ Module ForExtraction. Module UnsaturatedSolinas. Definition PipelineLines + (curve_description : string) (n : string) (s : string) (c : string) (machine_wordsize : string) : list (string * Pipeline.ErrorT (list string)) + list string - := let str_n := n in + := let prefix := ("fiat_" ++ curve_description ++ "_")%string in + let str_n := n in let n : nat := parse_n n in let str_machine_wordsize := machine_wordsize in let str_c := c in @@ -125,30 +127,34 @@ Module ForExtraction. | _, None => inr ["Could not parse c (" ++ c ++ ")"] | Some s, Some c - => let header := - ((["/* Autogenerated */"; + => let '(res, types_used) + := UnsaturatedSolinas.Synthesize n s c machine_wordsize prefix in + let header := + ((["/* Autogenerated */"; + "/* curve description: " ++ curve_description ++ " */"; "/* n = " ++ show false n ++ " (from """ ++ str_n ++ """) */"; "/* s = " ++ Hex.show_Z false s ++ " (from """ ++ str_s ++ """) */"; "/* c = " ++ show false c ++ " (from """ ++ str_c ++ """) */"; "/* machine_wordsize = " ++ show false machine_wordsize ++ " (from """ ++ str_machine_wordsize ++ """) */"; ""]%string) - ++ ToString.C.String.typedef_header + ++ ToString.C.String.typedef_header prefix types_used ++ [""])%list in inl ([("check_args" ++ NewLine ++ String.concat NewLine header, UnsaturatedSolinas.check_args n s c machine_wordsize (ErrorT.Success header))%string] - ++ UnsaturatedSolinas.Synthesize n s c machine_wordsize "fe")%list + ++ res)%list end. Definition ProcessedLines + (curve_description : string) (n : string) (s : string) (c : string) (machine_wordsize : string) : list string + list string - := match CollectErrors (PipelineLines n s c machine_wordsize) with + := match CollectErrors (PipelineLines curve_description n s c machine_wordsize) with | inl ls => inl (List.map (fun s => String.concat NewLine s ++ NewLine ++ NewLine) @@ -162,6 +168,7 @@ Module ForExtraction. Definition Pipeline {A} + (curve_description : string) (n : string) (s : string) (c : string) @@ -169,7 +176,7 @@ Module ForExtraction. (success : list string -> A) (error : list string -> A) : A - := match ProcessedLines n s c machine_wordsize with + := match ProcessedLines curve_description n s c machine_wordsize with | inl s => success s | inr s => error s end. @@ -181,24 +188,26 @@ Module ForExtraction. (error : list string -> A) : A := match argv with - | _::n::s::c::machine_wordsize::nil + | _::curve_description::n::s::c::machine_wordsize::nil => Pipeline - n s c machine_wordsize + curve_description n s c machine_wordsize success error | nil => error ["empty argv"] | prog::args - => error ["Expected arguments n, s, c, machine_wordsize, got " ++ show false (List.length args) ++ " arguments in " ++ prog] + => error ["Expected arguments curve_description, n, s, c, machine_wordsize, got " ++ show false (List.length args) ++ " arguments in " ++ prog] end. End UnsaturatedSolinas. Module SaturatedSolinas. Definition PipelineLines + (curve_description : string) (s : string) (c : string) (machine_wordsize : string) : list (string * Pipeline.ErrorT (list string)) + list string - := let str_machine_wordsize := machine_wordsize in + := let prefix := ("fiat_" ++ curve_description ++ "_")%string in + let str_machine_wordsize := machine_wordsize in let str_c := c in let str_s := s in let machine_wordsize := parse_machine_wordsize machine_wordsize in @@ -210,28 +219,32 @@ Module ForExtraction. | _, None => inr ["Could not parse c (" ++ c ++ ")"] | Some s, Some c - => let header := - ((["/* Autogenerated */"; - "/* s = " ++ Hex.show_Z false s ++ " (from """ ++ str_s ++ """) */"; - "/* c = " ++ show false c ++ " (from """ ++ str_c ++ """) */"; - "/* machine_wordsize = " ++ show false machine_wordsize ++ " (from """ ++ str_machine_wordsize ++ """) */"; - ""]%string) - ++ ToString.C.String.typedef_header - ++ [""])%list in - inl - ([("check_args" ++ NewLine ++ String.concat NewLine header, - SaturatedSolinas.check_args - s c machine_wordsize - (ErrorT.Success header))%string] - ++ SaturatedSolinas.Synthesize s c machine_wordsize "fe")%list + => let '(res, types_used) + := SaturatedSolinas.Synthesize s c machine_wordsize prefix in + let header := + ((["/* Autogenerated */"; + "/* curve description: " ++ curve_description ++ " */"; + "/* s = " ++ Hex.show_Z false s ++ " (from """ ++ str_s ++ """) */"; + "/* c = " ++ show false c ++ " (from """ ++ str_c ++ """) */"; + "/* machine_wordsize = " ++ show false machine_wordsize ++ " (from """ ++ str_machine_wordsize ++ """) */"; + ""]%string) + ++ ToString.C.String.typedef_header prefix types_used + ++ [""])%list in + inl + ([("check_args" ++ NewLine ++ String.concat NewLine header, + SaturatedSolinas.check_args + s c machine_wordsize + (ErrorT.Success header))%string] + ++ res)%list end. Definition ProcessedLines + (curve_description : string) (s : string) (c : string) (machine_wordsize : string) : list string + list string - := match CollectErrors (PipelineLines s c machine_wordsize) with + := match CollectErrors (PipelineLines curve_description s c machine_wordsize) with | inl ls => inl (List.map (fun s => String.concat NewLine s ++ NewLine ++ NewLine) @@ -245,13 +258,14 @@ Module ForExtraction. Definition Pipeline {A} + (curve_description : string) (s : string) (c : string) (machine_wordsize : string) (success : list string -> A) (error : list string -> A) : A - := match ProcessedLines s c machine_wordsize with + := match ProcessedLines curve_description s c machine_wordsize with | inl s => success s | inr s => error s end. @@ -263,14 +277,14 @@ Module ForExtraction. (error : list string -> A) : A := match argv with - | _::s::c::machine_wordsize::nil + | _::curve_description::s::c::machine_wordsize::nil => Pipeline - s c machine_wordsize + curve_description s c machine_wordsize success error | nil => error ["empty argv"] | prog::args - => error ["Expected arguments s, c, machine_wordsize, got " ++ show false (List.length args) ++ " arguments in " ++ prog] + => error ["Expected arguments curve_description, s, c, machine_wordsize, got " ++ show false (List.length args) ++ " arguments in " ++ prog] end. End SaturatedSolinas. End ForExtraction. diff --git a/src/Experiments/NewPipeline/CStringification.v b/src/Experiments/NewPipeline/CStringification.v index b880f665f..1cf7ee0fc 100644 --- a/src/Experiments/NewPipeline/CStringification.v +++ b/src/Experiments/NewPipeline/CStringification.v @@ -13,6 +13,7 @@ Require Import Crypto.Util.ZRange. Require Import Crypto.Util.ZRange.Operations. Require Import Crypto.Util.ZRange.Show. Require Import Crypto.Util.Option. +Require Import Crypto.Util.OptionList. Require Import Crypto.Experiments.NewPipeline.Language. Require Import Crypto.Experiments.NewPipeline.AbstractInterpretation. Require Import Crypto.Util.Bool.Equality. @@ -30,6 +31,55 @@ Module Compilers. Module ToString. Local Open Scope string_scope. + Module Import ZRange. + Module Export type. + Module Export base. + Fixpoint show_interp {t} : Show (ZRange.type.base.interp t) + := match t return bool -> ZRange.type.base.interp t -> string with + | base.type.unit => @show unit _ + | base.type.Z => @show zrange _ + | base.type.bool => @show bool _ + | base.type.nat => @show nat _ + | base.type.prod A B + => fun _ '(a, b) + => "(" ++ @show_interp A false a ++ ", " ++ @show_interp B true b ++ ")" + | base.type.list A + => let SA := @show_interp A in + @show (list (ZRange.type.base.interp A)) _ + end%string. + Global Existing Instance show_interp. + Module Export option. + Fixpoint show_interp {t} : Show (ZRange.type.base.option.interp t) + := match t return bool -> ZRange.type.base.option.interp t -> string with + | base.type.unit => @show unit _ + | base.type.Z => @show (option zrange) _ + | base.type.bool => @show (option bool) _ + | base.type.nat => @show (option nat) _ + | base.type.prod A B + => let SA := @show_interp A in + let SB := @show_interp B in + @show (ZRange.type.base.option.interp A * ZRange.type.base.option.interp B) _ + | base.type.list A + => let SA := @show_interp A in + @show (option (list (ZRange.type.base.option.interp A))) _ + end. + Global Existing Instance show_interp. + End option. + End base. + + Module option. + Global Instance show_interp {t} : Show (ZRange.type.option.interp t) + := fun parens + => match t return ZRange.type.option.interp t -> string with + | type.base t + => fun v : ZRange.type.base.option.interp t + => show parens v + | type.arrow s d => fun _ => "λ" + end. + End option. + End type. + End ZRange. + Module PHOAS. Module type. Module base. @@ -65,6 +115,17 @@ Module Compilers. Global Existing Instance show_interp. Global Instance show : Show base.type := show_type. End base. + Fixpoint show_for_each_lhs_of_arrow {base_type} {f : type.type base_type -> Type} {show_f : forall t, Show (f t)} {t : type.type base_type} : Show (type.for_each_lhs_of_arrow f t) + := match t return bool -> type.for_each_lhs_of_arrow f t -> string with + | type.base t => @show unit _ + | type.arrow s d + => fun with_parens '((x, xs) : f s * type.for_each_lhs_of_arrow f d) + => let _ : Show (f s) := show_f s in + let _ : Show (type.for_each_lhs_of_arrow f d) := @show_for_each_lhs_of_arrow base_type f show_f d in + show with_parens (x, xs) + end. + Global Existing Instance show_for_each_lhs_of_arrow. + Fixpoint show_type {base_type} {S : Show base_type} (with_parens : bool) (t : type.type base_type) : string := match t with | type.base t => S with_parens t @@ -139,28 +200,27 @@ Module Compilers. | ident.Z_modulo => "(mod)" | ident.Z_eqb => "(=)" | ident.Z_leb => "(≤)" + | ident.Z_geb => "(≥)" + | ident.Z_log2 => "log₂" + | ident.Z_log2_up => "⌈log₂⌉" | ident.Z_of_nat => "(ℕ→ℤ)" - | ident.Z_shiftr offset => "(>> " ++ decimal_string_of_Z offset ++ ")" - | ident.Z_shiftl offset => "(<< " ++ decimal_string_of_Z offset ++ ")" - | ident.Z_land mask => "(& " ++ HexString.of_Z mask ++ ")" + | ident.Z_to_nat => "(ℤ→ℕ)" + | ident.Z_shiftr => "(>>)" + | ident.Z_shiftl => "(<<)" + | ident.Z_land => "(&)" + | ident.Z_lor => "(|)" + | ident.Z_lnot_modulo => "~" + | ident.Z_bneg => "!" | ident.Z_mul_split => "Z.mul_split" - | ident.Z_mul_split_concrete s => maybe_wrap_parens with_parens ("Z.mul_split " ++ bitwidth_to_string s) | ident.Z_add_get_carry => "Z.add_get_carry" - | ident.Z_add_get_carry_concrete s => maybe_wrap_parens with_parens ("Z.add_get_carry " ++ bitwidth_to_string s) | ident.Z_add_with_carry => "Z.add_with_carry" | ident.Z_add_with_get_carry => "Z.add_with_get_carry" - | ident.Z_add_with_get_carry_concrete s => maybe_wrap_parens with_parens ("Z.add_with_get_carry " ++ bitwidth_to_string s) | ident.Z_sub_get_borrow => "Z.sub_get_borrow" - | ident.Z_sub_get_borrow_concrete s => maybe_wrap_parens with_parens ("Z.sub_get_borrow " ++ bitwidth_to_string s) | ident.Z_sub_with_get_borrow => "Z.sub_with_get_borrow" - | ident.Z_sub_with_get_borrow_concrete s => maybe_wrap_parens with_parens ("Z.sub_with_get_borrow " ++ bitwidth_to_string s) | ident.Z_zselect => "Z.zselect" | ident.Z_add_modulo => "Z.add_modulo" | ident.Z_rshi => "Z.rshi" - | ident.Z_rshi_concrete s offset => maybe_wrap_parens with_parens ("Z.rshi " ++ bitwidth_to_string s ++ " " ++ decimal_string_of_Z offset) | ident.Z_cc_m => "Z.cc_m" - | ident.Z_cc_m_concrete s => maybe_wrap_parens with_parens ("Z.cc_m " ++ bitwidth_to_string s) - | ident.Z_neg_snd => "Z.neg_snd" | ident.Z_cast range => "(" ++ show_range_or_ctype range ++ ")" | ident.Z_cast2 (r1, r2) => "(" ++ show_range_or_ctype r1 ++ ", " ++ show_range_or_ctype r2 ++ ")" | ident.fancy_add log2wordmax imm @@ -235,6 +295,20 @@ Module Compilers. ++ show_f end%list) end. + Fixpoint show_var_expr {var t} (with_parens : bool) (e : @expr.expr base_type ident var t) : string + := match e with + | expr.Ident t idc => show with_parens idc + | expr.Var t v => maybe_wrap_parens with_parens ("VAR_(" ++ show false t ++ ")") + | expr.Abs s d f => "λ_(" ++ show false t ++ ")" + | expr.App s d f x + => let show_f := @show_var_expr _ _ false f in + let show_x := @show_var_expr _ _ true x in + maybe_wrap_parens with_parens (show_f ++ " @ " ++ show_x) + | expr.LetIn A B x f + => let show_x := @show_var_expr _ _ false x in + maybe_wrap_parens with_parens ("expr_let _ := " ++ show_x ++ " in _") + end%string. + Definition partially_show_expr {var t} : Show (@expr.expr base_type ident var t) := show_var_expr. Global Instance show_lines_expr {t} : ShowLines (@expr.expr base_type ident (fun _ => string) t) := fun with_parens e => snd (@show_expr_lines t e 1%positive with_parens). Global Instance show_lines_Expr {t} : ShowLines (@expr.Expr base_type ident t) @@ -299,6 +373,12 @@ Module Compilers. Definition unsigned_counterpart_of (t : type) : type := unsigned (lgbitwidth_of t). + Global Instance show_type : Show type + := fun with_parens t + => maybe_wrap_parens + with_parens + ((if is_unsigned t then "u" else "") ++ "int" ++ decimal_string_of_Z (bitwidth_of t)). + Definition union (t1 t2 : type) : type := of_zrange_relaxed (ZRange.union (to_zrange t1) (to_zrange t2)). Fixpoint base_interp (t : base.type) : Set @@ -364,17 +444,18 @@ Module Compilers. | Dereference : ident Zptr Z | Z_shiftr (offset : BinInt.Z) : ident Z Z | Z_shiftl (offset : BinInt.Z) : ident Z Z - | Z_land (mask : BinInt.Z) : ident Z Z + | Z_land : ident (Z * Z) Z + | Z_lor : ident (Z * Z) Z | Z_add : ident (Z * Z) Z | Z_mul : ident (Z * Z) Z | Z_sub : ident (Z * Z) Z | Z_opp : ident Z Z - | Z_mul_split (lgs:BinInt.Z) : ident (Z * Z * Zptr) Z - | Z_add_get_carry (lgs:BinInt.Z) : ident (Z * Z * Zptr) Z - | Z_add_with_get_carry (lgs:BinInt.Z) : ident (Z * Z * Z * Zptr) Z - | Z_sub_get_borrow (lgs:BinInt.Z) : ident (Z * Z * Zptr) Z - | Z_sub_with_get_borrow (lgs:BinInt.Z) : ident (Z * Z * Z * Zptr) Z - | Z_zselect : ident (Z * Z * Z) Z + | Z_lnot (ty:int.type) : ident Z Z + | Z_bneg : ident Z Z + | Z_mul_split (lgs:BinInt.Z) : ident ((Zptr * Zptr) * (Z * Z)) unit + | Z_add_with_get_carry (lgs:BinInt.Z) : ident ((Zptr * Zptr) * (Z * Z * Z)) unit + | Z_sub_with_get_borrow (lgs:BinInt.Z) : ident ((Zptr * Zptr) * (Z * Z * Z)) unit + | Z_zselect (ty:int.type) : ident (Zptr * (Z * Z * Z)) unit | Z_add_modulo : ident (Z * Z * Z) Z | Z_static_cast (ty : int.type) : ident Z Z . @@ -387,6 +468,7 @@ Module Compilers. | TT : arith_expr type.unit. Inductive stmt := + | Call (val : arith_expr type.unit) | Assign (declare : bool) (t : type.primitive) (sz : option int.type) (name : string) (val : arith_expr t) | AssignZPtr (name : string) (sz : option int.type) (val : arith_expr type.Z) | DeclareVar (t : type.primitive) (sz : option int.type) (name : string) @@ -408,6 +490,12 @@ Module Compilers. Notation "x ;; y" := (@cons stmt x%Cexpr y%Cexpr) (at level 70, right associativity, format "'[v' x ;; '/' y ']'") : Cexpr_scope. End Notations. + Definition invert_literal {t} (e : arith_expr t) : option (BinInt.Z) + := match e with + | AppIdent s d (literal v) arg => Some v + | _ => None + end. + Module OfPHOAS. Fixpoint base_var_data (t : base.type) : Set := match t with @@ -628,14 +716,14 @@ Module Compilers. (s:=(base.type.Z * base.type.Z)%etype) (d:=base.type.Z) (idc : ident (type.Z * type.Z) type.Z) - : option int.type -> arith_expr_for s -> option (arith_expr_for d) + : option int.type -> arith_expr_for s -> arith_expr_for d := fun desired_type '((e1, t1), (e2, t2)) => let t1 := option_map integer_promote_type t1 in let t2 := option_map integer_promote_type t2 in let '((e1, t1), (e2, t2)) := cast_bigger_up_if_needed desired_type ((e1, t1), (e2, t2)) in let ct := (t1 <- t1; t2 <- t2; Some (common_type t1 t2))%option in - Some (Zcast_down_if_needed desired_type ((idc @@ (e1, e2))%Cexpr, ct)). + Zcast_down_if_needed desired_type ((idc @@ (e1, e2))%Cexpr, ct). Local Definition fakeprod (A B : Compilers.type.type base.type) : Compilers.type.type base.type := match A, B with @@ -650,598 +738,811 @@ Module Compilers. | type.arrow s d => arith_expr_for (type.uncurried_domain fakeprod s d) end. - Fixpoint arith_expr_of_PHOAS_ident - {t} - (idc : ident.ident t) - : int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => option T) arith_expr_for_base t - := match idc in ident.ident t return int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => option T) arith_expr_for_base t with - | ident.Literal base.type.Z v - => fun r => Some (cast_down_if_needed - r - (literal v @@ TT, Some (int.of_zrange_relaxed r[v~>v]))) - | ident.nil t - => fun _ => Some nil - | ident.cons t - => fun r x xs => Some (cast_down_if_needed r (cons x xs)) - | ident.fst A B => fun r xy => Some (cast_down_if_needed r (@fst _ _ xy)) - | ident.snd A B => fun r xy => Some (cast_down_if_needed r (@snd _ _ xy)) - | ident.List_nth_default base.type.Z - => fun r d ls n - => List.nth_default None (List.map (fun x => Some (cast_down_if_needed r x)) ls) n - | ident.Z_shiftr offset - => fun rout '(e, r) - => let rin := option_map integer_promote_type r in - Some (cast_down_if_needed rout (Z_shiftr offset @@ e, rin)) - | ident.Z_shiftl offset - => fun rout '(e, r) - => let rin := option_map integer_promote_type r in - let '(e', rin') := cast_up_if_needed rout (e, rin) in - Some (cast_down_if_needed rout (Z_shiftl offset @@ e', rin')) - | ident.Z_land mask - => fun rout '(e, r) - => Some (cast_down_if_needed - rout - (Z_land mask @@ e, - option_map integer_promote_type r)) - | ident.Z_add => fun r x y => arith_bin_arith_expr_of_PHOAS_ident Z_add r (x, y) - | ident.Z_mul => fun r x y => arith_bin_arith_expr_of_PHOAS_ident Z_mul r (x, y) - | ident.Z_sub => fun r x y => arith_bin_arith_expr_of_PHOAS_ident Z_sub r (x, y) - | ident.Z_zselect - => fun rout '(econd, _) '(e1, r1) '(e2, r2) - => let r1 := option_map integer_promote_type r1 in - let r2 := option_map integer_promote_type r2 in - let '((e1, r1), (e2, r2)) - := cast_bigger_up_if_needed rout ((e1, r1), (e2, r2)) in - let ct := (r1 <- r1; r2 <- r2; Some (common_type r1 r2))%option in - Some (cast_down_if_needed rout ((Z_zselect @@ (econd, e1, e2))%Cexpr, ct)) - | ident.pair A B - => fun _ _ _ => None - | ident.Z_opp - => fun _ _ => None - | ident.Literal _ v - => fun _ => Some v - | ident.Nat_succ - | ident.Nat_pred - | ident.Nat_max - | ident.Nat_mul - | ident.Nat_add - | ident.Nat_sub - | ident.pair_rect _ _ _ - | ident.bool_rect _ - | ident.nat_rect _ - | ident.list_rect _ _ - | ident.list_case _ _ - | ident.List_length _ - | ident.List_seq - | ident.List_repeat _ - | ident.List_combine _ _ - | ident.List_map _ _ - | ident.List_app _ - | ident.List_rev _ - | ident.List_flat_map _ _ - | ident.List_partition _ - | ident.List_fold_right _ _ - | ident.List_update_nth _ - | ident.List_nth_default _ - | ident.Z_pow - | ident.Z_div - | ident.Z_modulo - | ident.Z_eqb - | ident.Z_leb - | ident.Z_of_nat - | ident.Z_mul_split - | ident.Z_mul_split_concrete _ - | ident.Z_add_get_carry - | ident.Z_add_get_carry_concrete _ - | ident.Z_add_with_carry - | ident.Z_add_with_get_carry - | ident.Z_add_with_get_carry_concrete _ - | ident.Z_sub_get_borrow - | ident.Z_sub_get_borrow_concrete _ - | ident.Z_sub_with_get_borrow - | ident.Z_sub_with_get_borrow_concrete _ - | ident.Z_add_modulo - | ident.Z_rshi - | ident.Z_rshi_concrete _ _ - | ident.Z_cc_m - | ident.Z_cc_m_concrete _ - | ident.Z_neg_snd - | ident.Z_cast _ - | ident.Z_cast2 _ - | ident.fancy_add _ _ - | ident.fancy_addc _ _ - | ident.fancy_sub _ _ - | ident.fancy_subb _ _ - | ident.fancy_mulll _ - | ident.fancy_mullh _ - | ident.fancy_mulhl _ - | ident.fancy_mulhh _ - | ident.fancy_rshi _ _ - | ident.fancy_selc - | ident.fancy_selm _ - | ident.fancy_sell - | ident.fancy_addm - => fun _ => type.interpM_return _ _ _ None - end%core%Cexpr%option%zrange. - - Fixpoint collect_args_and_apply_unknown_casts {t} - : (int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => option T) arith_expr_for_base t) - -> type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) - t - := match t - return ((int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => option T) arith_expr_for_base t) - -> type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) - t) - with - | type.base t => fun v => Some v - | type.arrow (type.base s) d - => fun f - (x : (int.option.interp s -> option (arith_expr_for_base s))) - => match x int.option.None with - | Some x' - => @collect_args_and_apply_unknown_casts - d - (fun rout => f rout x') - | None => type.interpM_return _ _ _ None - end - | type.arrow (type.arrow _ _) _ - => fun _ => type.interpM_return _ _ _ None - end. + Section with_bind. + (* N.B. If we make the [bind*_err] notations, then Coq can't + infer types correctly; if we make them [Local + Definition]s or [Let]s, then [ocamlopt] fails with + "Error: The type of this module, [...], contains type + variables that cannot be generalized". We need to run + [cbv] below to actually unfold them. >.< *) + Local Notation ErrT T := (T + list string)%type. + Local Notation ret v := (@inl _ (list string) v) (only parsing). + Local Notation "x <- v ; f" := (match v with + | inl x => f + | inr err => inr err + end). + Reserved Notation "A1 ,, A2 <- X ; B" (at level 70, A2 at next level, right associativity, format "'[v' A1 ,, A2 <- X ; '/' B ']'"). + Reserved Notation "A1 ,, A2 <- X1 , X2 ; B" (at level 70, A2 at next level, right associativity, format "'[v' A1 ,, A2 <- X1 , X2 ; '/' B ']'"). + Reserved Notation "A1 ,, A2 ,, A3 <- X ; B" (at level 70, A2 at next level, A3 at next level, right associativity, format "'[v' A1 ,, A2 ,, A3 <- X ; '/' B ']'"). + Reserved Notation "A1 ,, A2 ,, A3 <- X1 , X2 , X3 ; B" (at level 70, A2 at next level, A3 at next level, right associativity, format "'[v' A1 ,, A2 ,, A3 <- X1 , X2 , X3 ; '/' B ']'"). + Reserved Notation "A1 ,, A2 ,, A3 ,, A4 <- X ; B" (at level 70, A2 at next level, A3 at next level, A4 at next level, right associativity, format "'[v' A1 ,, A2 ,, A3 ,, A4 <- X ; '/' B ']'"). + Reserved Notation "A1 ,, A2 ,, A3 ,, A4 <- X1 , X2 , X3 , X4 ; B" (at level 70, A2 at next level, A3 at next level, A4 at next level, right associativity, format "'[v' A1 ,, A2 ,, A3 ,, A4 <- X1 , X2 , X3 , X4 ; '/' B ']'"). + Reserved Notation "A1 ,, A2 ,, A3 ,, A4 ,, A5 <- X ; B" (at level 70, A2 at next level, A3 at next level, A4 at next level, A5 at next level, right associativity, format "'[v' A1 ,, A2 ,, A3 ,, A4 ,, A5 <- X ; '/' B ']'"). + Reserved Notation "A1 ,, A2 ,, A3 ,, A4 ,, A5 <- X1 , X2 , X3 , X4 , X5 ; B" (at level 70, A2 at next level, A3 at next level, A4 at next level, A5 at next level, right associativity, format "'[v' A1 ,, A2 ,, A3 ,, A4 ,, A5 <- X1 , X2 , X3 , X4 , X5 ; '/' B ']'"). + Let bind2_err {A B C} (v1 : ErrT A) (v2 : ErrT B) (f : A -> B -> ErrT C) : ErrT C + := match v1, v2 with + | inl x1, inl x2 => f x1 x2 + | inr err, inl _ | inl _, inr err => inr err + | inr err1, inr err2 => inr (List.app err1 err2) + end. + Local Notation "x1 ,, x2 <- v1 , v2 ; f" + := (bind2_err v1 v2 (fun x1 x2 => f)). + Local Notation "x1 ,, x2 <- v ; f" + := (x1 ,, x2 <- fst v , snd v; f). + Let bind3_err {A B C R} (v1 : ErrT A) (v2 : ErrT B) (v3 : ErrT C) (f : A -> B -> C -> ErrT R) : ErrT R + := (x12 ,, x3 <- (x1 ,, x2 <- v1, v2; inl (x1, x2)), v3; + let '(x1, x2) := x12 in + f x1 x2 x3). + Local Notation "x1 ,, x2 ,, x3 <- v1 , v2 , v3 ; f" + := (bind3_err v1 v2 v3 (fun x1 x2 x3 => f)). + Local Notation "x1 ,, x2 ,, x3 <- v ; f" + := (let '(v1, v2, v3) := v in x1 ,, x2 ,, x3 <- v1 , v2 , v3; f). + Let bind4_err {A B C D R} (v1 : ErrT A) (v2 : ErrT B) (v3 : ErrT C) (v4 : ErrT D) (f : A -> B -> C -> D -> ErrT R) : ErrT R + := (x12 ,, x34 <- (x1 ,, x2 <- v1, v2; inl (x1, x2)), (x3 ,, x4 <- v3, v4; inl (x3, x4)); + let '((x1, x2), (x3, x4)) := (x12, x34) in + f x1 x2 x3 x4). + Local Notation "x1 ,, x2 ,, x3 ,, x4 <- v1 , v2 , v3 , v4 ; f" + := (bind4_err v1 v2 v3 v4 (fun x1 x2 x3 x4 => f)). + Local Notation "x1 ,, x2 ,, x3 ,, x4 <- v ; f" + := (let '(v1, v2, v3, v4) := v in x1 ,, x2 ,, x3 ,, x4 <- v1 , v2 , v3 , v4; f). + Let bind5_err {A B C D E R} (v1 : ErrT A) (v2 : ErrT B) (v3 : ErrT C) (v4 : ErrT D) (v5 : ErrT E) (f : A -> B -> C -> D -> E -> ErrT R) : ErrT R + := (x12 ,, x345 <- (x1 ,, x2 <- v1, v2; inl (x1, x2)), (x3 ,, x4 ,, x5 <- v3, v4, v5; inl (x3, x4, x5)); + let '((x1, x2), (x3, x4, x5)) := (x12, x345) in + f x1 x2 x3 x4 x5). + Local Notation "x1 ,, x2 ,, x3 ,, x4 ,, x5 <- v1 , v2 , v3 , v4 , v5 ; f" + := (bind5_err v1 v2 v3 v4 v5 (fun x1 x2 x3 x4 x5 => f)). + Local Notation "x1 ,, x2 ,, x3 ,, x4 ,, x5 <- v ; f" + := (let '(v1, v2, v3, v4, v5) := v in x1 ,, x2 ,, x3 ,, x4 ,, x5 <- v1 , v2 , v3 , v4 , v5; f). + + Definition maybe_log2 (s : Z) : option Z + := if 2^Z.log2 s =? s then Some (Z.log2 s) else None. + Definition maybe_loglog2 (s : Z) : option nat + := (v <- maybe_log2 s; + v <- maybe_log2 v; + if Z.leb 0 v + then Some (Z.to_nat v) + else None). + + + Definition arith_expr_of_PHOAS_ident + {t} + (idc : ident.ident t) + : int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => ErrT T) arith_expr_for_base t + := match idc in ident.ident t return int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => ErrT T) arith_expr_for_base t with + | ident.Literal base.type.Z v + => fun r => ret (cast_down_if_needed + r + (literal v @@ TT, Some (int.of_zrange_relaxed r[v~>v]))) + | ident.nil t + => fun _ => ret nil + | ident.cons t + => fun r x xs => ret (cast_down_if_needed r (cons x xs)) + | ident.fst A B => fun r xy => ret (cast_down_if_needed r (@fst _ _ xy)) + | ident.snd A B => fun r xy => ret (cast_down_if_needed r (@snd _ _ xy)) + | ident.List_nth_default base.type.Z + => fun r d ls n + => List.nth_default (inr ["Invalid list index " ++ show false n]%string) + (List.map (fun x => ret (cast_down_if_needed r x)) ls) n + | ident.Z_shiftr + => fun rout '(e, r) '(offset, roffset) + => let rin := option_map integer_promote_type r in + match invert_literal offset with + | Some offset => ret (cast_down_if_needed rout (Z_shiftr offset @@ e, rin)) + | None => inr ["Invalid right-shift by a non-literal"]%string + end + | ident.Z_shiftl + => fun rout '(e, r) '(offset, roffset) + => let rin := option_map integer_promote_type r in + match invert_literal offset with + | Some offset + => let '(e', rin') := cast_up_if_needed rout (e, rin) in + ret (cast_down_if_needed rout (Z_shiftl offset @@ e', rin')) + | None => inr ["Invalid left-shift by a non-literal"]%string + end + | ident.Z_bneg + => fun rout '(e, r) + => (** N.B. We assume that C will implicitly cast the output of [!e] to whatever integer type it wants it in *) + ret (Z_bneg @@ e, rout) + | ident.Z_land => fun r x y => ret (arith_bin_arith_expr_of_PHOAS_ident Z_land r (x, y)) + | ident.Z_lor => fun r x y => ret (arith_bin_arith_expr_of_PHOAS_ident Z_lor r (x, y)) + | ident.Z_add => fun r x y => ret (arith_bin_arith_expr_of_PHOAS_ident Z_add r (x, y)) + | ident.Z_mul => fun r x y => ret (arith_bin_arith_expr_of_PHOAS_ident Z_mul r (x, y)) + | ident.Z_sub => fun r x y => ret (arith_bin_arith_expr_of_PHOAS_ident Z_sub r (x, y)) + | ident.Z_lnot_modulo + => fun rout '(e, r) '(modulus, _) + => let rin := option_map integer_promote_type r in + match invert_literal modulus with + | Some modulus + => match maybe_loglog2 modulus with + | Some lgbitwidth + => let ty := int.unsigned lgbitwidth in + let rin' := Some ty in + let '(e, _) := Zcast_up_if_needed rin' (e, r) in + ret (cast_down_if_needed rout (cast_up_if_needed rout (Z_lnot ty @@ e, rin'))) + | None => inr ["Invalid modulus for Z.lneg (not 2^(2^_)): " ++ show false modulus]%string + end + | None => inr ["Invalid non-literal modulus for Z.lneg"]%string + end + | ident.pair A B + => fun _ _ _ => inr ["Invalid identifier in arithmetic expression " ++ show true idc]%string + | ident.Z_opp (* we pretend this is [0 - _] *) + => fun r y => let zero := (literal 0 @@ TT, Some (int.of_zrange_relaxed r[0~>0])) in + ret (arith_bin_arith_expr_of_PHOAS_ident Z_sub r (zero, y)) + | ident.Literal _ v + => fun _ => ret v + | ident.Nat_succ + | ident.Nat_pred + | ident.Nat_max + | ident.Nat_mul + | ident.Nat_add + | ident.Nat_sub + | ident.pair_rect _ _ _ + | ident.bool_rect _ + | ident.nat_rect _ + | ident.list_rect _ _ + | ident.list_case _ _ + | ident.List_length _ + | ident.List_seq + | ident.List_repeat _ + | ident.List_combine _ _ + | ident.List_map _ _ + | ident.List_app _ + | ident.List_rev _ + | ident.List_flat_map _ _ + | ident.List_partition _ + | ident.List_fold_right _ _ + | ident.List_update_nth _ + | ident.List_nth_default _ + | ident.Z_pow + | ident.Z_div + | ident.Z_modulo + | ident.Z_eqb + | ident.Z_leb + | ident.Z_geb + | ident.Z_log2 + | ident.Z_log2_up + | ident.Z_of_nat + | ident.Z_to_nat + | ident.Z_zselect + | ident.Z_mul_split + | ident.Z_add_get_carry + | ident.Z_add_with_carry + | ident.Z_add_with_get_carry + | ident.Z_sub_get_borrow + | ident.Z_sub_with_get_borrow + | ident.Z_add_modulo + | ident.Z_rshi + | ident.Z_cc_m + | ident.Z_cast _ + | ident.Z_cast2 _ + | ident.fancy_add _ _ + | ident.fancy_addc _ _ + | ident.fancy_sub _ _ + | ident.fancy_subb _ _ + | ident.fancy_mulll _ + | ident.fancy_mullh _ + | ident.fancy_mulhl _ + | ident.fancy_mulhh _ + | ident.fancy_rshi _ _ + | ident.fancy_selc + | ident.fancy_selm _ + | ident.fancy_sell + | ident.fancy_addm + => fun _ => type.interpM_return _ _ _ (inr ["Invalid identifier in arithmetic expression " ++ show true idc]%string) + end%core%Cexpr%option%zrange. + + Fixpoint collect_args_and_apply_unknown_casts {t} + : (int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => ErrT T) arith_expr_for_base t) + -> type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) + t + := match t + return ((int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => ErrT T) arith_expr_for_base t) + -> type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) + t) + with + | type.base t => fun v => ret v + | type.arrow (type.base s) d + => fun f + (x : (int.option.interp s -> ErrT (arith_expr_for_base s))) + => match x int.option.None with + | inl x' + => @collect_args_and_apply_unknown_casts + d + (fun rout => f rout x') + | inr errs => type.interpM_return _ _ _ (inr errs) + end + | type.arrow (type.arrow _ _) _ + => fun _ => type.interpM_return _ _ _ (inr ["Invalid higher-order function"]) + end. - Definition collect_args_and_apply_known_casts {t} - (idc : ident.ident t) - : option (type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) + Definition collect_args_and_apply_known_casts {t} + (idc : ident.ident t) + : ErrT (type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) t) - := match idc in ident.ident t - return option - (type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) - t) - with - | ident.Z_cast r - => Some (fun arg => Some (fun r' => option_map (Zcast_down_if_needed r') (arg (Some (int.of_zrange_relaxed r))))) - | ident.Z_cast2 (r1, r2) - => Some (fun arg => Some (fun r' => option_map (cast_down_if_needed (t:=base.type.Z*base.type.Z) r') - (arg (Some (int.of_zrange_relaxed r1), Some (int.of_zrange_relaxed r2))))) - | ident.pair A B - => Some (fun ea eb - => Some - (fun '(ra, rb) - => (ea' <- ea ra; - eb' <- eb rb; - Some (ea', eb')))) - | ident.nil _ - => Some (Some (fun _ => Some nil)) - | ident.cons t - => Some - (fun x xs - => Some - (fun rls - => let mkcons (r : int.option.interp t) - (rs : int.option.interp (base.type.list t)) - := (x <- x r; - xs <- xs rs; - Some (cons x xs)) in - match rls with - | Some (cons r rs) => mkcons r (Some rs) - | Some nil - | None - => mkcons int.option.None int.option.None - end)) - | _ => None - end%option. - - Definition collect_args_and_apply_casts {t} (idc : ident.ident t) - (convert_no_cast : int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => option T) arith_expr_for_base t) - : type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) - t - := match collect_args_and_apply_known_casts idc with - | Some res => res - | None => collect_args_and_apply_unknown_casts convert_no_cast - end. + := match idc in ident.ident t + return ErrT + (type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) + t) + with + | ident.Z_cast r + => inl (fun arg => inl (fun r' => arg <- arg (Some (int.of_zrange_relaxed r)); ret (Zcast_down_if_needed r' arg))) + | ident.Z_cast2 (r1, r2) + => inl (fun arg => inl (fun r' => arg <- (arg (Some (int.of_zrange_relaxed r1), Some (int.of_zrange_relaxed r2))); + ret (cast_down_if_needed (t:=base.type.Z*base.type.Z) r' arg))) + | ident.pair A B + => inl (fun ea eb + => inl + (fun '(ra, rb) + => ea' ,, eb' <- ea ra, eb rb; + inl (ea', eb'))) + | ident.nil _ + => inl (inl (fun _ => inl nil)) + | ident.cons t + => inl + (fun x xs + => inl + (fun rls + => let mkcons (r : int.option.interp t) + (rs : int.option.interp (base.type.list t)) + := x ,, xs <- x r, xs rs; + inl (cons x xs) in + match rls with + | Some (cons r rs) => mkcons r (Some rs) + | Some nil + | None + => mkcons int.option.None int.option.None + end)) + | _ => inr ["Invalid identifier where cast or constructor was expected: " ++ show false idc]%string + end. - Fixpoint arith_expr_of_base_PHOAS_Var - {t} - : base_var_data t -> int.option.interp t -> option (arith_expr_for_base t) - := match t with - | base.type.Z - => fun '(n, r) r' => Some (cast_down_if_needed r' (Var type.Z n, r)) - | base.type.prod A B - => fun '(da, db) '(ra, rb) - => (ea <- @arith_expr_of_base_PHOAS_Var A da ra; - eb <- @arith_expr_of_base_PHOAS_Var B db rb; - Some (ea, eb))%option - | base.type.list base.type.Z - => fun '(n, r, len) r' - => Some (List.map - (fun i => (List_nth i @@ Var type.Zptr n, r))%core%Cexpr - (List.seq 0 len)) - | base.type.list _ - | base.type.type_base _ - => fun _ _ => None - end. + Definition collect_args_and_apply_casts {t} (idc : ident.ident t) + (convert_no_cast : int.option.interp (type.final_codomain t) -> type.interpM_final (fun T => ErrT T) arith_expr_for_base t) + : type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) + t + := match collect_args_and_apply_known_casts idc with + | inl res => res + | inr errs => collect_args_and_apply_unknown_casts convert_no_cast + end. - Fixpoint arith_expr_of_PHOAS - {t} - (e : @Compilers.expr.expr base.type ident.ident var_data t) - : type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) - t - := match e in expr.expr t - return type.interpM_final - (fun T => option T) - (fun t => int.option.interp t -> option (arith_expr_for_base t)) - t - with - | expr.Var (type.base _) v - => Some (arith_expr_of_base_PHOAS_Var v) - | expr.Ident t idc - => collect_args_and_apply_casts idc (arith_expr_of_PHOAS_ident idc) - | expr.App (type.base s) d f x - => let x' := @arith_expr_of_PHOAS s x in - match x' with - | Some x' => @arith_expr_of_PHOAS _ f x' - | None => type.interpM_return _ _ _ None - end - | expr.Var (type.arrow _ _) _ - | expr.App (type.arrow _ _) _ _ _ - | expr.LetIn _ _ _ _ - | expr.Abs _ _ _ - => type.interpM_return _ _ _ None - end. + Fixpoint arith_expr_of_base_PHOAS_Var + {t} + : base_var_data t -> int.option.interp t -> ErrT (arith_expr_for_base t) + := match t with + | base.type.Z + => fun '(n, r) r' => ret (cast_down_if_needed r' (Var type.Z n, r)) + | base.type.prod A B + => fun '(da, db) '(ra, rb) + => (ea,, eb <- @arith_expr_of_base_PHOAS_Var A da ra, @arith_expr_of_base_PHOAS_Var B db rb; + inl (ea, eb)) + | base.type.list base.type.Z + => fun '(n, r, len) r' + => ret (List.map + (fun i => (List_nth i @@ Var type.Zptr n, r))%core%Cexpr + (List.seq 0 len)) + | base.type.list _ + | base.type.type_base _ + => fun _ _ => inr ["Invalid type " ++ show false t]%string + end. - Definition arith_expr_of_base_PHOAS - {t:base.type} + Fixpoint arith_expr_of_PHOAS + {t} (e : @Compilers.expr.expr base.type ident.ident var_data t) - (rout : int.option.interp t) - : option (arith_expr_for_base t) - := (e' <- arith_expr_of_PHOAS e; e' rout)%option. - - Fixpoint make_return_assignment_of_base_arith {t} - : base_var_data t - -> @Compilers.expr.expr base.type ident.ident var_data t - -> option expr - := match t return base_var_data t -> expr.expr t -> option expr with - | base.type.Z - => fun '(n, r) e - => (rhs <- arith_expr_of_base_PHOAS e r; - let '(e, r) := rhs in - Some [AssignZPtr n r e]) - | base.type.type_base _ => fun _ _ => None - | base.type.prod A B - => fun '(rva, rvb) e - => match invert_pair e with - | Some (ea, eb) - => (ea' <- @make_return_assignment_of_base_arith A rva ea; - eb' <- @make_return_assignment_of_base_arith B rvb eb; - Some (ea' ++ eb')) - | None => None + : type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) + t + := match e in expr.expr t + return type.interpM_final + (fun T => ErrT T) + (fun t => int.option.interp t -> ErrT (arith_expr_for_base t)) + t + with + | expr.Var (type.base _) v + => ret (arith_expr_of_base_PHOAS_Var v) + | expr.Ident t idc + => collect_args_and_apply_casts idc (arith_expr_of_PHOAS_ident idc) + | expr.App (type.base s) d f x + => let x' := @arith_expr_of_PHOAS s x in + match x' with + | inl x' => @arith_expr_of_PHOAS _ f x' + | inr errs => type.interpM_return _ _ _ (inr errs) end - | base.type.list base.type.Z - => fun '(n, r, len) e - => (ls <- arith_expr_of_base_PHOAS e (Some (repeat r len)); - List.fold_right - (fun a b - => match b with - | Some b => Some (a ++ b) - | None => Some a - end) - None - (List.map - (fun '(i, (e, _)) => [AssignNth n i e]) - (List.combine (List.seq 0 len) ls))) - | base.type.list _ => fun _ _ => None - end%option%list. - Definition make_return_assignment_of_arith {t} - : var_data t - -> @Compilers.expr.expr base.type ident.ident var_data t - -> option expr - := match t with - | type.base t => make_return_assignment_of_base_arith - | type.arrow s d => fun _ _ => None - end. + | expr.Var (type.arrow _ _) _ + => type.interpM_return _ _ _ (inr ["Invalid non-arithmetic let-bound Var of type " ++ show false t]%string) + | expr.App (type.arrow _ _) _ _ _ + => type.interpM_return _ _ _ (inr ["Invalid non-arithmetic let-bound App of type " ++ show false t]%string) + | expr.LetIn _ _ _ _ + => type.interpM_return _ _ _ (inr ["Invalid non-arithmetic let-bound LetIn of type " ++ show false t]%string) + | expr.Abs _ _ _ + => type.interpM_return _ _ _ (inr ["Invalid non-arithmetic let-bound Abs of type: " ++ show false t]%string) + end. - Definition make_assign_2arg_1ref - {t1 t2 d t3} - (r1 r2 : option int.type) - (x1 : @Compilers.expr.expr base.type ident.ident var_data t1) - (x2 : @Compilers.expr.expr base.type ident.ident var_data t2) - (idc : ident (type.Z * type.Z * type.Zptr) type.Z) - (count : positive) - (make_name : positive -> option string) - (v : var_data t3) - (e2 : var_data d -> var_data (base.type.Z * base.type.Z)%etype -> option expr) - : option expr - := (v <- type.try_transport base.try_make_transport_cps var_data _ d v; - x1 <- type.try_transport base.try_make_transport_cps expr.expr _ base.type.Z x1; - x2 <- type.try_transport base.try_make_transport_cps expr.expr _ base.type.Z x2; - let e2 := e2 v in - x1 <- arith_expr_of_base_PHOAS x1 None; - x2 <- arith_expr_of_base_PHOAS x2 None; - let '(x1, x1r) := x1 in - let '(x2, x2r) := x2 in - n1 <- make_name count; - n2 <- make_name (Pos.succ count); - e2 <- e2 ((n1, r1), (n2, r2)); - Some ([DeclareVar type.Z r2 n2; - Assign true type.Z r1 n1 (idc @@ (x1, x2, Addr @@ Var type.Z n2))] - ++ e2))%option%list. - - Definition make_assign_3arg_1ref - {t1 t2 t3 d t4} - (r1 r2 : option int.type) - (x1 : @Compilers.expr.expr base.type ident.ident var_data t1) - (x2 : @Compilers.expr.expr base.type ident.ident var_data t2) - (x3 : @Compilers.expr.expr base.type ident.ident var_data t3) - (idc : ident (type.Z * type.Z * type.Z * type.Zptr) type.Z) - (count : positive) - (make_name : positive -> option string) - (v : var_data t4) - (e2 : var_data d -> var_data (base.type.Z * base.type.Z)%etype -> option expr) - : option expr - := (v <- type.try_transport base.try_make_transport_cps var_data _ d v; - x1 <- type.try_transport base.try_make_transport_cps expr.expr _ base.type.Z x1; - x2 <- type.try_transport base.try_make_transport_cps expr.expr _ base.type.Z x2; - x3 <- type.try_transport base.try_make_transport_cps expr.expr _ base.type.Z x3; - let e2 := e2 v in - x1 <- arith_expr_of_base_PHOAS x1 None; - x2 <- arith_expr_of_base_PHOAS x2 None; - x3 <- arith_expr_of_base_PHOAS x3 None; - let '(x1, x1r) := x1 in - let '(x2, x2r) := x2 in - let '(x3, x3r) := x3 in - n1 <- make_name count; - n2 <- make_name (Pos.succ count); - e2 <- e2 ((n1, r1), (n2, r2)); - Some ([DeclareVar type.Z r2 n2; - Assign true type.Z r1 n1 (idc @@ (x1, x2, x3, Addr @@ Var type.Z n2))] - ++ e2))%option%list. - - Fixpoint size_of_type (t : base.type) : positive - := match t with - | base.type.type_base t => 1 - | base.type.prod A B => size_of_type A + size_of_type B - | base.type.list A => 1 - end%positive. - - Definition maybe_log2 (s : Z) : option Z - := if 2^Z.log2 s =? s then Some (Z.log2 s) else None. - - Definition recognize_ident_2arg {t} (idc : ident.ident t) - : option (ident (type.type_primitive type.Z * type.type_primitive type.Z * type.type_primitive type.Zptr) (type.type_primitive type.Z)) - := match idc with - | ident.Z_mul_split_concrete s - => option_map Z_mul_split (maybe_log2 s) - | ident.Z_add_get_carry_concrete s - => option_map Z_add_get_carry (maybe_log2 s) - | ident.Z_sub_get_borrow_concrete s - => option_map Z_sub_get_borrow (maybe_log2 s) - | _ => None - end. - Definition recognize_ident_3arg {t} (idc : ident.ident t) - : option (ident (type.type_primitive type.Z * type.type_primitive type.Z * type.type_primitive type.Z * type.type_primitive type.Zptr) (type.type_primitive type.Z)) - := match idc with - | ident.Z_add_with_get_carry_concrete s - => option_map Z_add_with_get_carry (maybe_log2 s) - | ident.Z_sub_with_get_borrow_concrete s - => option_map Z_sub_with_get_borrow (maybe_log2 s) - | _ => None - end. + Definition arith_expr_of_base_PHOAS + {t:base.type} + (e : @Compilers.expr.expr base.type ident.ident var_data t) + (rout : int.option.interp t) + : ErrT (arith_expr_for_base t) + := (e' <- arith_expr_of_PHOAS e; e' rout). + + Fixpoint make_return_assignment_of_base_arith {t} + : base_var_data t + -> @Compilers.expr.expr base.type ident.ident var_data t + -> ErrT expr + := match t return base_var_data t -> expr.expr t -> ErrT expr with + | base.type.Z + => fun '(n, r) e + => (rhs <- arith_expr_of_base_PHOAS e r; + let '(e, r) := rhs in + ret [AssignZPtr n r e]) + | base.type.type_base _ => fun _ _ => inr ["Invalid type " ++ show false t]%string + | base.type.prod A B + => fun '(rva, rvb) e + => match invert_pair e with + | Some (ea, eb) + => ea',, eb' <- @make_return_assignment_of_base_arith A rva ea, @make_return_assignment_of_base_arith B rvb eb; + ret (ea' ++ eb') + | None => inr ["Invalid non-pair expr of type " ++ show false t]%string + end + | base.type.list base.type.Z + => fun '(n, r, len) e + => (ls <- arith_expr_of_base_PHOAS e (Some (repeat r len)); + ret (List.map + (fun '(i, (e, _)) => AssignNth n i e) + (List.combine (List.seq 0 len) ls))) + | base.type.list _ => fun _ _ => inr ["Invalid type of expr: " ++ show false t]%string + end%list. + Definition make_return_assignment_of_arith {t} + : var_data t + -> @Compilers.expr.expr base.type ident.ident var_data t + -> ErrT expr + := match t with + | type.base t => make_return_assignment_of_base_arith + | type.arrow s d => fun _ _ => inr ["Invalid type of expr: " ++ show false t]%string + end. - Definition make_uniform_assign_expr_of_PHOAS - {s} (e1 : @Compilers.expr.expr base.type ident.ident var_data s) - {d} (e2 : var_data s -> var_data d -> option expr) - (count : positive) - (make_name : positive -> option string) - (v : var_data d) - : option expr - := match s return (@Compilers.expr.expr base.type ident.ident var_data s) - -> (var_data s -> var_data d -> option expr) - -> option expr - with - | type.base (base.type.type_base base.type.Z) - => fun e1 e2 - => (e1 <- arith_expr_of_base_PHOAS e1 None; - let '(e1, r1) := e1 in - n1 <- make_name count; - e2 <- e2 (n1, r1) v; - Some ((Assign true type.Z r1 n1 e1) - :: e2)) - | type.base (base.type.Z * base.type.Z)%etype - => fun e1 e2 - => let '((r1, r2), e1)%core - := match invert_Z_cast2 e1 with - | Some ((r1, r2), e) => ((Some (int.of_zrange_relaxed r1), Some (int.of_zrange_relaxed r2)), e) - | None => ((None, None), e1) - end%core in - match e1 with - | (#idc @ x1 @ x2) - => idc <- recognize_ident_2arg idc; - make_assign_2arg_1ref - r1 r2 - x1 x2 idc count make_name v - (fun v rv => e2 rv v) - | (#idc @ x1 @ x2 @ x3) - => idc <- recognize_ident_3arg idc; - make_assign_3arg_1ref - r1 r2 - x1 x2 x3 idc count make_name v - (fun v rv => e2 rv v) - | _ => None - end%expr_pat - | _ => fun _ _ => None - end%option e1 e2. - Definition make_assign_expr_of_PHOAS - {s} (e1 : @Compilers.expr.expr base.type ident.ident var_data s) - {s' d} (e2 : var_data s' -> var_data d -> option expr) - (count : positive) - (make_name : positive -> option string) - (v : var_data d) - : option expr - := (e1 <- type.try_transport base.try_make_transport_cps _ _ _ e1; - make_uniform_assign_expr_of_PHOAS e1 e2 count make_name v). + Definition report_type_mismatch (expected : defaults.type) (given : defaults.type) : string + := ("Type mismatch; expected " ++ show false expected ++ " but given " ++ show false given ++ ".")%string. - Fixpoint expr_of_base_PHOAS - {t} - (e : @Compilers.expr.expr base.type ident.ident var_data t) - (count : positive) - (make_name : positive -> option string) - {struct e} - : forall (ret_val : var_data t), option expr - := match e in expr.expr t return var_data t -> option expr with - | expr.LetIn (type.base s) d e1 e2 - => make_assign_expr_of_PHOAS - e1 - (fun vs vd => @expr_of_base_PHOAS d (e2 vs) (size_of_type s + count)%positive make_name vd) - count make_name - | expr.LetIn (type.arrow _ _) _ _ _ as e - | expr.Var _ _ as e - | expr.Ident _ _ as e - | expr.App _ _ _ _ as e - | expr.Abs _ _ _ as e - => fun v => make_return_assignment_of_arith v e - end%expr_pat%option. - - Fixpoint base_var_data_of_bounds {t} - (count : positive) - (make_name : positive -> option string) - {struct t} - : ZRange.type.base.option.interp t -> option (positive * var_data t) - := match t return ZRange.type.base.option.interp t -> option (positive * var_data t) with - | base.type.Z - => fun r => (n <- make_name count; - Some (Pos.succ count, (n, option_map int.of_zrange_relaxed r))) - | base.type.prod A B - => fun '(ra, rb) - => (va <- @base_var_data_of_bounds A count make_name ra; - let '(count, va) := va in - vb <- @base_var_data_of_bounds B count make_name rb; - let '(count, vb) := vb in - Some (count, (va, vb))) - | base.type.list base.type.Z - => fun r - => (ls <- r; - n <- make_name count; - Some (Pos.succ count, - (n, - match List.map (option_map int.of_zrange_relaxed) ls with - | nil => None - | cons x xs - => List.fold_right - (fun r1 r2 => r1 <- r1; r2 <- r2; Some (int.union r1 r2)) - x - xs - end, - length ls))) - | base.type.unit - => fun _ => Some (count, tt) - | base.type.list _ - | base.type.type_base _ - => fun _ => None - end%option. + Fixpoint arith_expr_of_PHOAS_args + {t} + : type.for_each_lhs_of_arrow (@Compilers.expr.expr base.type ident.ident var_data) t + -> ErrT (type.for_each_lhs_of_arrow (fun t => @Compilers.expr.expr base.type ident.ident var_data t * (arith_expr type.Z * option int.type)) t) + := match t with + | type.base t => fun 'tt => inl tt + | type.arrow (type.base base.type.Z) d + => fun '(arg, args) + => arg' ,, args' <- arith_expr_of_base_PHOAS arg int.option.None , @arith_expr_of_PHOAS_args d args; + inl ((arg, arg'), args') + | type.arrow s d + => fun '(arg, args) + => arg' ,, args' <- @inr unit _ ["Invalid argument of non-ℤ type " ++ show false s]%string , @arith_expr_of_PHOAS_args d args; + inr ["Impossible! (type error got lost somewhere)"] + end. - Definition var_data_of_bounds {t} + Let recognize_1ref_ident + {t} + (idc : ident.ident t) + (rout : option int.type) + : type.for_each_lhs_of_arrow (fun t => @Compilers.expr.expr base.type ident.ident var_data t * (arith_expr type.Z * option int.type))%type t + -> ErrT (arith_expr type.Zptr -> expr) + := let _ := @PHOAS.expr.partially_show_expr in + match idc with + | ident.Z_zselect + => fun '((econdv, (econd, rcond)), ((e1v, (e1, r1)), ((e2v, (e2, r2)), tt))) + => match rcond with + | Some (int.unsigned 0) + => let r1 := option_map integer_promote_type r1 in + let r2 := option_map integer_promote_type r2 in + let '((e1, r1), (e2, r2)) + := cast_bigger_up_if_needed rout ((e1, r1), (e2, r2)) in + let ct := (r1 <- r1; r2 <- r2; Some (common_type r1 r2))%option in + ty <- match ct, rout with + | Some ct, Some rout + => if int.type_beq ct rout + then inl ct + else inr ["Casting the result of Z.zselect to a type other than the common type is not currently supported. (Cannot cast a pointer to " ++ show false rout ++ " to a pointer to " ++ show false ct ++ ".)"]%string + | None, _ => inr ["Unexpected None result of common-type calculation for Z.zselect"] + | _, None => inr ["Missing cast annotation on return of Z.zselect"] + end; + ret (fun retptr => [Call (Z_zselect ty @@ (retptr, (econd, e1, e2)))]%Cexpr) + | _ => inr ["Invalid argument to Z.zselect not known to be 0 or 1: " ++ show false econdv]%string + end + | _ => fun _ => inr ["Unrecognized identifier (expecting a 1-pointer-returning function): " ++ show false idc]%string + end. + + Definition bounds_check (descr : string) {t} (idc : ident.ident t) (s : BinInt.Z) {t'} (ev : @Compilers.expr.expr base.type ident.ident var_data t') (found : option int.type) + : ErrT unit + := let _ := @PHOAS.expr.partially_show_expr in + match found with + | None => inr ["Missing range on " ++ descr ++ " " ++ show true idc ++ ": " ++ show true ev]%string + | Some ty + => if int.is_tighter_than ty (int.of_zrange_relaxed r[0~>2^s-1]) + then ret tt + else inr ["Final bounds check failed on " ++ descr ++ " " ++ show false idc ++ "; expected an unsigned " ++ decimal_string_of_Z s ++ "-bit number (" ++ show false (int.of_zrange_relaxed r[0~>2^s-1]) ++ "), but found a " ++ show false ty ++ "."]%string + end. + + Let recognize_3arg_2ref_ident + (t:=(base.type.Z -> base.type.Z -> base.type.Z -> base.type.Z * base.type.Z)%etype) + (idc : ident.ident t) + (rout : option int.type * option int.type) + (args : type.for_each_lhs_of_arrow (fun t => @Compilers.expr.expr base.type ident.ident var_data t * (arith_expr type.Z * option int.type))%type t) + : ErrT (arith_expr (type.Zptr * type.Zptr) -> expr) + := let _ := @PHOAS.expr.partially_show_expr in + let '((s, _), ((e1v, (e1, r1)), ((e2v, (e2, r2)), tt))) := args in + match (s <- invert_Literal s; maybe_log2 s)%option, idc return ErrT (arith_expr (type.Zptr * type.Zptr) -> expr) + with + | Some s, ident.Z_mul_split + => (_ ,, _ ,, _ ,, _ + <- bounds_check "first argument to" idc s e1v r1, + bounds_check "second argument to" idc s e2v r2, + bounds_check "first return value of" idc s e2v (fst rout), + bounds_check "second return value of" idc s e2v (snd rout); + inl (fun retptr => [Call (Z_mul_split s @@ (retptr, (e1, e2)))%Cexpr])) + | Some s, ident.Z_add_get_carry as idc + | Some s, ident.Z_sub_get_borrow as idc + => let idc' : ident _ _ := invert_Some match idc with + | ident.Z_add_get_carry => Some (Z_add_with_get_carry s) + | ident.Z_sub_get_borrow => Some (Z_sub_with_get_borrow s) + | _ => None + end in + (_ ,, _ ,, _ ,, _ + <- bounds_check "first argument to" idc s e1v r1, + bounds_check "second argument to" idc s e2v r2, + bounds_check "first return value of" idc s e2v (fst rout), + bounds_check "second return value of" idc 1 (* boolean carry/borrow *) e2v (snd rout); + inl (fun retptr => [Call (idc' @@ (retptr, (literal 0 @@ TT, e1, e2)))%Cexpr])) + | Some _, _ => inr ["Unrecognized identifier when attempting to construct an assignment with 2 arguments: " ++ show true idc]%string + | None, _ => inr ["Expression is not a literal power of two of type ℤ: " ++ show true s ++ " (when trying to parse the first argument of " ++ show true idc ++ ")"]%string + end. + + Let recognize_4arg_2ref_ident + (t:=(base.type.Z -> base.type.Z -> base.type.Z -> base.type.Z -> base.type.Z * base.type.Z)%etype) + (idc : ident.ident t) + (rout : option int.type * option int.type) + (args : type.for_each_lhs_of_arrow (fun t => @Compilers.expr.expr base.type ident.ident var_data t * (arith_expr type.Z * option int.type))%type t) + : ErrT (arith_expr (type.Zptr * type.Zptr) -> expr) + := let _ := @PHOAS.expr.partially_show_expr in + let '((s, _), ((e1v, (e1, r1)), ((e2v, (e2, r2)), ((e3v, (e3, r3)), tt)))) := args in + match (s <- invert_Literal s; maybe_log2 s)%option, idc return ErrT (arith_expr (type.Zptr * type.Zptr) -> expr) + with + | Some s, ident.Z_add_with_get_carry as idc + | Some s, ident.Z_sub_with_get_borrow as idc + => let idc' : ident _ _ := invert_Some match idc with + | ident.Z_add_with_get_carry => Some (Z_add_with_get_carry s) + | ident.Z_sub_with_get_borrow => Some (Z_sub_with_get_borrow s) + | _ => None + end in + (_ ,, _ ,, _ ,, _ ,, _ + <- (bounds_check "first (carry) argument to" idc 1 e1v r1, + bounds_check "second argument to" idc s e2v r2, + bounds_check "third argument to" idc s e2v r2, + bounds_check "first return value of" idc s e2v (fst rout), + bounds_check "second (carry) return value of" idc 1 (* boolean carry/borrow *) e2v (snd rout)); + inl (fun retptr => [Call (idc' @@ (retptr, (e1, e2, e3)))%Cexpr])) + | Some _, _ => inr ["Unrecognized identifier when attempting to construct an assignment with 2 arguments: " ++ show true idc]%string + | None, _ => inr ["Expression is not a literal power of two of type ℤ: " ++ show true s ++ " (when trying to parse the first argument of " ++ show true idc ++ ")"]%string + end. + + Let recognize_2ref_ident + {t} + : forall (idc : ident.ident t) + (rout : option int.type * option int.type) + (args : type.for_each_lhs_of_arrow (fun t => @Compilers.expr.expr base.type ident.ident var_data t * (arith_expr type.Z * option int.type))%type t), + ErrT (arith_expr (type.Zptr * type.Zptr) -> expr) + := match t with + | (type.base base.type.Z -> type.base base.type.Z -> type.base base.type.Z -> type.base (base.type.Z * base.type.Z))%etype + => recognize_3arg_2ref_ident + | (type.base base.type.Z -> type.base base.type.Z -> type.base base.type.Z -> type.base base.type.Z -> type.base (base.type.Z * base.type.Z))%etype + => recognize_4arg_2ref_ident + | _ => fun idc rout args => inr ["Unrecognized type for function call: " ++ show false t ++ " (when trying to handle the identifer " ++ show false idc ++ ")"]%string + end. + + Definition declare_name + (descr : string) + (count : positive) + (make_name : positive -> option string) + (r : option int.type) + : ErrT (expr * string * arith_expr type.Zptr) + := (n ,, r <- (match make_name count with + | Some n => ret n + | None => inr ["Variable naming-function does not support the index " ++ show false count]%string + end, + match r with + | Some r => ret r + | None => inr ["Missing return type annotation for " ++ descr]%string + end); + ret ([DeclareVar type.Z (Some r) n], n, (Addr @@ Var type.Z n)%Cexpr)). + + Let make_assign_arg_1ref_opt + (e : @Compilers.expr.expr base.type ident.ident var_data base.type.Z) + (count : positive) + (make_name : positive -> option string) + : ErrT (expr * var_data base.type.Z) + := let _ := @PHOAS.expr.partially_show_expr in + let e1 := e in + let '(rout, e) := match invert_Z_cast e with + | Some (r, e) => (Some (int.of_zrange_relaxed r), e) + | None => (None, e) + end%core in + let res_ref + := match invert_AppIdent_curried e with + | Some (existT _ (idc, args)) + => args <- arith_expr_of_PHOAS_args args; + idce <- recognize_1ref_ident idc rout args; + v <- declare_name (show false idc) count make_name rout; + let '(decls, n, retv) := v in + ret ((decls ++ (idce retv))%list, (n, rout)) + | None => inr ["Invalid assignment of non-identifier-application: " ++ show false e]%string + end in + match res_ref with + | inl res => ret res + | inr _ + => e1 <- arith_expr_of_base_PHOAS e1 None; + let '(e1, r1) := e1 in + match make_name count with + | Some n1 + => ret ([Assign true type.Z r1 n1 e1], (n1, r1)) + | None => inr ["Variable naming-function does not support the index " ++ show false count]%string + end + end. + + Let make_assign_arg_2ref + (e : @Compilers.expr.expr base.type ident.ident var_data (base.type.Z * base.type.Z)) + (count : positive) + (make_name : positive -> option string) + : ErrT (expr * var_data (base.type.Z * base.type.Z)) + := let _ := @PHOAS.expr.partially_show_expr in + let '((rout1, rout2), e) + := match invert_Z_cast2 e with + | Some ((r1, r2), e) => ((Some (int.of_zrange_relaxed r1), Some (int.of_zrange_relaxed r2)), e) + | None => ((None, None), e) + end%core in + match invert_AppIdent_curried e with + | Some (existT t (idc, args)) + => args <- arith_expr_of_PHOAS_args args; + idce <- recognize_2ref_ident idc (rout1, rout2) args; + v1,, v2 <- (declare_name (show false idc) count make_name rout1, + declare_name (show false idc) (Pos.succ count) make_name rout2); + let '(decls1, n1, retv1) := v1 in + let '(decls2, n2, retv2) := v2 in + ret (decls1 ++ decls2 ++ (idce (retv1, retv2)%Cexpr), ((n1, rout1), (n2, rout2)))%list + | None => inr ["Invalid assignment of non-identifier-application: " ++ show false e]%string + end. + + Let make_assign_arg_ref + {t} + : forall (e : @Compilers.expr.expr base.type ident.ident var_data t) + (count : positive) + (make_name : positive -> option string), + ErrT (expr * var_data t) + := let _ := @PHOAS.expr.partially_show_expr in + match t with + | type.base base.type.Z => make_assign_arg_1ref_opt + | type.base (base.type.Z * base.type.Z)%etype => make_assign_arg_2ref + | _ => fun e _ _ => inr ["Invalid type of assignment expression: " ++ show false t ++ " (with expression " ++ show true e ++ ")"] + end. + + Fixpoint size_of_type (t : base.type) : positive + := match t with + | base.type.type_base t => 1 + | base.type.prod A B => size_of_type A + size_of_type B + | base.type.list A => 1 + end%positive. + + Let make_uniform_assign_expr_of_PHOAS + {s} (e1 : @Compilers.expr.expr base.type ident.ident var_data s) + {d} (e2 : var_data s -> var_data d -> ErrT expr) + (count : positive) + (make_name : positive -> option string) + (vd : var_data d) + : ErrT expr + := let _ := @PHOAS.expr.partially_show_expr in (* for TC resolution *) + e1_vs <- make_assign_arg_ref e1 count make_name; + let '(e1, vs) := e1_vs in + e2 <- e2 vs vd; + ret (e1 ++ e2)%list. + + (* See above comment about extraction issues *) + Definition make_assign_expr_of_PHOAS + {s} (e1 : @Compilers.expr.expr base.type ident.ident var_data s) + {s' d} (e2 : var_data s' -> var_data d -> ErrT expr) + (count : positive) + (make_name : positive -> option string) + (v : var_data d) + : ErrT expr + := Eval cbv beta iota delta [bind2_err bind3_err bind4_err bind5_err recognize_1ref_ident recognize_3arg_2ref_ident recognize_4arg_2ref_ident recognize_2ref_ident make_assign_arg_1ref_opt make_assign_arg_2ref make_assign_arg_ref make_uniform_assign_expr_of_PHOAS make_uniform_assign_expr_of_PHOAS] in + match type.try_transport base.try_make_transport_cps _ _ s' e1 with + | Some e1 => make_uniform_assign_expr_of_PHOAS e1 e2 count make_name v + | None => inr [report_type_mismatch s' s] + end. + + Fixpoint expr_of_base_PHOAS + {t} + (e : @Compilers.expr.expr base.type ident.ident var_data t) (count : positive) (make_name : positive -> option string) - : ZRange.type.option.interp t -> option (positive * var_data t) - := match t with - | type.base t => base_var_data_of_bounds count make_name - | type.arrow s d => fun _ => None - end. + {struct e} + : forall (ret_val : var_data t), ErrT expr + := match e in expr.expr t return var_data t -> ErrT expr with + | expr.LetIn (type.base s) d e1 e2 + => make_assign_expr_of_PHOAS + e1 + (fun vs vd => @expr_of_base_PHOAS d (e2 vs) (size_of_type s + count)%positive make_name vd) + count make_name + | expr.LetIn (type.arrow _ _) _ _ _ as e + | expr.Var _ _ as e + | expr.Ident _ _ as e + | expr.App _ _ _ _ as e + | expr.Abs _ _ _ as e + => fun v => make_return_assignment_of_arith v e + end%expr_pat%option. + + Fixpoint base_var_data_of_bounds {t} + (count : positive) + (make_name : positive -> option string) + {struct t} + : ZRange.type.base.option.interp t -> option (positive * var_data t) + := match t return ZRange.type.base.option.interp t -> option (positive * var_data t) with + | base.type.Z + => fun r => (n <- make_name count; + Some (Pos.succ count, (n, option_map int.of_zrange_relaxed r))) + | base.type.prod A B + => fun '(ra, rb) + => (va <- @base_var_data_of_bounds A count make_name ra; + let '(count, va) := va in + vb <- @base_var_data_of_bounds B count make_name rb; + let '(count, vb) := vb in + Some (count, (va, vb))) + | base.type.list base.type.Z + => fun r + => (ls <- r; + n <- make_name count; + Some (Pos.succ count, + (n, + match List.map (option_map int.of_zrange_relaxed) ls with + | nil => None + | cons x xs + => List.fold_right + (fun r1 r2 => r1 <- r1; r2 <- r2; Some (int.union r1 r2)) + x + xs + end, + length ls))) + | base.type.unit + => fun _ => Some (count, tt) + | base.type.list _ + | base.type.type_base _ + => fun _ => None + end%option. + + Definition var_data_of_bounds {t} + (count : positive) + (make_name : positive -> option string) + : ZRange.type.option.interp t -> option (positive * var_data t) + := match t with + | type.base t => base_var_data_of_bounds count make_name + | type.arrow s d => fun _ => None + end. - Fixpoint expr_of_PHOAS - {t} - (e : @Compilers.expr.expr base.type ident.ident var_data t) - (make_name : positive -> option string) - (inbounds : type.for_each_lhs_of_arrow ZRange.type.option.interp t) - (outbounds : ZRange.type.option.interp (type.final_codomain t)) - (count : positive) - {struct t} - : option (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) - := match t return @Compilers.expr.expr base.type ident.ident var_data t -> type.for_each_lhs_of_arrow ZRange.type.option.interp t -> ZRange.type.option.interp (type.final_codomain t) -> option (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) with - | type.base t - => fun e tt outbounds - => vd <- var_data_of_bounds count make_name outbounds; - let '(count, vd) := vd in - rv <- expr_of_base_PHOAS e count make_name vd; - Some (tt, vd, rv) - | type.arrow s d - => fun e '(inbound, inbounds) outbounds - => vs <- var_data_of_bounds count make_name inbound; - let '(count, vs) := vs in - f <- invert_Abs e; - ret <- @expr_of_PHOAS d (f vs) make_name inbounds outbounds count; - let '(vss, vd, rv) := ret in - Some (vs, vss, vd, rv) - end%option%core%expr e inbounds outbounds. - - Definition ExprOfPHOAS + Fixpoint expr_of_PHOAS' {t} - (e : @Compilers.expr.Expr base.type ident.ident t) - (name_list : option (list string)) + (e : @Compilers.expr.expr base.type ident.ident var_data t) + (make_in_name : positive -> option string) + (make_name : positive -> option string) (inbounds : type.for_each_lhs_of_arrow ZRange.type.option.interp t) - : option (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) - := (let outbounds := partial.Extract e inbounds in - let make_name := match name_list with - | None => fun p => Some ("x" ++ decimal_string_of_Z (Zpos p)) - | Some ls => fun p => List.nth_error ls (pred (Pos.to_nat p)) - end in - expr_of_PHOAS (e _) make_name inbounds outbounds 1). - End OfPHOAS. + (out_data : var_data (type.final_codomain t)) + (count : positive) + (in_to_body_count : positive -> positive) + {struct t} + : ErrT (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) + := let _ := @PHOAS.expr.partially_show_expr in (* for TC resolution *) + match t return @Compilers.expr.expr base.type ident.ident var_data t -> type.for_each_lhs_of_arrow ZRange.type.option.interp t -> var_data (type.final_codomain t) -> ErrT (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) with + | type.base t + => fun e tt vd + => rv <- expr_of_base_PHOAS e (in_to_body_count count) make_name vd; + ret (tt, vd, rv) + | type.arrow s d + => fun e '(inbound, inbounds) vd + => match var_data_of_bounds count make_in_name inbound, invert_Abs e with + | Some (count, vs), Some f + => retv <- @expr_of_PHOAS' d (f vs) make_in_name make_name inbounds vd count in_to_body_count; + let '(vss, vd, rv) := retv in + ret (vs, vss, vd, rv) + | None, _ => inr ["Unable to bind names for all arguments and bounds at type " ++ show false s]%string%list + | _, None => inr ["Invalid non-λ expression of arrow type (" ++ show false t ++ "): " ++ show true e]%string%list + end + end%core%expr e inbounds out_data. + + Definition expr_of_PHOAS + {t} + (e : @Compilers.expr.expr base.type ident.ident var_data t) + (make_in_name : positive -> option string) + (make_out_name : positive -> option string) + (make_name : positive -> option string) + (inbounds : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (outbounds : ZRange.type.option.interp (type.final_codomain t)) + (count : positive) + (in_to_body_count out_to_in_count : positive -> positive) + : ErrT (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) + := match var_data_of_bounds count make_out_name outbounds with + | Some vd + => let '(count, vd) := vd in + let count := out_to_in_count count in + @expr_of_PHOAS' t e make_in_name make_name inbounds vd count in_to_body_count + | None => inr ["Unable to bind names for all return arguments and bounds at type " ++ show false (type.final_codomain t)]%string + end. - Module primitive. - Definition small_enough (v : Z) : bool - := Z.log2_up (Z.abs v + 1) <=? 128. - Definition to_UL_postfix (r : zrange) : string - := let lower := lower r in - let upper := upper r in - let u := (if lower >=? 0 then "U" else "") in - let sz := Z.log2_up (Z.max (Z.abs upper + 1) (Z.abs lower)) in - if sz <=? 32 - then "" - else if sz <=? 64 - then u ++ "L" - else if sz <=? 128 - then u ++ "LL" - else " /* " ++ HexString.of_Z lower ++ " <= val <= " ++ HexString.of_Z upper ++ " */". - - Definition to_string {t : type.primitive} (v : BinInt.Z) : string - := match t with - | type.Z => HexString.of_Z v ++ (if small_enough v - then to_UL_postfix r[v~>v] - else "ℤ") - | type.Zptr => "#error ""literal address " ++ HexString.of_Z v ++ """;" - end. - End primitive. + Definition ExprOfPHOAS + {t} + (e : @Compilers.expr.Expr base.type ident.ident t) + (name_list : option (list string)) + (inbounds : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : ErrT (type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) + := (let outbounds := partial.Extract e inbounds in + let make_name_gen prefix := match name_list with + | None => fun p => Some (prefix ++ decimal_string_of_Z (Zpos p)) + | Some ls => fun p => List.nth_error ls (pred (Pos.to_nat p)) + end in + let make_in_name := make_name_gen "arg" in + let make_out_name := make_name_gen "out" in + let make_name := make_name_gen "x" in + let reset_if_names_given := match name_list with + | Some _ => fun p : positive => p + | None => fun _ : positive => 1%positive + end in + expr_of_PHOAS (e _) make_in_name make_out_name make_name inbounds outbounds 1 reset_if_names_given reset_if_names_given). + End with_bind. + End OfPHOAS. Module String. - Definition typedef_header : list string - := ["typedef unsigned char uint1_t;"]%string. + Definition typedef_header (prefix : string) (bitwidths_used : PositiveSet.t) + : list string + := (["#include "] + ++ (if PositiveSet.mem 1 bitwidths_used + then ["typedef unsigned char " ++ prefix ++ "uint1;"; + "typedef signed char " ++ prefix ++ "int1;"]%string + else []) + ++ (if PositiveSet.mem 128 bitwidths_used + then ["typedef signed __int128 " ++ prefix ++ "int128;"; + "typedef unsigned __int128 " ++ prefix ++ "uint128;"]%string + else []))%list. + + Definition stdint_bitwidths : list Z := [8; 16; 32; 64]. + Definition is_special_bitwidth (bw : Z) := negb (existsb (Z.eqb bw) stdint_bitwidths). + Module int. Module type. - Definition to_string (t : int.type) : string - := ((if int.is_unsigned t then "u" else "") + Definition to_string (prefix : string) (t : int.type) : string + := ((if is_special_bitwidth (int.bitwidth_of t) then prefix else "") + ++ (if int.is_unsigned t then "u" else "") ++ "int" ++ decimal_string_of_Z (int.bitwidth_of t) - ++ "_t")%string. + ++ (if is_special_bitwidth (int.bitwidth_of t) then "" else "_t"))%string. + Definition to_literal_macro_string (t : int.type) : option string + := if Z.ltb (int.bitwidth_of t) 8 + then None + else Some ((if int.is_unsigned t then "U" else "") + ++ "INT" + ++ decimal_string_of_Z (int.bitwidth_of t) + ++ "_C")%string. End type. End int. Module type. Module primitive. - Definition to_string (t : type.primitive) (r : option int.type) : string + Definition to_string (prefix : string) (t : type.primitive) (r : option int.type) : string := match r with - | Some int_t => int.type.to_string int_t + | Some int_t => int.type.to_string prefix int_t | None => "ℤ" end ++ match t with | type.Zptr => "*" @@ -1251,183 +1552,333 @@ Module Compilers. End type. End String. - Fixpoint arith_to_string {t} (e : arith_expr t) : string + Module primitive. + Definition to_string (prefix : string) (t : type.primitive) (v : BinInt.Z) : string + := match t, String.int.type.to_literal_macro_string (int.of_zrange_relaxed r[v ~> v]) with + | type.Z, Some macro + => macro ++ "(" ++ HexString.of_Z v ++ ")" + | type.Z, None => HexString.of_Z v + | type.Zptr, _ => "#error ""literal address " ++ HexString.of_Z v ++ """;" + end. + End primitive. + + Fixpoint arith_to_string (prefix : string) {t} (e : arith_expr t) : string := match e with - | (literal v @@ _) => primitive.to_string (t:=type.Z) v + | (literal v @@ _) => primitive.to_string prefix type.Z v | (List_nth n @@ Var _ v) => "(" ++ v ++ "[" ++ decimal_string_of_Z (Z.of_nat n) ++ "])" | (Addr @@ Var _ v) => "&" ++ v - | (Dereference @@ e) => "( *" ++ @arith_to_string _ e ++ " )" + | (Dereference @@ e) => "( *" ++ @arith_to_string prefix _ e ++ " )" | (Z_shiftr offset @@ e) - => "(" ++ @arith_to_string _ e ++ " >> " ++ decimal_string_of_Z offset ++ ")" + => "(" ++ @arith_to_string prefix _ e ++ " >> " ++ decimal_string_of_Z offset ++ ")" | (Z_shiftl offset @@ e) - => "(" ++ @arith_to_string _ e ++ " << " ++ decimal_string_of_Z offset ++ ")" - | (Z_land mask @@ e) - => "(" ++ @arith_to_string _ e ++ " & " ++ primitive.to_string (t:=type.Z) mask ++ ")" + => "(" ++ @arith_to_string prefix _ e ++ " << " ++ decimal_string_of_Z offset ++ ")" + | (Z_land @@ (e1, e2)) + => "(" ++ @arith_to_string prefix _ e1 ++ " & " ++ @arith_to_string prefix _ e2 ++ ")" + | (Z_lor @@ (e1, e2)) + => "(" ++ @arith_to_string prefix _ e1 ++ " | " ++ @arith_to_string prefix _ e2 ++ ")" | (Z_add @@ (x1, x2)) - => "(" ++ @arith_to_string _ x1 ++ " + " ++ @arith_to_string _ x2 ++ ")" + => "(" ++ @arith_to_string prefix _ x1 ++ " + " ++ @arith_to_string prefix _ x2 ++ ")" | (Z_mul @@ (x1, x2)) - => "(" ++ @arith_to_string _ x1 ++ " * " ++ @arith_to_string _ x2 ++ ")" + => "(" ++ @arith_to_string prefix _ x1 ++ " * " ++ @arith_to_string prefix _ x2 ++ ")" | (Z_sub @@ (x1, x2)) - => "(" ++ @arith_to_string _ x1 ++ " - " ++ @arith_to_string _ x2 ++ ")" + => "(" ++ @arith_to_string prefix _ x1 ++ " - " ++ @arith_to_string prefix _ x2 ++ ")" | (Z_opp @@ e) - => "(-" ++ @arith_to_string _ e ++ ")" - | (Z_mul_split lg2s @@ (x1, x2, x3)) - => "_mulx_u" - ++ decimal_string_of_Z lg2s ++ "(" - ++ @arith_to_string _ x1 ++ ", " - ++ @arith_to_string _ x2 ++ ", " - ++ @arith_to_string _ x3 ++ ")" - | (Z_add_get_carry lg2s @@ (x1, x2, x3)) - => "_add_carryx_u" - ++ decimal_string_of_Z lg2s ++ "(0, " - ++ @arith_to_string _ x1 ++ ", " - ++ @arith_to_string _ x2 ++ ", " - ++ @arith_to_string _ x3 ++ ")" - | (Z_add_with_get_carry lg2s @@ (x1, x2, x3, x4)) - => "_add_carryx_u" - ++ decimal_string_of_Z lg2s ++ "(" - ++ @arith_to_string _ x1 ++ ", " - ++ @arith_to_string _ x2 ++ ", " - ++ @arith_to_string _ x3 ++ ", " - ++ @arith_to_string _ x4 ++ ")" - | (Z_sub_get_borrow lg2s @@ (x1, x2, x3)) - => "_subborrow_u" - ++ decimal_string_of_Z lg2s ++ "(0, " - ++ @arith_to_string _ x1 ++ ", " - ++ @arith_to_string _ x2 ++ ", " - ++ @arith_to_string _ x3 ++ ")" - | (Z_sub_with_get_borrow lg2s @@ (x1, x2, x3, x4)) - => "_subborrow_u" - ++ decimal_string_of_Z lg2s ++ "(" - ++ @arith_to_string _ x1 ++ ", " - ++ @arith_to_string _ x2 ++ ", " - ++ @arith_to_string _ x3 ++ ", " - ++ @arith_to_string _ x4 ++ ")" - | (Z_zselect @@ (cond, et, ef)) => "#error zselect;" + => "(-" ++ @arith_to_string prefix _ e ++ ")" + | (Z_lnot _ @@ e) + => "(~" ++ @arith_to_string prefix _ e ++ ")" + | (Z_bneg @@ e) + => "(!" ++ @arith_to_string prefix _ e ++ ")" + | (Z_mul_split lg2s @@ args) + => prefix + ++ "mulx_u" + ++ decimal_string_of_Z lg2s ++ "(" ++ @arith_to_string prefix _ args ++ ")" + | (Z_add_with_get_carry lg2s @@ args) + => prefix + ++ "addcarryx_u" + ++ decimal_string_of_Z lg2s ++ "(" ++ @arith_to_string prefix _ args ++ ")" + | (Z_sub_with_get_borrow lg2s @@ args) + => prefix + ++ "subborrowx_u" + ++ decimal_string_of_Z lg2s ++ "(" ++ @arith_to_string prefix _ args ++ ")" + | (Z_zselect ty @@ args) + => prefix + ++ "cmovznz_" + ++ (if int.is_unsigned ty then "u" else "") + ++ decimal_string_of_Z (int.bitwidth_of ty) ++ "(" ++ @arith_to_string prefix _ args ++ ")" | (Z_add_modulo @@ (x1, x2, x3)) => "#error addmodulo;" | (Z_static_cast int_t @@ e) - => "(" ++ String.type.primitive.to_string type.Z (Some int_t) ++ ")" - ++ @arith_to_string _ e + => "(" ++ String.type.primitive.to_string prefix type.Z (Some int_t) ++ ")" + ++ @arith_to_string prefix _ e | Var _ v => v + | Pair A B a b + => @arith_to_string prefix A a ++ ", " ++ @arith_to_string prefix B b | (List_nth _ @@ _) | (Addr @@ _) | (Z_add @@ _) | (Z_mul @@ _) | (Z_sub @@ _) - | (Z_mul_split _ @@ _) - | (Z_add_get_carry _ @@ _) - | (Z_add_with_get_carry _ @@ _) - | (Z_sub_get_borrow _ @@ _) - | (Z_sub_with_get_borrow _ @@ _) - | (Z_zselect @@ _) + | (Z_land @@ _) + | (Z_lor @@ _) | (Z_add_modulo @@ _) => "#error bad_arg;" - | Pair A B a b - => "#error pair;" | TT => "#error tt;" end%core%Cexpr. - Fixpoint stmt_to_string (e : stmt) : string + Fixpoint stmt_to_string (prefix : string) (e : stmt) : string := match e with + | Call val + => arith_to_string prefix val ++ ";" | Assign true t sz name val - => String.type.primitive.to_string t sz ++ " " ++ name ++ " = " ++ arith_to_string val ++ ";" + => String.type.primitive.to_string prefix t sz ++ " " ++ name ++ " = " ++ arith_to_string prefix val ++ ";" | Assign false _ sz name val - => name ++ " = " ++ arith_to_string val ++ ";" + => name ++ " = " ++ arith_to_string prefix val ++ ";" | AssignZPtr name sz val - => "*" ++ name ++ " = " ++ arith_to_string val ++ ";" + => "*" ++ name ++ " = " ++ arith_to_string prefix val ++ ";" | DeclareVar t sz name - => String.type.primitive.to_string t sz ++ " " ++ name ++ ";" + => String.type.primitive.to_string prefix t sz ++ " " ++ name ++ ";" | AssignNth name n val - => name ++ "[" ++ decimal_string_of_Z (Z.of_nat n) ++ "] = " ++ arith_to_string val ++ ";" + => name ++ "[" ++ decimal_string_of_Z (Z.of_nat n) ++ "] = " ++ arith_to_string prefix val ++ ";" + end. + Definition to_strings (prefix : string) (e : expr) : list string + := List.map (stmt_to_string prefix) e. + + Record ident_infos := + { bitwidths_used : PositiveSet.t; + addcarryx_lg_splits : PositiveSet.t; + mulx_lg_splits : PositiveSet.t; + cmovznz_bitwidths : PositiveSet.t }. + Definition ident_info_empty : ident_infos + := Build_ident_infos PositiveSet.empty PositiveSet.empty PositiveSet.empty PositiveSet.empty. + Definition ident_info_union (x y : ident_infos) : ident_infos + := let (x0, x1, x2, x3) := x in + let (y0, y1, y2, y3) := y in + Build_ident_infos + (PositiveSet.union x0 y0) + (PositiveSet.union x1 y1) + (PositiveSet.union x2 y2) + (PositiveSet.union x3 y3). + Definition ident_info_of_bitwidths_used (v : PositiveSet.t) : ident_infos + := {| bitwidths_used := v; + addcarryx_lg_splits := PositiveSet.empty; + mulx_lg_splits := PositiveSet.empty; + cmovznz_bitwidths := PositiveSet.empty |}. + Definition ident_info_of_addcarryx (v : PositiveSet.t) : ident_infos + := {| bitwidths_used := PositiveSet.empty; + addcarryx_lg_splits := v; + mulx_lg_splits := PositiveSet.empty; + cmovznz_bitwidths := PositiveSet.empty |}. + Definition ident_info_of_mulx (v : PositiveSet.t) : ident_infos + := {| bitwidths_used := PositiveSet.empty; + addcarryx_lg_splits := PositiveSet.empty; + mulx_lg_splits := v; + cmovznz_bitwidths := PositiveSet.empty |}. + Definition ident_info_of_cmovznz (v : PositiveSet.t) : ident_infos + := {| bitwidths_used := PositiveSet.empty; + addcarryx_lg_splits := PositiveSet.empty; + mulx_lg_splits := PositiveSet.empty; + cmovznz_bitwidths := v |}. + + Definition collect_bitwidths_of_int_type (t : int.type) : PositiveSet.t + := PositiveSet.add (Z.to_pos (int.bitwidth_of t)) PositiveSet.empty. + Definition collect_infos_of_ident {s d} (idc : ident s d) : ident_infos + := match idc with + | Z_static_cast ty => ident_info_of_bitwidths_used (collect_bitwidths_of_int_type ty) + | Z_mul_split lg2s + => ident_info_of_mulx (PositiveSet.add (Z.to_pos lg2s) PositiveSet.empty) + | Z_add_with_get_carry lg2s + | Z_sub_with_get_borrow lg2s + => ident_info_of_addcarryx (PositiveSet.add (Z.to_pos lg2s) PositiveSet.empty) + | Z_zselect ty + => ident_info_of_cmovznz (collect_bitwidths_of_int_type ty) + | literal _ + | List_nth _ + | Addr + | Dereference + | Z_shiftr _ + | Z_shiftl _ + | Z_land + | Z_lor + | Z_add + | Z_mul + | Z_sub + | Z_opp + | Z_bneg + | Z_lnot _ + | Z_add_modulo + => ident_info_empty + end. + Fixpoint collect_infos_of_arith_expr {t} (e : arith_expr t) : ident_infos + := match e with + | AppIdent s d idc arg => ident_info_union (collect_infos_of_ident idc) (@collect_infos_of_arith_expr _ arg) + | Var t v => ident_info_empty + | Pair A B a b => ident_info_union (@collect_infos_of_arith_expr _ a) (@collect_infos_of_arith_expr _ b) + | TT => ident_info_empty + end. + + Fixpoint collect_infos_of_stmt (e : stmt) : ident_infos + := match e with + | Assign _ _ (Some sz) _ val + | AssignZPtr _ (Some sz) val + => ident_info_union (ident_info_of_bitwidths_used (collect_bitwidths_of_int_type sz)) (collect_infos_of_arith_expr val) + | Call val + | Assign _ _ None _ val + | AssignZPtr _ None val + | AssignNth _ _ val + => collect_infos_of_arith_expr val + | DeclareVar _ (Some sz) _ + => ident_info_of_bitwidths_used (collect_bitwidths_of_int_type sz) + | DeclareVar _ None _ + => ident_info_empty end. - Definition to_strings (e : expr) : list string - := List.map stmt_to_string e. + + Fixpoint collect_infos (e : expr) : ident_infos + := fold_right + ident_info_union + ident_info_empty + (List.map + collect_infos_of_stmt + e). Import OfPHOAS. - Fixpoint to_base_arg_list {t} : base_var_data t -> list string + Fixpoint to_base_arg_list (prefix : string) {t} : base_var_data t -> list string := match t return base_var_data t -> _ with | base.type.Z - => fun '(n, r) => [String.type.primitive.to_string type.Z r ++ " " ++ n] + => fun '(n, r) => [String.type.primitive.to_string prefix type.Z r ++ " " ++ n] | base.type.prod A B - => fun '(va, vb) => (@to_base_arg_list A va ++ @to_base_arg_list B vb)%list + => fun '(va, vb) => (@to_base_arg_list prefix A va ++ @to_base_arg_list prefix B vb)%list | base.type.list base.type.Z - => fun '(n, r, len) => [String.type.primitive.to_string type.Z r ++ " " ++ n ++ "[" ++ decimal_string_of_Z (Z.of_nat len) ++ "]"] + => fun '(n, r, len) => ["const " ++ String.type.primitive.to_string prefix type.Z r ++ " " ++ n ++ "[" ++ decimal_string_of_Z (Z.of_nat len) ++ "]"] | base.type.list _ => fun _ => ["#error ""complex list"";"] | base.type.unit => fun _ => ["#error unit;"] | base.type.nat => fun _ => ["#error ℕ;"] | base.type.bool => fun _ => ["#error bool;"] end. - Definition to_arg_list {t} : var_data t -> list string + Definition to_arg_list (prefix : string) {t} : var_data t -> list string := match t return var_data t -> _ with - | type.base t => to_base_arg_list + | type.base t => to_base_arg_list prefix | type.arrow _ _ => fun _ => ["#error arrow;"] end. - Fixpoint to_arg_list_for_each_lhs_of_arrow {t} : type.for_each_lhs_of_arrow var_data t -> list string + Fixpoint to_arg_list_for_each_lhs_of_arrow (prefix : string) {t} : type.for_each_lhs_of_arrow var_data t -> list string := match t return type.for_each_lhs_of_arrow var_data t -> _ with | type.base t => fun _ => nil | type.arrow s d => fun '(x, xs) - => to_arg_list x ++ @to_arg_list_for_each_lhs_of_arrow d xs + => to_arg_list prefix x ++ @to_arg_list_for_each_lhs_of_arrow prefix d xs end%list. - Fixpoint to_base_retarg_list {t} : base_var_data t -> list string + Fixpoint to_base_retarg_list prefix {t} : base_var_data t -> list string := match t return base_var_data t -> _ with | base.type.Z - => fun '(n, r) => [String.type.primitive.to_string type.Zptr r ++ " " ++ n] + => fun '(n, r) => [String.type.primitive.to_string prefix type.Zptr r ++ " " ++ n] | base.type.prod A B - => fun '(va, vb) => (@to_base_retarg_list A va ++ @to_base_retarg_list B vb)%list + => fun '(va, vb) => (@to_base_retarg_list prefix A va ++ @to_base_retarg_list prefix B vb)%list | base.type.list base.type.Z - => fun '(n, r, len) => [String.type.primitive.to_string type.Z r ++ " " ++ n ++ "[" ++ decimal_string_of_Z (Z.of_nat len) ++ "]"] + => fun '(n, r, len) => [String.type.primitive.to_string prefix type.Z r ++ " " ++ n ++ "[" ++ decimal_string_of_Z (Z.of_nat len) ++ "]"] | base.type.list _ => fun _ => ["#error ""complex list"";"] | base.type.unit => fun _ => ["#error unit;"] | base.type.nat => fun _ => ["#error ℕ;"] | base.type.bool => fun _ => ["#error bool;"] end. - Definition to_retarg_list {t} : var_data t -> list string + Definition to_retarg_list (prefix : string) {t} : var_data t -> list string := match t return var_data t -> _ with - | type.base _ => to_base_arg_list + | type.base _ => to_base_retarg_list prefix | type.arrow _ _ => fun _ => ["#error arrow;"] end. - Definition to_function_lines (name : string) + Fixpoint bound_to_string {t : base.type} : var_data t -> ZRange.type.base.option.interp t -> list string + := match t return var_data t -> ZRange.type.base.option.interp t -> list string with + | base.type.Z + => fun '(name, _) arg + => [(name ++ ": ") + ++ match arg with + | Some arg => show false arg + | None => show false arg + end]%string + | base.type.prod A B + => fun '(va, vb) '(a, b) + => @bound_to_string A va a ++ @bound_to_string B vb b + | base.type.list A + => fun '(name, _, _) arg + => [(name ++ ": ") + ++ match ZRange.type.base.option.lift_Some arg with + | Some arg => show false arg + | None => show false arg + end]%string + | base.type.unit + | base.type.bool + | base.type.nat + => fun _ _ => nil + end%list. + + Fixpoint input_bounds_to_string {t} : type.for_each_lhs_of_arrow var_data t -> type.for_each_lhs_of_arrow ZRange.type.option.interp t -> list string + := match t return type.for_each_lhs_of_arrow var_data t -> type.for_each_lhs_of_arrow ZRange.type.option.interp t -> list string with + | type.base t => fun _ _ => nil + | type.arrow (type.base s) d + => fun '(v, vs) '(arg, args) + => (bound_to_string v arg) + ++ @input_bounds_to_string d vs args + | type.arrow s d + => fun '(absurd, _) => match absurd : Empty_set with end + end%list. + + Definition to_function_lines (prefix : string) (name : string) {t} + (f : type.for_each_lhs_of_arrow var_data t * var_data (type.final_codomain t) * expr) : list string := let '(args, rets, body) := f in (((("void " ++ name ++ "(" - ++ (String.concat ", " (to_arg_list_for_each_lhs_of_arrow args ++ to_retarg_list rets)) + ++ (String.concat ", " (to_retarg_list prefix rets ++ to_arg_list_for_each_lhs_of_arrow prefix args)) ++ ") {")%string) - :: (List.map (fun s => " " ++ s)%string (to_strings body))) + :: (List.map (fun s => " " ++ s)%string (to_strings prefix body))) ++ ["}"])%list. - Definition ToFunctionLines (name : string) + Definition ToFunctionLines (prefix : string) (name : string) {t} (e : @Compilers.expr.Expr base.type ident.ident t) (name_list : option (list string)) (inbounds : type.for_each_lhs_of_arrow ZRange.type.option.interp t) - : option (list string) - := (f <- ExprOfPHOAS e name_list inbounds; - Some (to_function_lines name f)). + (outbounds : ZRange.type.base.option.interp (type.final_codomain t)) + : (list string * ident_infos) + string + := match ExprOfPHOAS e name_list inbounds with + | inl (indata, outdata, f) + => inl ((["/*"; " * Input Bounds:"] + ++ List.map (fun v => " * " ++ v)%string (input_bounds_to_string indata inbounds) + ++ [" * Output Bounds:"] + ++ List.map (fun v => " * " ++ v)%string (bound_to_string outdata outbounds) + ++ [" */"] + ++ to_function_lines prefix name (indata, outdata, f))%list, + collect_infos f) + | inr nil + => inr ("Unknown internal error in converting " ++ name ++ " to C")%string + | inr [err] + => inr ("Error in converting " ++ name ++ " to C:" ++ String.NewLine ++ err)%string + | inr errs + => inr ("Errors in converting " ++ name ++ " to C:" ++ String.NewLine ++ String.concat String.NewLine errs)%string + end. Definition LinesToString (lines : list string) : string := String.concat String.NewLine lines. - Definition ToFunctionString (name : string) + Definition ToFunctionString (prefix : string) (name : string) {t} (e : @Compilers.expr.Expr base.type ident.ident t) (name_list : option (list string)) (inbounds : type.for_each_lhs_of_arrow ZRange.type.option.interp t) - : option string - := (ls <- ToFunctionLines name e name_list inbounds; - Some (LinesToString ls)). + (outbounds : ZRange.type.option.interp (type.final_codomain t)) + : (string * ident_infos) + string + := match ToFunctionLines prefix name e name_list inbounds outbounds with + | inl (ls, used_types) => inl (LinesToString ls, used_types) + | inr err => inr err + end. End C. Notation ToFunctionLines := C.ToFunctionLines. Notation ToFunctionString := C.ToFunctionString. diff --git a/src/Experiments/NewPipeline/GENERATEDIdentifiersWithoutTypes.v b/src/Experiments/NewPipeline/GENERATEDIdentifiersWithoutTypes.v index 94257f793..589b5ad26 100644 --- a/src/Experiments/NewPipeline/GENERATEDIdentifiersWithoutTypes.v +++ b/src/Experiments/NewPipeline/GENERATEDIdentifiersWithoutTypes.v @@ -9,396 +9,550 @@ Module Compilers. Module pattern. Module ident. Set Boolean Equality Schemes. - (*Print Compilers.ident.ident.*) - (*Show Match Compilers.ident.ident.*) + (* + Set Printing Coercions. + Redirect "/tmp/pr" Print Compilers.ident.ident. + Redirect "/tmp/sm" Show Match Compilers.ident.ident. + *) (* <<< #!/usr/bin/env python2 -import re +import re, os -print_ident = r"""Inductive ident : Compilers.type Compilers.base.type.type -> Set := +print_ident = r"""Inductive ident : defaults.type -> Set := Literal : forall t : base.type.base, - base.interp (Compilers.base.type.type_base t) -> - ident ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base t)) + base.interp (base.type.type_base t) -> + ident + ((fun x : base.type => type.base x) (base.type.type_base t)) | Nat_succ : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) | Nat_pred : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) | Nat_max : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) | Nat_mul : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) | Nat_add : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) | Nat_sub : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype - | nil : forall t : Compilers.base.type.type, ident ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list t)) - | cons : forall t : Compilers.base.type.type, + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) + | nil : forall t : base.type, + ident ((fun x : base.type => type.base x) (base.type.list t)) + | cons : forall t : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) t -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list t) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list t))%etype - | pair : forall A B : Compilers.base.type.type, + ((fun x : base.type => type.base x) t -> + (fun x : base.type => type.base x) (base.type.list t) -> + (fun x : base.type => type.base x) (base.type.list t)) + | pair : forall A B : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) B -> (fun x : Compilers.base.type.type => type.base x) (A * B)%etype)%etype - | fst : forall A B : Compilers.base.type.type, + ((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) B -> + (fun x : base.type => type.base x) (A * B)%etype) + | fst : forall A B : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (A * B)%etype -> (fun x : Compilers.base.type.type => type.base x) A)%etype - | snd : forall A B : Compilers.base.type.type, + ((fun x : base.type => type.base x) (A * B)%etype -> + (fun x : base.type => type.base x) A) + | snd : forall A B : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (A * B)%etype -> (fun x : Compilers.base.type.type => type.base x) B)%etype - | pair_rect : forall A B T : Compilers.base.type.type, + ((fun x : base.type => type.base x) (A * B)%etype -> + (fun x : base.type => type.base x) B) + | pair_rect : forall A B T : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) B -> (fun x : Compilers.base.type.type => type.base x) T) -> - (fun x : Compilers.base.type.type => type.base x) (A * B)%etype -> - (fun x : Compilers.base.type.type => type.base x) T)%etype - | bool_rect : forall T : Compilers.base.type.type, + (((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) B -> + (fun x : base.type => type.base x) T) -> + (fun x : base.type => type.base x) (A * B)%etype -> + (fun x : base.type => type.base x) T) + | bool_rect : forall T : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) ()%etype -> (fun x : Compilers.base.type.type => type.base x) T) -> - ((fun x : Compilers.base.type.type => type.base x) ()%etype -> (fun x : Compilers.base.type.type => type.base x) T) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.bool) -> - (fun x : Compilers.base.type.type => type.base x) T)%etype - | nat_rect : forall P : Compilers.base.type.type, + (((fun x : base.type => type.base x) ()%etype -> + (fun x : base.type => type.base x) T) -> + ((fun x : base.type => type.base x) ()%etype -> + (fun x : base.type => type.base x) T) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.bool) -> + (fun x : base.type => type.base x) T) + | nat_rect : forall P : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) ()%etype -> (fun x : Compilers.base.type.type => type.base x) P) -> - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) P -> (fun x : Compilers.base.type.type => type.base x) P) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) P)%etype - | list_rect : forall A P : Compilers.base.type.type, + (((fun x : base.type => type.base x) ()%etype -> + (fun x : base.type => type.base x) P) -> + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) P -> + (fun x : base.type => type.base x) P) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) P) + | list_rect : forall A P : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) ()%etype -> (fun x : Compilers.base.type.type => type.base x) P) -> - ((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) P -> (fun x : Compilers.base.type.type => type.base x) P) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) P)%etype - | list_case : forall A P : Compilers.base.type.type, + (((fun x : base.type => type.base x) ()%etype -> + (fun x : base.type => type.base x) P) -> + ((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) P -> + (fun x : base.type => type.base x) P) -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) P) + | list_case : forall A P : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) ()%etype -> (fun x : Compilers.base.type.type => type.base x) P) -> - ((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) P) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) P)%etype - | List_length : forall T : Compilers.base.type.type, + (((fun x : base.type => type.base x) ()%etype -> + (fun x : base.type => type.base x) P) -> + ((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) P) -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) P) + | List_length : forall T : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list T) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat))%etype + ((fun x : base.type => type.base x) (base.type.list T) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) | List_seq : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.list (Compilers.base.type.type_base base.type.nat)))%etype - | List_repeat : forall A : Compilers.base.type.type, + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.list (base.type.type_base base.type.nat))) + | List_repeat : forall A : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A))%etype - | List_combine : forall A B : Compilers.base.type.type, + ((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) (base.type.list A)) + | List_combine : forall A B : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list B) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list (A * B)))%etype - | List_map : forall A B : Compilers.base.type.type, + ((fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) (base.type.list B) -> + (fun x : base.type => type.base x) + (base.type.list (A * B))) + | List_map : forall A B : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) A -> (fun x : Compilers.base.type.type => type.base x) B) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list B))%etype - | List_app : forall A : Compilers.base.type.type, + (((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) B) -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) (base.type.list B)) + | List_app : forall A : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A))%etype - | List_rev : forall A : Compilers.base.type.type, + ((fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) (base.type.list A)) + | List_rev : forall A : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A))%etype - | List_flat_map : forall A B : Compilers.base.type.type, + ((fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) (base.type.list A)) + | List_flat_map : forall A B : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list B)) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list B))%etype - | List_partition : forall A : Compilers.base.type.type, + (((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) (base.type.list B)) -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) (base.type.list B)) + | List_partition : forall A : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.bool)) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list A) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.list A * Compilers.base.type.list A)%etype)%etype - | List_fold_right : forall A B : Compilers.base.type.type, + (((fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.bool)) -> + (fun x : base.type => type.base x) (base.type.list A) -> + (fun x : base.type => type.base x) + (base.type.list A * base.type.list A)%etype) + | List_fold_right : forall A B : base.type, ident - (((fun x : Compilers.base.type.type => type.base x) B -> - (fun x : Compilers.base.type.type => type.base x) A -> (fun x : Compilers.base.type.type => type.base x) A) -> - (fun x : Compilers.base.type.type => type.base x) A -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list B) -> - (fun x : Compilers.base.type.type => type.base x) A)%etype - | List_update_nth : forall T : Compilers.base.type.type, + (((fun x : base.type => type.base x) B -> + (fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) A) -> + (fun x : base.type => type.base x) A -> + (fun x : base.type => type.base x) + (base.type.list B) -> + (fun x : base.type => type.base x) A) + | List_update_nth : forall T : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - ((fun x : Compilers.base.type.type => type.base x) T -> (fun x : Compilers.base.type.type => type.base x) T) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list T) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list T))%etype - | List_nth_default : forall T : Compilers.base.type.type, + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + ((fun x : base.type => type.base x) T -> + (fun x : base.type => type.base x) T) -> + (fun x : base.type => type.base x) + (base.type.list T) -> + (fun x : base.type => type.base x) + (base.type.list T)) + | List_nth_default : forall T : base.type, ident - ((fun x : Compilers.base.type.type => type.base x) T -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.list T) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) T)%etype + ((fun x : base.type => type.base x) T -> + (fun x : base.type => type.base x) + (base.type.list T) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) T) | Z_add : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_mul : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_pow : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_sub : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_opp : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_div : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_modulo : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_log2 : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_log2_up : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_eqb : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.bool))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.bool)) | Z_leb : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.bool))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.bool)) + | Z_geb : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.bool)) | Z_of_nat : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.nat) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype - | Z_shiftr : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype - | Z_shiftl : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype - | Z_land : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.nat) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_to_nat : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.nat)) + | Z_shiftr : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_shiftl : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_land : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_lor : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_bneg : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + | Z_lnot_modulo : ident + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_mul_split : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_mul_split_concrete : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | Z_add_get_carry : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_add_get_carry_concrete : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | Z_add_with_carry : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_add_with_get_carry : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_add_with_get_carry_concrete : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | Z_sub_get_borrow : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_sub_get_borrow_concrete : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | Z_sub_with_get_borrow : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_sub_with_get_borrow_concrete : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | Z_zselect : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_add_modulo : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_rshi : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype - | Z_rshi_concrete : Z -> - Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_cc_m : ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype - | Z_cc_m_concrete : Z -> - ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype - | Z_neg_snd : ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_cast : zrange -> ident - ((fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z) -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z) -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | Z_cast2 : zrange * zrange -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | fancy_add : Z -> Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | fancy_addc : Z -> Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * - Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | fancy_sub : Z -> Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | fancy_subb : Z -> Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * - Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype) | fancy_mulll : Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_mullh : Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_mulhl : Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_mulhh : Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_rshi : Z -> Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_selc : ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * - Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_selm : Z -> ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * - Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_sell : ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * - Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) | fancy_addm : ident - ((fun x : Compilers.base.type.type => type.base x) - (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * - Compilers.base.type.type_base base.type.Z)%etype -> - (fun x : Compilers.base.type.type => type.base x) (Compilers.base.type.type_base base.type.Z))%etype + ((fun x : base.type => type.base x) + (base.type.type_base base.type.Z * + base.type.type_base base.type.Z * + base.type.type_base base.type.Z)%etype -> + (fun x : base.type => type.base x) + (base.type.type_base base.type.Z)) + """ show_match_ident = r"""match # with | ident.Literal t v => @@ -437,30 +591,29 @@ show_match_ident = r"""match # with | ident.Z_opp => | ident.Z_div => | ident.Z_modulo => + | ident.Z_log2 => + | ident.Z_log2_up => | ident.Z_eqb => | ident.Z_leb => + | ident.Z_geb => | ident.Z_of_nat => - | ident.Z_shiftr offset => - | ident.Z_shiftl offset => - | ident.Z_land mask => + | ident.Z_to_nat => + | ident.Z_shiftr => + | ident.Z_shiftl => + | ident.Z_land => + | ident.Z_lor => + | ident.Z_bneg => + | ident.Z_lnot_modulo => | ident.Z_mul_split => - | ident.Z_mul_split_concrete s => | ident.Z_add_get_carry => - | ident.Z_add_get_carry_concrete s => | ident.Z_add_with_carry => | ident.Z_add_with_get_carry => - | ident.Z_add_with_get_carry_concrete s => | ident.Z_sub_get_borrow => - | ident.Z_sub_get_borrow_concrete s => | ident.Z_sub_with_get_borrow => - | ident.Z_sub_with_get_borrow_concrete s => | ident.Z_zselect => | ident.Z_add_modulo => | ident.Z_rshi => - | ident.Z_rshi_concrete s offset => | ident.Z_cc_m => - | ident.Z_cc_m_concrete s => - | ident.Z_neg_snd => | ident.Z_cast range => | ident.Z_cast2 range => | ident.fancy_add log2wordmax imm => @@ -479,6 +632,13 @@ show_match_ident = r"""match # with end """ +remake = False +if remake: + assert(os.path.exists('/tmp/pr.out')) + assert(os.path.exists('/tmp/sm.out')) + with open('/tmp/pr.out', 'r') as f: print_ident = re.sub(r'^For.*\n', '', f.read(), flags=re.MULTILINE) + with open('/tmp/sm.out', 'r') as f: show_match_ident = f.read() + prefix = 'Compilers.' indent = ' ' exts = ('Unit', 'Z', 'Bool', 'Nat') @@ -500,7 +660,7 @@ ttypes = ([[] for ty in tys] + [get_dep_types(case) for case in print_ident.replace('\n', ' ').split('|')[1:]]) ctypes = ([['base.interp ' + ty] for ty in tys] - + [[i.strip() for i in re.sub(r'forall [^:]+ : %sbase.type.type,' % prefix, '', i[i.find(':')+1:i.find('ident')]).strip(' ->').split('->') if i.strip()] + + [[i.strip() for i in re.sub(r'forall [^:]+ : base.type,', '', i[i.find(':')+1:i.find('ident')]).strip(' ->').split('->') if i.strip()] for i in print_ident.replace('\n', ' ').split('|')[1:]]) crettypes = ([('%sident.ident (type.base (%sbase.type.type_base ' + ty + '))') % (prefix, prefix) for ty in tys] + [prefix + 'ident.' + re.sub(r'\(fun x : [^ ]+ => ([^ ]+) x\)', r'\1', re.sub(' +', ' ', i[i.find('ident'):])) @@ -517,18 +677,28 @@ Module Compilers. Module pattern. Module ident. Set Boolean Equality Schemes. - (*Print Compilers.ident.ident.*) - (*Show Match Compilers.ident.ident.*) + (""" + """* + Set Printing Coercions. + Redirect "/tmp/pr" Print Compilers.ident.ident. + Redirect "/tmp/sm" Show Match Compilers.ident.ident. + *""" + """) """ -def addnewline(s): return s + '\n' +def addnewline(s): return re.sub(r' +$', '', s + '\n', flags=re.MULTILINE) + +def get_updated_contents(): + contents = open(__file__).read() + contents = re.sub(r'^remake = True', r'remake = False', contents, flags=re.MULTILINE) + contents = re.sub(r'^print_ident = r""".*?"""', r'print_ident = r"""' + print_ident + r'"""', contents, flags=re.MULTILINE | re.DOTALL) + contents = re.sub(r'^show_match_ident = r""".*?"""', r'show_match_ident = r"""' + show_match_ident + r'"""', contents, flags=re.MULTILINE | re.DOTALL) + return contents retcode += addnewline(r"""%s(* <<< %s >>> %s*) -""" % (indent, open(__file__).read(), indent)) +""" % (indent, get_updated_contents(), indent)) retcode += addnewline(r"""%sInductive ident := %s| %s. """ % (indent, indent, ('\n' + indent + '| ').join(pctors))) @@ -628,7 +798,7 @@ def do_adjust_type(ctor, ctype): retcode += addnewline((r"""%sDefinition type_of (pidc : ident) : full_types pidc -> %stype %sbase.type := match pidc return full_types pidc -> _ with | %s - end. + end%%etype. """ % (indent, prefix, prefix, '\n | '.join(pctor + ' => ' + 'fun ' + ('_ => ' if len(ttype + ctype) == 0 else ((ctor[-1] + ' => ') if len(ttype + ctype) == 1 else "arg => let '(%s) := eta%d arg in " % (', '.join(ctor[-len(ttype + ctype):]), len(ttype + ctype)))) + cretty.replace(prefix + 'ident.ident ', '') @@ -710,30 +880,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp | Z_div | Z_modulo + | Z_log2 + | Z_log2_up | Z_eqb | Z_leb + | Z_geb | Z_of_nat + | Z_to_nat | Z_shiftr | Z_shiftl | Z_land + | Z_lor + | Z_bneg + | Z_lnot_modulo | Z_mul_split - | Z_mul_split_concrete | Z_add_get_carry - | Z_add_get_carry_concrete | Z_add_with_carry | Z_add_with_get_carry - | Z_add_with_get_carry_concrete | Z_sub_get_borrow - | Z_sub_get_borrow_concrete | Z_sub_with_get_borrow - | Z_sub_with_get_borrow_concrete | Z_zselect | Z_add_modulo | Z_rshi - | Z_rshi_concrete | Z_cc_m - | Z_cc_m_concrete - | Z_neg_snd | Z_cast | Z_cast2 | fancy_add @@ -791,30 +960,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp, Z_opp | Z_div, Z_div | Z_modulo, Z_modulo + | Z_log2, Z_log2 + | Z_log2_up, Z_log2_up | Z_eqb, Z_eqb | Z_leb, Z_leb + | Z_geb, Z_geb | Z_of_nat, Z_of_nat + | Z_to_nat, Z_to_nat | Z_shiftr, Z_shiftr | Z_shiftl, Z_shiftl | Z_land, Z_land + | Z_lor, Z_lor + | Z_bneg, Z_bneg + | Z_lnot_modulo, Z_lnot_modulo | Z_mul_split, Z_mul_split - | Z_mul_split_concrete, Z_mul_split_concrete | Z_add_get_carry, Z_add_get_carry - | Z_add_get_carry_concrete, Z_add_get_carry_concrete | Z_add_with_carry, Z_add_with_carry | Z_add_with_get_carry, Z_add_with_get_carry - | Z_add_with_get_carry_concrete, Z_add_with_get_carry_concrete | Z_sub_get_borrow, Z_sub_get_borrow - | Z_sub_get_borrow_concrete, Z_sub_get_borrow_concrete | Z_sub_with_get_borrow, Z_sub_with_get_borrow - | Z_sub_with_get_borrow_concrete, Z_sub_with_get_borrow_concrete | Z_zselect, Z_zselect | Z_add_modulo, Z_add_modulo | Z_rshi, Z_rshi - | Z_rshi_concrete, Z_rshi_concrete | Z_cc_m, Z_cc_m - | Z_cc_m_concrete, Z_cc_m_concrete - | Z_neg_snd, Z_neg_snd | Z_cast, Z_cast | Z_cast2, Z_cast2 | fancy_add, fancy_add @@ -870,30 +1038,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp, _ | Z_div, _ | Z_modulo, _ + | Z_log2, _ + | Z_log2_up, _ | Z_eqb, _ | Z_leb, _ + | Z_geb, _ | Z_of_nat, _ + | Z_to_nat, _ | Z_shiftr, _ | Z_shiftl, _ | Z_land, _ + | Z_lor, _ + | Z_bneg, _ + | Z_lnot_modulo, _ | Z_mul_split, _ - | Z_mul_split_concrete, _ | Z_add_get_carry, _ - | Z_add_get_carry_concrete, _ | Z_add_with_carry, _ | Z_add_with_get_carry, _ - | Z_add_with_get_carry_concrete, _ | Z_sub_get_borrow, _ - | Z_sub_get_borrow_concrete, _ | Z_sub_with_get_borrow, _ - | Z_sub_with_get_borrow_concrete, _ | Z_zselect, _ | Z_add_modulo, _ | Z_rshi, _ - | Z_rshi_concrete, _ | Z_cc_m, _ - | Z_cc_m_concrete, _ - | Z_neg_snd, _ | Z_cast, _ | Z_cast2, _ | fancy_add, _ @@ -911,7 +1078,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | fancy_addm, _ => fun T k => k None end%cps. - + Definition eta_ident_cps {T : Compilers.type Compilers.base.type -> Type} {t} (idc : Compilers.ident.ident t) (f : forall t', Compilers.ident.ident t' -> T t') : T t @@ -955,30 +1122,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.Z_opp => f _ Compilers.ident.Z_opp | Compilers.ident.Z_div => f _ Compilers.ident.Z_div | Compilers.ident.Z_modulo => f _ Compilers.ident.Z_modulo + | Compilers.ident.Z_log2 => f _ Compilers.ident.Z_log2 + | Compilers.ident.Z_log2_up => f _ Compilers.ident.Z_log2_up | Compilers.ident.Z_eqb => f _ Compilers.ident.Z_eqb | Compilers.ident.Z_leb => f _ Compilers.ident.Z_leb + | Compilers.ident.Z_geb => f _ Compilers.ident.Z_geb | Compilers.ident.Z_of_nat => f _ Compilers.ident.Z_of_nat - | Compilers.ident.Z_shiftr offset => f _ (@Compilers.ident.Z_shiftr offset) - | Compilers.ident.Z_shiftl offset => f _ (@Compilers.ident.Z_shiftl offset) - | Compilers.ident.Z_land mask => f _ (@Compilers.ident.Z_land mask) + | Compilers.ident.Z_to_nat => f _ Compilers.ident.Z_to_nat + | Compilers.ident.Z_shiftr => f _ Compilers.ident.Z_shiftr + | Compilers.ident.Z_shiftl => f _ Compilers.ident.Z_shiftl + | Compilers.ident.Z_land => f _ Compilers.ident.Z_land + | Compilers.ident.Z_lor => f _ Compilers.ident.Z_lor + | Compilers.ident.Z_bneg => f _ Compilers.ident.Z_bneg + | Compilers.ident.Z_lnot_modulo => f _ Compilers.ident.Z_lnot_modulo | Compilers.ident.Z_mul_split => f _ Compilers.ident.Z_mul_split - | Compilers.ident.Z_mul_split_concrete s => f _ (@Compilers.ident.Z_mul_split_concrete s) | Compilers.ident.Z_add_get_carry => f _ Compilers.ident.Z_add_get_carry - | Compilers.ident.Z_add_get_carry_concrete s => f _ (@Compilers.ident.Z_add_get_carry_concrete s) | Compilers.ident.Z_add_with_carry => f _ Compilers.ident.Z_add_with_carry | Compilers.ident.Z_add_with_get_carry => f _ Compilers.ident.Z_add_with_get_carry - | Compilers.ident.Z_add_with_get_carry_concrete s => f _ (@Compilers.ident.Z_add_with_get_carry_concrete s) | Compilers.ident.Z_sub_get_borrow => f _ Compilers.ident.Z_sub_get_borrow - | Compilers.ident.Z_sub_get_borrow_concrete s => f _ (@Compilers.ident.Z_sub_get_borrow_concrete s) | Compilers.ident.Z_sub_with_get_borrow => f _ Compilers.ident.Z_sub_with_get_borrow - | Compilers.ident.Z_sub_with_get_borrow_concrete s => f _ (@Compilers.ident.Z_sub_with_get_borrow_concrete s) | Compilers.ident.Z_zselect => f _ Compilers.ident.Z_zselect | Compilers.ident.Z_add_modulo => f _ Compilers.ident.Z_add_modulo | Compilers.ident.Z_rshi => f _ Compilers.ident.Z_rshi - | Compilers.ident.Z_rshi_concrete s offset => f _ (@Compilers.ident.Z_rshi_concrete s offset) | Compilers.ident.Z_cc_m => f _ Compilers.ident.Z_cc_m - | Compilers.ident.Z_cc_m_concrete s => f _ (@Compilers.ident.Z_cc_m_concrete s) - | Compilers.ident.Z_neg_snd => f _ Compilers.ident.Z_neg_snd | Compilers.ident.Z_cast range => f _ (@Compilers.ident.Z_cast range) | Compilers.ident.Z_cast2 range => f _ (@Compilers.ident.Z_cast2 range) | Compilers.ident.fancy_add log2wordmax imm => f _ (@Compilers.ident.fancy_add log2wordmax imm) @@ -995,7 +1161,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.fancy_sell => f _ Compilers.ident.fancy_sell | Compilers.ident.fancy_addm => f _ Compilers.ident.fancy_addm end. - + Definition of_typed_ident {t} (idc : Compilers.ident.ident t) : ident := match idc with | Compilers.ident.LiteralUnit v => LiteralUnit @@ -1037,30 +1203,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.Z_opp => Z_opp | Compilers.ident.Z_div => Z_div | Compilers.ident.Z_modulo => Z_modulo + | Compilers.ident.Z_log2 => Z_log2 + | Compilers.ident.Z_log2_up => Z_log2_up | Compilers.ident.Z_eqb => Z_eqb | Compilers.ident.Z_leb => Z_leb + | Compilers.ident.Z_geb => Z_geb | Compilers.ident.Z_of_nat => Z_of_nat - | Compilers.ident.Z_shiftr offset => Z_shiftr - | Compilers.ident.Z_shiftl offset => Z_shiftl - | Compilers.ident.Z_land mask => Z_land + | Compilers.ident.Z_to_nat => Z_to_nat + | Compilers.ident.Z_shiftr => Z_shiftr + | Compilers.ident.Z_shiftl => Z_shiftl + | Compilers.ident.Z_land => Z_land + | Compilers.ident.Z_lor => Z_lor + | Compilers.ident.Z_bneg => Z_bneg + | Compilers.ident.Z_lnot_modulo => Z_lnot_modulo | Compilers.ident.Z_mul_split => Z_mul_split - | Compilers.ident.Z_mul_split_concrete s => Z_mul_split_concrete | Compilers.ident.Z_add_get_carry => Z_add_get_carry - | Compilers.ident.Z_add_get_carry_concrete s => Z_add_get_carry_concrete | Compilers.ident.Z_add_with_carry => Z_add_with_carry | Compilers.ident.Z_add_with_get_carry => Z_add_with_get_carry - | Compilers.ident.Z_add_with_get_carry_concrete s => Z_add_with_get_carry_concrete | Compilers.ident.Z_sub_get_borrow => Z_sub_get_borrow - | Compilers.ident.Z_sub_get_borrow_concrete s => Z_sub_get_borrow_concrete | Compilers.ident.Z_sub_with_get_borrow => Z_sub_with_get_borrow - | Compilers.ident.Z_sub_with_get_borrow_concrete s => Z_sub_with_get_borrow_concrete | Compilers.ident.Z_zselect => Z_zselect | Compilers.ident.Z_add_modulo => Z_add_modulo | Compilers.ident.Z_rshi => Z_rshi - | Compilers.ident.Z_rshi_concrete s offset => Z_rshi_concrete | Compilers.ident.Z_cc_m => Z_cc_m - | Compilers.ident.Z_cc_m_concrete s => Z_cc_m_concrete - | Compilers.ident.Z_neg_snd => Z_neg_snd | Compilers.ident.Z_cast range => Z_cast | Compilers.ident.Z_cast2 range => Z_cast2 | Compilers.ident.fancy_add log2wordmax imm => fancy_add @@ -1077,7 +1242,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.fancy_sell => fancy_sell | Compilers.ident.fancy_addm => fancy_addm end. - + Definition arg_types (idc : ident) : option Type := match idc return option Type with | LiteralUnit => @Some Type (base.interp Compilers.base.type.unit) @@ -1119,30 +1284,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp => None | Z_div => None | Z_modulo => None + | Z_log2 => None + | Z_log2_up => None | Z_eqb => None | Z_leb => None + | Z_geb => None | Z_of_nat => None - | Z_shiftr => @Some Type Z - | Z_shiftl => @Some Type Z - | Z_land => @Some Type Z + | Z_to_nat => None + | Z_shiftr => None + | Z_shiftl => None + | Z_land => None + | Z_lor => None + | Z_bneg => None + | Z_lnot_modulo => None | Z_mul_split => None - | Z_mul_split_concrete => @Some Type Z | Z_add_get_carry => None - | Z_add_get_carry_concrete => @Some Type Z | Z_add_with_carry => None | Z_add_with_get_carry => None - | Z_add_with_get_carry_concrete => @Some Type Z | Z_sub_get_borrow => None - | Z_sub_get_borrow_concrete => @Some Type Z | Z_sub_with_get_borrow => None - | Z_sub_with_get_borrow_concrete => @Some Type Z | Z_zselect => None | Z_add_modulo => None | Z_rshi => None - | Z_rshi_concrete => @Some Type (Z * Z) | Z_cc_m => None - | Z_cc_m_concrete => @Some Type Z - | Z_neg_snd => None | Z_cast => @Some Type zrange | Z_cast2 => @Some Type (zrange * zrange) | fancy_add => @Some Type (Z * Z) @@ -1159,7 +1323,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | fancy_sell => None | fancy_addm => None end%type. - + Definition full_types (idc : ident) : Type := match idc return Type with | LiteralUnit => base.interp Compilers.base.type.unit @@ -1172,28 +1336,28 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Nat_mul => unit | Nat_add => unit | Nat_sub => unit - | nil => Compilers.base.type.type - | cons => Compilers.base.type.type - | pair => Compilers.base.type.type * Compilers.base.type.type - | fst => Compilers.base.type.type * Compilers.base.type.type - | snd => Compilers.base.type.type * Compilers.base.type.type - | pair_rect => Compilers.base.type.type * Compilers.base.type.type * Compilers.base.type.type - | bool_rect => Compilers.base.type.type - | nat_rect => Compilers.base.type.type - | list_rect => Compilers.base.type.type * Compilers.base.type.type - | list_case => Compilers.base.type.type * Compilers.base.type.type - | List_length => Compilers.base.type.type + | nil => base.type + | cons => base.type + | pair => base.type * base.type + | fst => base.type * base.type + | snd => base.type * base.type + | pair_rect => base.type * base.type * base.type + | bool_rect => base.type + | nat_rect => base.type + | list_rect => base.type * base.type + | list_case => base.type * base.type + | List_length => base.type | List_seq => unit - | List_repeat => Compilers.base.type.type - | List_combine => Compilers.base.type.type * Compilers.base.type.type - | List_map => Compilers.base.type.type * Compilers.base.type.type - | List_app => Compilers.base.type.type - | List_rev => Compilers.base.type.type - | List_flat_map => Compilers.base.type.type * Compilers.base.type.type - | List_partition => Compilers.base.type.type - | List_fold_right => Compilers.base.type.type * Compilers.base.type.type - | List_update_nth => Compilers.base.type.type - | List_nth_default => Compilers.base.type.type + | List_repeat => base.type + | List_combine => base.type * base.type + | List_map => base.type * base.type + | List_app => base.type + | List_rev => base.type + | List_flat_map => base.type * base.type + | List_partition => base.type + | List_fold_right => base.type * base.type + | List_update_nth => base.type + | List_nth_default => base.type | Z_add => unit | Z_mul => unit | Z_pow => unit @@ -1201,30 +1365,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp => unit | Z_div => unit | Z_modulo => unit + | Z_log2 => unit + | Z_log2_up => unit | Z_eqb => unit | Z_leb => unit + | Z_geb => unit | Z_of_nat => unit - | Z_shiftr => Z - | Z_shiftl => Z - | Z_land => Z + | Z_to_nat => unit + | Z_shiftr => unit + | Z_shiftl => unit + | Z_land => unit + | Z_lor => unit + | Z_bneg => unit + | Z_lnot_modulo => unit | Z_mul_split => unit - | Z_mul_split_concrete => Z | Z_add_get_carry => unit - | Z_add_get_carry_concrete => Z | Z_add_with_carry => unit | Z_add_with_get_carry => unit - | Z_add_with_get_carry_concrete => Z | Z_sub_get_borrow => unit - | Z_sub_get_borrow_concrete => Z | Z_sub_with_get_borrow => unit - | Z_sub_with_get_borrow_concrete => Z | Z_zselect => unit | Z_add_modulo => unit | Z_rshi => unit - | Z_rshi_concrete => Z * Z | Z_cc_m => unit - | Z_cc_m_concrete => Z - | Z_neg_snd => unit | Z_cast => zrange | Z_cast2 => zrange * zrange | fancy_add => Z * Z @@ -1241,7 +1404,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | fancy_sell => unit | fancy_addm => unit end%type. - + Definition bind_args {t} (idc : Compilers.ident.ident t) : match arg_types (of_typed_ident idc) return Type with Some t => t | None => unit end := match idc return match arg_types (of_typed_ident idc) return Type with Some t => t | None => unit end with | Compilers.ident.LiteralUnit v => v @@ -1283,30 +1446,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.Z_opp => tt | Compilers.ident.Z_div => tt | Compilers.ident.Z_modulo => tt + | Compilers.ident.Z_log2 => tt + | Compilers.ident.Z_log2_up => tt | Compilers.ident.Z_eqb => tt | Compilers.ident.Z_leb => tt + | Compilers.ident.Z_geb => tt | Compilers.ident.Z_of_nat => tt - | Compilers.ident.Z_shiftr offset => offset - | Compilers.ident.Z_shiftl offset => offset - | Compilers.ident.Z_land mask => mask + | Compilers.ident.Z_to_nat => tt + | Compilers.ident.Z_shiftr => tt + | Compilers.ident.Z_shiftl => tt + | Compilers.ident.Z_land => tt + | Compilers.ident.Z_lor => tt + | Compilers.ident.Z_bneg => tt + | Compilers.ident.Z_lnot_modulo => tt | Compilers.ident.Z_mul_split => tt - | Compilers.ident.Z_mul_split_concrete s => s | Compilers.ident.Z_add_get_carry => tt - | Compilers.ident.Z_add_get_carry_concrete s => s | Compilers.ident.Z_add_with_carry => tt | Compilers.ident.Z_add_with_get_carry => tt - | Compilers.ident.Z_add_with_get_carry_concrete s => s | Compilers.ident.Z_sub_get_borrow => tt - | Compilers.ident.Z_sub_get_borrow_concrete s => s | Compilers.ident.Z_sub_with_get_borrow => tt - | Compilers.ident.Z_sub_with_get_borrow_concrete s => s | Compilers.ident.Z_zselect => tt | Compilers.ident.Z_add_modulo => tt | Compilers.ident.Z_rshi => tt - | Compilers.ident.Z_rshi_concrete s offset => (s, offset) | Compilers.ident.Z_cc_m => tt - | Compilers.ident.Z_cc_m_concrete s => s - | Compilers.ident.Z_neg_snd => tt | Compilers.ident.Z_cast range => range | Compilers.ident.Z_cast2 range => range | Compilers.ident.fancy_add log2wordmax imm => (log2wordmax, imm) @@ -1323,7 +1485,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.fancy_sell => tt | Compilers.ident.fancy_addm => tt end%cps. - + Definition invert_bind_args {t} (idc : Compilers.ident.ident t) (pidc : ident) : option (full_types pidc) := match pidc, idc return option (full_types pidc) with | LiteralUnit, Compilers.ident.LiteralUnit v => Some v @@ -1365,30 +1527,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp, Compilers.ident.Z_opp => Some tt | Z_div, Compilers.ident.Z_div => Some tt | Z_modulo, Compilers.ident.Z_modulo => Some tt + | Z_log2, Compilers.ident.Z_log2 => Some tt + | Z_log2_up, Compilers.ident.Z_log2_up => Some tt | Z_eqb, Compilers.ident.Z_eqb => Some tt | Z_leb, Compilers.ident.Z_leb => Some tt + | Z_geb, Compilers.ident.Z_geb => Some tt | Z_of_nat, Compilers.ident.Z_of_nat => Some tt - | Z_shiftr, Compilers.ident.Z_shiftr offset => Some offset - | Z_shiftl, Compilers.ident.Z_shiftl offset => Some offset - | Z_land, Compilers.ident.Z_land mask => Some mask + | Z_to_nat, Compilers.ident.Z_to_nat => Some tt + | Z_shiftr, Compilers.ident.Z_shiftr => Some tt + | Z_shiftl, Compilers.ident.Z_shiftl => Some tt + | Z_land, Compilers.ident.Z_land => Some tt + | Z_lor, Compilers.ident.Z_lor => Some tt + | Z_bneg, Compilers.ident.Z_bneg => Some tt + | Z_lnot_modulo, Compilers.ident.Z_lnot_modulo => Some tt | Z_mul_split, Compilers.ident.Z_mul_split => Some tt - | Z_mul_split_concrete, Compilers.ident.Z_mul_split_concrete s => Some s | Z_add_get_carry, Compilers.ident.Z_add_get_carry => Some tt - | Z_add_get_carry_concrete, Compilers.ident.Z_add_get_carry_concrete s => Some s | Z_add_with_carry, Compilers.ident.Z_add_with_carry => Some tt | Z_add_with_get_carry, Compilers.ident.Z_add_with_get_carry => Some tt - | Z_add_with_get_carry_concrete, Compilers.ident.Z_add_with_get_carry_concrete s => Some s | Z_sub_get_borrow, Compilers.ident.Z_sub_get_borrow => Some tt - | Z_sub_get_borrow_concrete, Compilers.ident.Z_sub_get_borrow_concrete s => Some s | Z_sub_with_get_borrow, Compilers.ident.Z_sub_with_get_borrow => Some tt - | Z_sub_with_get_borrow_concrete, Compilers.ident.Z_sub_with_get_borrow_concrete s => Some s | Z_zselect, Compilers.ident.Z_zselect => Some tt | Z_add_modulo, Compilers.ident.Z_add_modulo => Some tt | Z_rshi, Compilers.ident.Z_rshi => Some tt - | Z_rshi_concrete, Compilers.ident.Z_rshi_concrete s offset => Some (s, offset) | Z_cc_m, Compilers.ident.Z_cc_m => Some tt - | Z_cc_m_concrete, Compilers.ident.Z_cc_m_concrete s => Some s - | Z_neg_snd, Compilers.ident.Z_neg_snd => Some tt | Z_cast, Compilers.ident.Z_cast range => Some range | Z_cast2, Compilers.ident.Z_cast2 range => Some range | fancy_add, Compilers.ident.fancy_add log2wordmax imm => Some (log2wordmax, imm) @@ -1443,30 +1604,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp, _ | Z_div, _ | Z_modulo, _ + | Z_log2, _ + | Z_log2_up, _ | Z_eqb, _ | Z_leb, _ + | Z_geb, _ | Z_of_nat, _ + | Z_to_nat, _ | Z_shiftr, _ | Z_shiftl, _ | Z_land, _ + | Z_lor, _ + | Z_bneg, _ + | Z_lnot_modulo, _ | Z_mul_split, _ - | Z_mul_split_concrete, _ | Z_add_get_carry, _ - | Z_add_get_carry_concrete, _ | Z_add_with_carry, _ | Z_add_with_get_carry, _ - | Z_add_with_get_carry_concrete, _ | Z_sub_get_borrow, _ - | Z_sub_get_borrow_concrete, _ | Z_sub_with_get_borrow, _ - | Z_sub_with_get_borrow_concrete, _ | Z_zselect, _ | Z_add_modulo, _ | Z_rshi, _ - | Z_rshi_concrete, _ | Z_cc_m, _ - | Z_cc_m_concrete, _ - | Z_neg_snd, _ | Z_cast, _ | Z_cast2, _ | fancy_add, _ @@ -1484,7 +1644,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | fancy_addm, _ => None end%cps. - + Local Notation eta2 x := (Datatypes.fst x, Datatypes.snd x) (only parsing). Local Notation eta3 x := (eta2 (Datatypes.fst x), Datatypes.snd x) (only parsing). @@ -1494,82 +1654,81 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | LiteralZ => fun v => (type.base (Compilers.base.type.type_base Compilers.base.type.Z)) | LiteralBool => fun v => (type.base (Compilers.base.type.type_base Compilers.base.type.bool)) | LiteralNat => fun v => (type.base (Compilers.base.type.type_base Compilers.base.type.nat)) - | Nat_succ => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | Nat_pred => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | Nat_max => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | Nat_mul => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | Nat_add => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | Nat_sub => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | nil => fun t => (type.base (Compilers.base.type.list t)) - | cons => fun t => (type.base t -> type.base (Compilers.base.type.list t) -> type.base (Compilers.base.type.list t))%etype - | pair => fun arg => let '(A, B) := eta2 arg in (type.base A -> type.base B -> type.base (A * B)%etype)%etype - | fst => fun arg => let '(A, B) := eta2 arg in (type.base (A * B)%etype -> type.base A)%etype - | snd => fun arg => let '(A, B) := eta2 arg in (type.base (A * B)%etype -> type.base B)%etype - | pair_rect => fun arg => let '(A, B, T) := eta3 arg in ((type.base A -> type.base B -> type.base T) -> type.base (A * B)%etype -> type.base T)%etype - | bool_rect => fun T => ((type.base ()%etype -> type.base T) -> (type.base ()%etype -> type.base T) -> type.base (Compilers.base.type.type_base base.type.bool) -> type.base T)%etype - | nat_rect => fun P => ((type.base ()%etype -> type.base P) -> (type.base (Compilers.base.type.type_base base.type.nat) -> type.base P -> type.base P) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base P)%etype - | list_rect => fun arg => let '(A, P) := eta2 arg in ((type.base ()%etype -> type.base P) -> (type.base A -> type.base (Compilers.base.type.list A) -> type.base P -> type.base P) -> type.base (Compilers.base.type.list A) -> type.base P)%etype - | list_case => fun arg => let '(A, P) := eta2 arg in ((type.base ()%etype -> type.base P) -> (type.base A -> type.base (Compilers.base.type.list A) -> type.base P) -> type.base (Compilers.base.type.list A) -> type.base P)%etype - | List_length => fun T => (type.base (Compilers.base.type.list T) -> type.base (Compilers.base.type.type_base base.type.nat))%etype - | List_seq => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.list (Compilers.base.type.type_base base.type.nat)))%etype - | List_repeat => fun A => (type.base A -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.list A))%etype - | List_combine => fun arg => let '(A, B) := eta2 arg in (type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list B) -> type.base (Compilers.base.type.list (A * B)))%etype - | List_map => fun arg => let '(A, B) := eta2 arg in ((type.base A -> type.base B) -> type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list B))%etype - | List_app => fun A => (type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list A))%etype - | List_rev => fun A => (type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list A))%etype - | List_flat_map => fun arg => let '(A, B) := eta2 arg in ((type.base A -> type.base (Compilers.base.type.list B)) -> type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list B))%etype - | List_partition => fun A => ((type.base A -> type.base (Compilers.base.type.type_base base.type.bool)) -> type.base (Compilers.base.type.list A) -> type.base (Compilers.base.type.list A * Compilers.base.type.list A)%etype)%etype - | List_fold_right => fun arg => let '(A, B) := eta2 arg in ((type.base B -> type.base A -> type.base A) -> type.base A -> type.base (Compilers.base.type.list B) -> type.base A)%etype - | List_update_nth => fun T => (type.base (Compilers.base.type.type_base base.type.nat) -> (type.base T -> type.base T) -> type.base (Compilers.base.type.list T) -> type.base (Compilers.base.type.list T))%etype - | List_nth_default => fun T => (type.base T -> type.base (Compilers.base.type.list T) -> type.base (Compilers.base.type.type_base base.type.nat) -> type.base T)%etype - | Z_add => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_mul => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_pow => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_sub => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_opp => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_div => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_modulo => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_eqb => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.bool))%etype - | Z_leb => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.bool))%etype - | Z_of_nat => fun _ => (type.base (Compilers.base.type.type_base base.type.nat) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_shiftr => fun offset => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_shiftl => fun offset => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_land => fun mask => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_mul_split => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_mul_split_concrete => fun s => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_add_get_carry => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_add_get_carry_concrete => fun s => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_add_with_carry => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_add_with_get_carry => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_add_with_get_carry_concrete => fun s => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_sub_get_borrow => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_sub_get_borrow_concrete => fun s => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_sub_with_get_borrow => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_sub_with_get_borrow_concrete => fun s => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_zselect => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_add_modulo => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_rshi => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_rshi_concrete => fun arg => let '(s, offset) := eta2 arg in (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_cc_m => fun _ => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_cc_m_concrete => fun s => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_neg_snd => fun _ => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | Z_cast => fun range => (type.base (Compilers.base.type.type_base base.type.Z) -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | Z_cast2 => fun range => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | fancy_add => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | fancy_addc => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | fancy_sub => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | fancy_subb => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype)%etype - | fancy_mulll => fun log2wordmax => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_mullh => fun log2wordmax => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_mulhl => fun log2wordmax => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_mulhh => fun log2wordmax => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_rshi => fun arg => let '(log2wordmax, x) := eta2 arg in (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_selc => fun _ => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_selm => fun log2wordmax => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_sell => fun _ => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - | fancy_addm => fun _ => (type.base (Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z * Compilers.base.type.type_base base.type.Z)%etype -> type.base (Compilers.base.type.type_base base.type.Z))%etype - end. - + | Nat_succ => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat)) + | Nat_pred => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat)) + | Nat_max => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat)) + | Nat_mul => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat)) + | Nat_add => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat)) + | Nat_sub => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat)) + | nil => fun t => (type.base (base.type.list t)) + | cons => fun t => (type.base t -> type.base (base.type.list t) -> type.base (base.type.list t)) + | pair => fun arg => let '(A, B) := eta2 arg in (type.base A -> type.base B -> type.base (A * B)%etype) + | fst => fun arg => let '(A, B) := eta2 arg in (type.base (A * B)%etype -> type.base A) + | snd => fun arg => let '(A, B) := eta2 arg in (type.base (A * B)%etype -> type.base B) + | pair_rect => fun arg => let '(A, B, T) := eta3 arg in ((type.base A -> type.base B -> type.base T) -> type.base (A * B)%etype -> type.base T) + | bool_rect => fun T => ((type.base ()%etype -> type.base T) -> (type.base ()%etype -> type.base T) -> type.base (base.type.type_base base.type.bool) -> type.base T) + | nat_rect => fun P => ((type.base ()%etype -> type.base P) -> (type.base (base.type.type_base base.type.nat) -> type.base P -> type.base P) -> type.base (base.type.type_base base.type.nat) -> type.base P) + | list_rect => fun arg => let '(A, P) := eta2 arg in ((type.base ()%etype -> type.base P) -> (type.base A -> type.base (base.type.list A) -> type.base P -> type.base P) -> type.base (base.type.list A) -> type.base P) + | list_case => fun arg => let '(A, P) := eta2 arg in ((type.base ()%etype -> type.base P) -> (type.base A -> type.base (base.type.list A) -> type.base P) -> type.base (base.type.list A) -> type.base P) + | List_length => fun T => (type.base (base.type.list T) -> type.base (base.type.type_base base.type.nat)) + | List_seq => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.nat) -> type.base (base.type.list (base.type.type_base base.type.nat))) + | List_repeat => fun A => (type.base A -> type.base (base.type.type_base base.type.nat) -> type.base (base.type.list A)) + | List_combine => fun arg => let '(A, B) := eta2 arg in (type.base (base.type.list A) -> type.base (base.type.list B) -> type.base (base.type.list (A * B))) + | List_map => fun arg => let '(A, B) := eta2 arg in ((type.base A -> type.base B) -> type.base (base.type.list A) -> type.base (base.type.list B)) + | List_app => fun A => (type.base (base.type.list A) -> type.base (base.type.list A) -> type.base (base.type.list A)) + | List_rev => fun A => (type.base (base.type.list A) -> type.base (base.type.list A)) + | List_flat_map => fun arg => let '(A, B) := eta2 arg in ((type.base A -> type.base (base.type.list B)) -> type.base (base.type.list A) -> type.base (base.type.list B)) + | List_partition => fun A => ((type.base A -> type.base (base.type.type_base base.type.bool)) -> type.base (base.type.list A) -> type.base (base.type.list A * base.type.list A)%etype) + | List_fold_right => fun arg => let '(A, B) := eta2 arg in ((type.base B -> type.base A -> type.base A) -> type.base A -> type.base (base.type.list B) -> type.base A) + | List_update_nth => fun T => (type.base (base.type.type_base base.type.nat) -> (type.base T -> type.base T) -> type.base (base.type.list T) -> type.base (base.type.list T)) + | List_nth_default => fun T => (type.base T -> type.base (base.type.list T) -> type.base (base.type.type_base base.type.nat) -> type.base T) + | Z_add => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_mul => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_pow => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_sub => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_opp => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_div => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_modulo => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_log2 => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_log2_up => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_eqb => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.bool)) + | Z_leb => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.bool)) + | Z_geb => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.bool)) + | Z_of_nat => fun _ => (type.base (base.type.type_base base.type.nat) -> type.base (base.type.type_base base.type.Z)) + | Z_to_nat => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.nat)) + | Z_shiftr => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_shiftl => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_land => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_lor => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_bneg => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_lnot_modulo => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_mul_split => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | Z_add_get_carry => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | Z_add_with_carry => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_add_with_get_carry => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | Z_sub_get_borrow => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | Z_sub_with_get_borrow => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | Z_zselect => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_add_modulo => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_rshi => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_cc_m => fun _ => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_cast => fun range => (type.base (base.type.type_base base.type.Z) -> type.base (base.type.type_base base.type.Z)) + | Z_cast2 => fun range => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | fancy_add => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | fancy_addc => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | fancy_sub => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | fancy_subb => fun arg => let '(log2wordmax, imm) := eta2 arg in (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype) + | fancy_mulll => fun log2wordmax => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_mullh => fun log2wordmax => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_mulhl => fun log2wordmax => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_mulhh => fun log2wordmax => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_rshi => fun arg => let '(log2wordmax, x) := eta2 arg in (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_selc => fun _ => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_selm => fun log2wordmax => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_sell => fun _ => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + | fancy_addm => fun _ => (type.base (base.type.type_base base.type.Z * base.type.type_base base.type.Z * base.type.type_base base.type.Z)%etype -> type.base (base.type.type_base base.type.Z)) + end%etype. + Definition to_typed (pidc : ident) : forall args : full_types pidc, Compilers.ident.ident (type_of pidc args) := match pidc return forall args : full_types pidc, Compilers.ident.ident (type_of pidc args) with | LiteralUnit => fun v => @Compilers.ident.LiteralUnit v @@ -1611,30 +1770,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Z_opp => fun _ => @Compilers.ident.Z_opp | Z_div => fun _ => @Compilers.ident.Z_div | Z_modulo => fun _ => @Compilers.ident.Z_modulo + | Z_log2 => fun _ => @Compilers.ident.Z_log2 + | Z_log2_up => fun _ => @Compilers.ident.Z_log2_up | Z_eqb => fun _ => @Compilers.ident.Z_eqb | Z_leb => fun _ => @Compilers.ident.Z_leb + | Z_geb => fun _ => @Compilers.ident.Z_geb | Z_of_nat => fun _ => @Compilers.ident.Z_of_nat - | Z_shiftr => fun offset => @Compilers.ident.Z_shiftr offset - | Z_shiftl => fun offset => @Compilers.ident.Z_shiftl offset - | Z_land => fun mask => @Compilers.ident.Z_land mask + | Z_to_nat => fun _ => @Compilers.ident.Z_to_nat + | Z_shiftr => fun _ => @Compilers.ident.Z_shiftr + | Z_shiftl => fun _ => @Compilers.ident.Z_shiftl + | Z_land => fun _ => @Compilers.ident.Z_land + | Z_lor => fun _ => @Compilers.ident.Z_lor + | Z_bneg => fun _ => @Compilers.ident.Z_bneg + | Z_lnot_modulo => fun _ => @Compilers.ident.Z_lnot_modulo | Z_mul_split => fun _ => @Compilers.ident.Z_mul_split - | Z_mul_split_concrete => fun s => @Compilers.ident.Z_mul_split_concrete s | Z_add_get_carry => fun _ => @Compilers.ident.Z_add_get_carry - | Z_add_get_carry_concrete => fun s => @Compilers.ident.Z_add_get_carry_concrete s | Z_add_with_carry => fun _ => @Compilers.ident.Z_add_with_carry | Z_add_with_get_carry => fun _ => @Compilers.ident.Z_add_with_get_carry - | Z_add_with_get_carry_concrete => fun s => @Compilers.ident.Z_add_with_get_carry_concrete s | Z_sub_get_borrow => fun _ => @Compilers.ident.Z_sub_get_borrow - | Z_sub_get_borrow_concrete => fun s => @Compilers.ident.Z_sub_get_borrow_concrete s | Z_sub_with_get_borrow => fun _ => @Compilers.ident.Z_sub_with_get_borrow - | Z_sub_with_get_borrow_concrete => fun s => @Compilers.ident.Z_sub_with_get_borrow_concrete s | Z_zselect => fun _ => @Compilers.ident.Z_zselect | Z_add_modulo => fun _ => @Compilers.ident.Z_add_modulo | Z_rshi => fun _ => @Compilers.ident.Z_rshi - | Z_rshi_concrete => fun arg => match eta2 arg as args' return Compilers.ident.ident (type_of Z_rshi_concrete args') with (s, offset) => @Compilers.ident.Z_rshi_concrete s offset end | Z_cc_m => fun _ => @Compilers.ident.Z_cc_m - | Z_cc_m_concrete => fun s => @Compilers.ident.Z_cc_m_concrete s - | Z_neg_snd => fun _ => @Compilers.ident.Z_neg_snd | Z_cast => fun range => @Compilers.ident.Z_cast range | Z_cast2 => fun range => @Compilers.ident.Z_cast2 range | fancy_add => fun arg => match eta2 arg as args' return Compilers.ident.ident (type_of fancy_add args') with (log2wordmax, imm) => @Compilers.ident.fancy_add log2wordmax imm end @@ -1651,7 +1809,7 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | fancy_sell => fun _ => @Compilers.ident.fancy_sell | fancy_addm => fun _ => @Compilers.ident.fancy_addm end. - + Definition retype_ident {t} (idc : Compilers.ident.ident t) : match arg_types (of_typed_ident idc) return Type with Some t => t | None => unit end -> Compilers.ident.ident t := match idc in Compilers.ident.ident t return match arg_types (of_typed_ident idc) return Type with Some t => t | None => unit end -> Compilers.ident.ident t with | Compilers.ident.LiteralUnit v => (fun v => @Compilers.ident.LiteralUnit v) : match arg_types (of_typed_ident (@Compilers.ident.LiteralUnit v)) return Type with Some t => t | None => unit end -> _ (* COQBUG(https://github.com/coq/coq/issues/7726) *) @@ -1693,30 +1851,29 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.Z_opp => fun _ => @Compilers.ident.Z_opp | Compilers.ident.Z_div => fun _ => @Compilers.ident.Z_div | Compilers.ident.Z_modulo => fun _ => @Compilers.ident.Z_modulo + | Compilers.ident.Z_log2 => fun _ => @Compilers.ident.Z_log2 + | Compilers.ident.Z_log2_up => fun _ => @Compilers.ident.Z_log2_up | Compilers.ident.Z_eqb => fun _ => @Compilers.ident.Z_eqb | Compilers.ident.Z_leb => fun _ => @Compilers.ident.Z_leb + | Compilers.ident.Z_geb => fun _ => @Compilers.ident.Z_geb | Compilers.ident.Z_of_nat => fun _ => @Compilers.ident.Z_of_nat - | Compilers.ident.Z_shiftr offset => fun offset => @Compilers.ident.Z_shiftr offset - | Compilers.ident.Z_shiftl offset => fun offset => @Compilers.ident.Z_shiftl offset - | Compilers.ident.Z_land mask => fun mask => @Compilers.ident.Z_land mask + | Compilers.ident.Z_to_nat => fun _ => @Compilers.ident.Z_to_nat + | Compilers.ident.Z_shiftr => fun _ => @Compilers.ident.Z_shiftr + | Compilers.ident.Z_shiftl => fun _ => @Compilers.ident.Z_shiftl + | Compilers.ident.Z_land => fun _ => @Compilers.ident.Z_land + | Compilers.ident.Z_lor => fun _ => @Compilers.ident.Z_lor + | Compilers.ident.Z_bneg => fun _ => @Compilers.ident.Z_bneg + | Compilers.ident.Z_lnot_modulo => fun _ => @Compilers.ident.Z_lnot_modulo | Compilers.ident.Z_mul_split => fun _ => @Compilers.ident.Z_mul_split - | Compilers.ident.Z_mul_split_concrete s => fun s => @Compilers.ident.Z_mul_split_concrete s | Compilers.ident.Z_add_get_carry => fun _ => @Compilers.ident.Z_add_get_carry - | Compilers.ident.Z_add_get_carry_concrete s => fun s => @Compilers.ident.Z_add_get_carry_concrete s | Compilers.ident.Z_add_with_carry => fun _ => @Compilers.ident.Z_add_with_carry | Compilers.ident.Z_add_with_get_carry => fun _ => @Compilers.ident.Z_add_with_get_carry - | Compilers.ident.Z_add_with_get_carry_concrete s => fun s => @Compilers.ident.Z_add_with_get_carry_concrete s | Compilers.ident.Z_sub_get_borrow => fun _ => @Compilers.ident.Z_sub_get_borrow - | Compilers.ident.Z_sub_get_borrow_concrete s => fun s => @Compilers.ident.Z_sub_get_borrow_concrete s | Compilers.ident.Z_sub_with_get_borrow => fun _ => @Compilers.ident.Z_sub_with_get_borrow - | Compilers.ident.Z_sub_with_get_borrow_concrete s => fun s => @Compilers.ident.Z_sub_with_get_borrow_concrete s | Compilers.ident.Z_zselect => fun _ => @Compilers.ident.Z_zselect | Compilers.ident.Z_add_modulo => fun _ => @Compilers.ident.Z_add_modulo | Compilers.ident.Z_rshi => fun _ => @Compilers.ident.Z_rshi - | Compilers.ident.Z_rshi_concrete s offset => fun arg => let '(s, offset) := eta2 arg in @Compilers.ident.Z_rshi_concrete s offset | Compilers.ident.Z_cc_m => fun _ => @Compilers.ident.Z_cc_m - | Compilers.ident.Z_cc_m_concrete s => fun s => @Compilers.ident.Z_cc_m_concrete s - | Compilers.ident.Z_neg_snd => fun _ => @Compilers.ident.Z_neg_snd | Compilers.ident.Z_cast range => fun range => @Compilers.ident.Z_cast range | Compilers.ident.Z_cast2 range => fun range => @Compilers.ident.Z_cast2 range | Compilers.ident.fancy_add log2wordmax imm => fun arg => let '(log2wordmax, imm) := eta2 arg in @Compilers.ident.fancy_add log2wordmax imm @@ -1733,8 +1890,8 @@ with open('GENERATEDIdentifiersWithoutTypes.v', 'w') as f: | Compilers.ident.fancy_sell => fun _ => @Compilers.ident.fancy_sell | Compilers.ident.fancy_addm => fun _ => @Compilers.ident.fancy_addm end. - - + + (*===*) End ident. End pattern. diff --git a/src/Experiments/NewPipeline/Language.v b/src/Experiments/NewPipeline/Language.v index f839d8949..af9780418 100644 --- a/src/Experiments/NewPipeline/Language.v +++ b/src/Experiments/NewPipeline/Language.v @@ -629,7 +629,14 @@ Module Compilers. end in constr:(I : I) end in - reify_rec term) + match constr:(Set) with + | _ => reify_rec term + | _ => let __ := match goal with + | _ => idtac "Error: Failed to reify" term' "via unfolding"; + fail 2 "Failed to reify" term' "via unfolding" + end in + constr:(I : I) + end) end) end end. @@ -739,7 +746,6 @@ Module Compilers. Section with_scope. Import base.type. Notation type := ttype. - (* N.B. [ident] must have essentially flat structure for the python script constructing [pattern.ident] to work *) Inductive ident : type -> Type := @@ -779,30 +785,29 @@ Module Compilers. | Z_opp : ident (Z -> Z) | Z_div : ident (Z -> Z -> Z) | Z_modulo : ident (Z -> Z -> Z) + | Z_log2 : ident (Z -> Z) + | Z_log2_up : ident (Z -> Z) | Z_eqb : ident (Z -> Z -> bool) | Z_leb : ident (Z -> Z -> bool) + | Z_geb : ident (Z -> Z -> bool) | Z_of_nat : ident (nat -> Z) - | Z_shiftr (offset : BinInt.Z) : ident (Z -> Z) - | Z_shiftl (offset : BinInt.Z) : ident (Z -> Z) - | Z_land (mask : BinInt.Z) : ident (Z -> Z) + | Z_to_nat : ident (Z -> nat) + | Z_shiftr : ident (Z -> Z -> Z) + | Z_shiftl : ident (Z -> Z -> Z) + | Z_land : ident (Z -> Z -> Z) + | Z_lor : ident (Z -> Z -> Z) + | Z_bneg : ident (Z -> Z) + | Z_lnot_modulo : ident (Z -> Z -> Z) | Z_mul_split : ident (Z -> Z -> Z -> Z * Z) - | Z_mul_split_concrete (s:BinInt.Z) : ident (Z -> Z -> Z * Z) | Z_add_get_carry : ident (Z -> Z -> Z -> (Z * Z)) - | Z_add_get_carry_concrete (s:BinInt.Z) : ident (Z -> Z -> (Z * Z)) | Z_add_with_carry : ident (Z -> Z -> Z -> Z) | Z_add_with_get_carry : ident (Z -> Z -> Z -> Z -> (Z * Z)) - | Z_add_with_get_carry_concrete (s:BinInt.Z) : ident (Z -> Z -> Z -> Z * Z) | Z_sub_get_borrow : ident (Z -> Z -> Z -> (Z * Z)) - | Z_sub_get_borrow_concrete (s:BinInt.Z) : ident (Z -> Z -> Z * Z) | Z_sub_with_get_borrow : ident (Z -> Z -> Z -> Z -> (Z * Z)) - | Z_sub_with_get_borrow_concrete (s:BinInt.Z) : ident (Z -> Z -> Z -> Z * Z) | Z_zselect : ident (Z -> Z -> Z -> Z) | Z_add_modulo : ident (Z -> Z -> Z -> Z) | Z_rshi : ident (Z -> Z -> Z -> Z -> Z) - | Z_rshi_concrete (s offset:BinInt.Z) : ident (Z -> Z -> Z) | Z_cc_m : ident (Z -> Z -> Z) - | Z_cc_m_concrete (s:BinInt.Z) : ident (Z -> Z) - | Z_neg_snd : ident ((Z * Z) -> Z * Z) (** TODO(jadep): This is only here for demonstration purposes; remove it once you no longer need it as a template; N.B. the type signature here says "given any amount of information about a thing of type [ℤ * ℤ], we promise to return a concrete pair of some amount of information about a thing of type ℤ and a thing of type ℤ" *) | Z_cast (range : zrange) : ident (Z -> Z) | Z_cast2 (range : zrange * zrange) : ident ((Z * Z) -> (Z * Z)) | fancy_add (log2wordmax : BinInt.Z) (imm : BinInt.Z) : ident (Z * Z -> Z * Z) @@ -907,28 +912,27 @@ Module Compilers. | Z_modulo => Z.modulo | Z_eqb => Z.eqb | Z_leb => Z.leb + | Z_geb => Z.geb + | Z_log2 => Z.log2 + | Z_log2_up => Z.log2_up | Z_of_nat => Z.of_nat - | Z_shiftr offset => fun v => Z.shiftr v offset - | Z_shiftl offset => fun v => Z.shiftl v offset - | Z_land mask => fun v => Z.land v mask + | Z_to_nat => Z.to_nat + | Z_shiftr => Z.shiftr + | Z_shiftl => Z.shiftl + | Z_land => Z.land + | Z_lor => Z.lor | Z_mul_split => Z.mul_split - | Z_mul_split_concrete s => Z.mul_split s | Z_add_get_carry => Z.add_get_carry_full - | Z_add_get_carry_concrete s => Z.add_get_carry_full s | Z_add_with_carry => Z.add_with_carry | Z_add_with_get_carry => Z.add_with_get_carry_full - | Z_add_with_get_carry_concrete s => Z.add_with_get_carry_full s | Z_sub_get_borrow => Z.sub_get_borrow_full - | Z_sub_get_borrow_concrete s => Z.sub_get_borrow_full s | Z_sub_with_get_borrow => Z.sub_with_get_borrow_full - | Z_sub_with_get_borrow_concrete s => Z.sub_with_get_borrow_full s | Z_zselect => Z.zselect | Z_add_modulo => Z.add_modulo + | Z_bneg => Z.bneg + | Z_lnot_modulo => Z.lnot_modulo | Z_rshi => Z.rshi - | Z_rshi_concrete s offset => fun x y => Z.rshi s x y offset | Z_cc_m => Z.cc_m - | Z_cc_m_concrete s => Z.cc_m s - | Z_neg_snd => fun '(x, y) => (x, -y) (** TODO(jadep): This is only here for demonstration purposes; remove it once you no longer need it as a template *) | Z_cast r => cast r | Z_cast2 (r1, r2) => fun '(x1, x2) => (cast r1 x1, cast r2 x2) | fancy_add _ _ as idc @@ -1110,7 +1114,17 @@ Module Compilers. | Z.modulo => then_tac ident.Z_modulo | Z.eqb => then_tac ident.Z_eqb | Z.leb => then_tac ident.Z_leb + | Z.geb => then_tac ident.Z_geb + | Z.log2 => then_tac ident.Z_log2 + | Z.log2_up => then_tac ident.Z_log2_up + | Z.shiftl => then_tac ident.Z_shiftl + | Z.shiftr => then_tac ident.Z_shiftr + | Z.land => then_tac ident.Z_land + | Z.lor => then_tac ident.Z_lor + | Z.bneg => then_tac ident.Z_bneg + | Z.lnot_modulo => then_tac ident.Z_lnot_modulo | Z.of_nat => then_tac ident.Z_of_nat + | Z.to_nat => then_tac ident.Z_to_nat | Z.mul_split => then_tac ident.Z_mul_split | Z.add_get_carry_full => then_tac ident.Z_add_get_carry | Z.add_with_carry => then_tac ident.Z_add_with_carry @@ -1166,8 +1180,10 @@ Module Compilers. Notation "x + y" := (#Z_add @ x @ y)%expr : expr_scope. Notation "x / y" := (#Z_div @ x @ y)%expr : expr_scope. Notation "x * y" := (#Z_mul @ x @ y)%expr : expr_scope. - Notation "x >> y" := (#(Z_shiftr y) @ x)%expr : expr_scope. - Notation "x << y" := (#(Z_shiftl y) @ x)%expr : expr_scope. + Notation "x >> y" := (#Z_shiftr @ x @ y)%expr : expr_scope. + Notation "x << y" := (#Z_shiftl @ x @ y)%expr : expr_scope. + Notation "x &' y" := (#Z_land @ x @ y)%expr : expr_scope. + Notation "x || y" := (#Z_lor @ x @ y)%expr : expr_scope. Notation "x 'mod' y" := (#Z_modulo @ x @ y)%expr : expr_scope. Notation "- x" := (#Z_opp @ x)%expr : expr_scope. End Notations. @@ -1262,6 +1278,22 @@ Module Compilers. | e => fun v => @App_curried _ e (type.map_for_each_lhs_of_arrow (fun _ v => expr.Var v) v) end. + Fixpoint invert_App_curried {t} (e : expr t) + : type.for_each_lhs_of_arrow expr t -> { t' : _ & expr t' * type.for_each_lhs_of_arrow expr t' }%type + := match e in expr.expr t return type.for_each_lhs_of_arrow expr t -> { t' : _ & expr t' * type.for_each_lhs_of_arrow expr t' }%type with + | expr.App s d f x + => fun args + => @invert_App_curried _ f (x, args) + | e => fun args => existT _ _ (e, args) + end. + Definition invert_AppIdent_curried {t} (e : expr t) + : option { t' : _ & ident t' * type.for_each_lhs_of_arrow expr t' }%type + := match t return expr t -> _ with + | type.base _ => fun e => let 'existT t (f, args) := invert_App_curried e tt in + (idc <- invert_Ident f; + Some (existT _ t (idc, args)))%option + | _ => fun _ => None + end e. End with_var_gen. Section with_var. diff --git a/src/Experiments/NewPipeline/Rewriter.v b/src/Experiments/NewPipeline/Rewriter.v index eecf89832..1d0354713 100644 --- a/src/Experiments/NewPipeline/Rewriter.v +++ b/src/Experiments/NewPipeline/Rewriter.v @@ -80,6 +80,9 @@ Module Compilers. Notation "x + y" := (#ident.Z_add @ x @ y) : pattern_scope. Notation "x / y" := (#ident.Z_div @ x @ y) : pattern_scope. Notation "x * y" := (#ident.Z_mul @ x @ y) : pattern_scope. + Notation "x >> y" := (#ident.Z_shiftr @ x @ y) : pattern_scope. + Notation "x << y" := (#ident.Z_shiftl @ x @ y) : pattern_scope. + Notation "x &' y" := (#ident.Z_land @ x @ y) : pattern_scope. Notation "x 'mod' y" := (#ident.Z_modulo @ x @ y)%pattern : pattern_scope. Notation "- x" := (#ident.Z_opp @ x) : pattern_scope. End Notations. @@ -1293,20 +1296,12 @@ In the RHS, the follow notation applies: ; make_rewrite ((-??{ℤ}) * ??{ℤ} ) (fun x y => -(x * y)) ; make_rewrite ( ??{ℤ} * (-??{ℤ})) (fun x y => -(x * y)) - ; make_rewrite (??{ℤ} * #?ℤ) (fun x y => x << (Z.log2 y) when Z.eqb y (2^Z.log2 y)) - ; make_rewrite (#?ℤ * ??{ℤ}) (fun y x => x << (Z.log2 y) when Z.eqb y (2^Z.log2 y)) - ; make_rewrite (??{ℤ} / #?ℤ) (fun x y => x >> (Z.log2 y) when Z.eqb y (2^Z.log2 y)) - ; make_rewrite (??{ℤ} mod #?ℤ) (fun x y => #(ident.Z_land (y-1)) @ x when Z.eqb y (2^Z.log2 y)) + ; make_rewrite (??{ℤ} * #?ℤ) (fun x y => x << ##(Z.log2 y) when Z.eqb y (2^Z.log2 y)) + ; make_rewrite (#?ℤ * ??{ℤ}) (fun y x => x << ##(Z.log2 y) when Z.eqb y (2^Z.log2 y)) + ; make_rewrite (??{ℤ} / #?ℤ) (fun x y => x >> ##(Z.log2 y) when Z.eqb y (2^Z.log2 y)) + ; make_rewrite (??{ℤ} mod #?ℤ) (fun x y => x &' ##(y-1)%Z when Z.eqb y (2^Z.log2 y)) ; make_rewrite (-(-??{ℤ})) (fun v => v) - (** TODO(jadep): These next two are only here for demonstration purposes; remove them once you no longer need it as a template *) - (* if it's a concrete pair, we can opp the second value *) - ; make_rewrite (#pident.Z_neg_snd @ (??{ℤ}, ??{ℤ})) (fun x y => (x, -y)) - (* if it's not a concrete pair, let-bind the pair and negate the second element *) - ; make_rewrite - (#pident.Z_neg_snd @ ??{ℤ * ℤ}) - (fun xy => ret (UnderLets.UnderLet xy (fun xyv => UnderLets.Base (#ident.fst @ $xyv, -(#ident.snd @ $xyv))))) - ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (##0, ##0)%Z when Z.eqb xx 0) ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (##0, ##0)%Z when Z.eqb xx 0) ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (y, ##0)%Z when Z.eqb xx 1) @@ -1314,104 +1309,69 @@ In the RHS, the follow notation applies: ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (-y, ##0%Z) when Z.eqb xx (-1)) ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (-y, ##0%Z) when Z.eqb xx (-1)) - ; make_rewrite (#pident.Z_add_get_carry @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s xx y => (y, ##0%Z) when Z.eqb xx 0) - ; make_rewrite (#pident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s y xx => (y, ##0%Z) when Z.eqb xx 0) - - ; make_rewrite (#pident.Z_add_with_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun c x y => x + y when Z.eqb c 0) + ; make_rewrite (#pident.Z_add_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ}) (fun s xx y => (y, ##0) when xx =? 0) + ; make_rewrite (#pident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ) (fun s y xx => (y, ##0) when xx =? 0) + ; make_rewrite (#pident.Z_add_with_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun cc x y => x + y when cc =? 0) ; make_rewrite (#pident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ #?ℤ @ ??{ℤ}) (fun s cc xx y => (y, ##0) when (cc =? 0) && (xx =? 0)) ; make_rewrite (#pident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??{ℤ} @ #?ℤ) (fun s cc y xx => (y, ##0) when (cc =? 0) && (xx =? 0)) ; make_rewrite (* carry = 0: ADC x y -> ADD x y *) - (#pident.Z_add_with_get_carry @ #?ℤ @ #?ℤ @ ??{ℤ} @ ??{ℤ}) - (fun s cc x y => #(ident.Z_add_get_carry_concrete s) @ x @ y when cc =? 0) + (#pident.Z_add_with_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ} @ ??{ℤ}) + (fun s cc x y => #ident.Z_add_get_carry @ s @ x @ y when cc =? 0) ; make_rewrite (* ADC 0 0 -> (ADX 0 0, 0) *) - (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ #?ℤ @ #?ℤ) - (fun s c xx yy => #ident.Z_add_with_carry @ ##s @ ##xx @ ##yy when (xx =? 0) && (yy =? 0)) + (#pident.Z_add_with_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ @ #?ℤ) + (fun s c xx yy => #ident.Z_add_with_carry @ s @ ##xx @ ##yy when (xx =? 0) && (yy =? 0)) ; make_rewrite - (#pident.Z_add_get_carry @ #?ℤ @ (-??{ℤ}) @ ??{ℤ}) + (#pident.Z_add_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ}) (fun s y x => ret (UnderLets.UnderLet - (#(ident.Z_sub_get_borrow_concrete s) @ x @ y) + (#ident.Z_sub_get_borrow @ s @ x @ y) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) ; make_rewrite - (#pident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ (-??{ℤ})) + (#pident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ (-??{ℤ})) (fun s x y => ret (UnderLets.UnderLet - (#(ident.Z_sub_get_borrow_concrete s) @ x @ y) - (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) - - - ; make_rewrite - (#pident.Z_add_with_get_carry @ #?ℤ @ (-??{ℤ}) @ (-??{ℤ}) @ ??{ℤ}) - (fun s c y x => ret (UnderLets.UnderLet - (#(ident.Z_sub_with_get_borrow_concrete s) @ c @ x @ y) - (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) - ; make_rewrite - (#pident.Z_add_with_get_carry @ #?ℤ @ (-??{ℤ}) @ ??{ℤ} @ (-??{ℤ})) - (fun s c x y => ret (UnderLets.UnderLet - (#(ident.Z_sub_with_get_borrow_concrete s) @ c @ x @ y) - (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) - - ; make_rewrite - (#pident.Z_add_get_carry_concrete @ (-??{ℤ}) @ ??{ℤ}) - (fun s y x => ret (UnderLets.UnderLet - (#(ident.Z_sub_get_borrow_concrete s) @ x @ y) + (#ident.Z_sub_get_borrow @ s @ x @ y) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) ; make_rewrite - (#pident.Z_add_get_carry_concrete @ ??{ℤ} @ (-??{ℤ})) - (fun s x y => ret (UnderLets.UnderLet - (#(ident.Z_sub_get_borrow_concrete s) @ x @ y) - (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) - ; make_rewrite - (#pident.Z_add_get_carry_concrete @ #?ℤ @ ??{ℤ}) + (#pident.Z_add_get_carry @ ??{ℤ} @ #?ℤ @ ??{ℤ}) (fun s yy x => ret (UnderLets.UnderLet - (#(ident.Z_sub_get_borrow_concrete s) @ x @ ##(-yy)%Z) + (#ident.Z_sub_get_borrow @ s @ x @ ##(-yy)%Z) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc)))) when yy <=? 0) ; make_rewrite - (#pident.Z_add_get_carry_concrete @ ??{ℤ} @ #?ℤ) + (#pident.Z_add_get_carry @ ??{ℤ} @ ??{ℤ} @ #?ℤ) (fun s x yy => ret (UnderLets.UnderLet - (#(ident.Z_sub_get_borrow_concrete s) @ x @ ##(-yy)%Z) + (#ident.Z_sub_get_borrow @ s @ x @ ##(-yy)%Z) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc)))) when yy <=? 0) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ (-??{ℤ}) @ (-??{ℤ}) @ ??{ℤ}) + (#pident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ (-??{ℤ}) @ ??{ℤ}) (fun s c y x => ret (UnderLets.UnderLet - (#(ident.Z_sub_with_get_borrow_concrete s) @ c @ x @ y) + (#ident.Z_sub_with_get_borrow @ s @ c @ x @ y) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ (-??{ℤ}) @ ??{ℤ} @ (-??{ℤ})) + (#pident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ} @ (-??{ℤ})) (fun s c x y => ret (UnderLets.UnderLet - (#(ident.Z_sub_with_get_borrow_concrete s) @ c @ x @ y) + (#ident.Z_sub_with_get_borrow @ s @ c @ x @ y) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc))))) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ (-??{ℤ}) @ #?ℤ @ ??{ℤ}) + (#pident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ #?ℤ @ ??{ℤ}) (fun s c yy x => ret (UnderLets.UnderLet - (#(ident.Z_sub_with_get_borrow_concrete s) @ c @ x @ ##(-yy)%Z) + (#ident.Z_sub_with_get_borrow @ s @ c @ x @ ##(-yy)%Z) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc)))) when yy <=? 0) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ (-??{ℤ}) @ ??{ℤ} @ #?ℤ) + (#pident.Z_add_with_get_carry @ ??{ℤ} @ (-??{ℤ}) @ ??{ℤ} @ #?ℤ) (fun s c x yy => ret (UnderLets.UnderLet - (#(ident.Z_sub_with_get_borrow_concrete s) @ c @ x @ ##(-yy)%Z) + (#ident.Z_sub_with_get_borrow @ s @ c @ x @ ##(-yy)%Z) (fun vc => UnderLets.Base (#ident.fst @ $vc, -(#ident.snd @ $vc)))) when yy <=? 0) - ; make_rewrite (#pident.Z_add_get_carry_concrete @ #?ℤ @ ??{ℤ}) (fun s xx y => (y, ##0) when xx =? 0) - ; make_rewrite (#pident.Z_add_get_carry_concrete @ ??{ℤ} @ #?ℤ) (fun s y xx => (y, ##0) when xx =? 0) - - (** XXX TODO: Do we still need the _concrete versions? *) - ; make_rewrite (#pident.Z_mul_split @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun s x y => #(ident.Z_mul_split_concrete s) @ x @ y) - ; make_rewrite (#pident.Z_rshi @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ #?ℤ) (fun x y z a => #(ident.Z_rshi_concrete x a) @ y @ z) - ; make_rewrite (#pident.Z_cc_m @ #?ℤ @ ??{ℤ}) (fun x y => #(ident.Z_cc_m_concrete x) @ y) - ; make_rewrite (#pident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun s x y => #(ident.Z_add_get_carry_concrete s) @ x @ y) - ; make_rewrite (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun s c x y => #(ident.Z_add_with_get_carry_concrete s) @ c @ x @ y) - ; make_rewrite (#pident.Z_sub_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun s x y => #(ident.Z_sub_get_borrow_concrete s) @ x @ y) - ; make_rewrite (#pident.Z_sub_with_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun s x y b => #(ident.Z_sub_with_get_borrow_concrete s) @ x @ y @ b) ; make_rewrite_step (* _step, so that if one of the arguments is concrete, we automatically get the rewrite rule for [Z_cast] applying to it *) (#pident.Z_cast2 @ (??{ℤ}, ??{ℤ})) (fun r x y => (#(ident.Z_cast (fst r)) @ $x, #(ident.Z_cast (snd r)) @ $y)) @@ -1458,19 +1418,19 @@ In the RHS, the follow notation applies: (Z.add_get_carry_concrete 2^256) @@ (?x, ?y) --> (add 0) @@ (y, x) *) make_rewrite - (#pident.Z_add_get_carry_concrete @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ})) - (fun s x offset y => #(ident.fancy_add (Z.log2 s) offset) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ)) + (fun s x y offset => #(ident.fancy_add (Z.log2 s) offset) @ (x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_get_carry_concrete @ (#pident.Z_shiftl @ ??{ℤ}) @ ??{ℤ}) - (fun s offset y x => #(ident.fancy_add (Z.log2 s) offset) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_get_carry @ #?ℤ @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ) @ ??{ℤ}) + (fun s y offset x => #(ident.fancy_add (Z.log2 s) offset) @ (x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_get_carry_concrete @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ})) - (fun s x offset y => #(ident.fancy_add (Z.log2 s) (-offset)) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun s x y offset => #(ident.fancy_add (Z.log2 s) (-offset)) @ (x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_get_carry_concrete @ (#pident.Z_shiftr @ ??{ℤ}) @ ??{ℤ}) - (fun s offset y x => #(ident.fancy_add (Z.log2 s) (-offset)) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_get_carry @ #?ℤ @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) @ ??{ℤ}) + (fun s y offset x => #(ident.fancy_add (Z.log2 s) (-offset)) @ (x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_get_carry_concrete @ ??{ℤ} @ ??{ℤ}) + (#pident.Z_add_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun s x y => #(ident.fancy_add (Z.log2 s) 0) @ (x, y) when s =? 2^Z.log2 s) (* (Z.add_with_get_carry_concrete 2^256) @@ (?c, ?x, ?y << 128) --> (addc 128) @@ (c, x, y) @@ -1480,19 +1440,19 @@ In the RHS, the follow notation applies: (Z.add_with_get_carry_concrete 2^256) @@ (?c, ?x, ?y) --> (addc 0) @@ (c, y, x) *) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ})) - (fun s c x offset y => #(ident.fancy_addc (Z.log2 s) offset) @ (c, x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ)) + (fun s c x y offset => #(ident.fancy_addc (Z.log2 s) offset) @ (c, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ}) @ ??{ℤ}) - (fun s c offset y x => #(ident.fancy_addc (Z.log2 s) offset) @ (c, x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ) @ ??{ℤ}) + (fun s c y offset x => #(ident.fancy_addc (Z.log2 s) offset) @ (c, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ})) - (fun s c x offset y => #(ident.fancy_addc (Z.log2 s) (-offset)) @ (c, x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun s c x y offset => #(ident.fancy_addc (Z.log2 s) (-offset)) @ (c, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ}) @ ??{ℤ}) - (fun s c offset y x => #(ident.fancy_addc (Z.log2 s) (-offset)) @ (c, x, y) when s =? 2^Z.log2 s) + (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) @ ??{ℤ}) + (fun s c y offset x => #(ident.fancy_addc (Z.log2 s) (-offset)) @ (c, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_add_with_get_carry_concrete @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) + (#pident.Z_add_with_get_carry @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun s c x y => #(ident.fancy_addc (Z.log2 s) 0) @ (c, x, y) when s =? 2^Z.log2 s) (* (Z.sub_get_borrow_concrete 2^256) @@ (?x, ?y << 128) --> (sub 128) @@ (x, y) @@ -1500,13 +1460,13 @@ In the RHS, the follow notation applies: (Z.sub_get_borrow_concrete 2^256) @@ (?x, ?y) --> (sub 0) @@ (y, x) *) ; make_rewrite - (#pident.Z_sub_get_borrow_concrete @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ})) - (fun s x offset y => #(ident.fancy_sub (Z.log2 s) offset) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_sub_get_borrow @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ)) + (fun s x y offset => #(ident.fancy_sub (Z.log2 s) offset) @ (x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_sub_get_borrow_concrete @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ})) - (fun s x offset y => #(ident.fancy_sub (Z.log2 s) (-offset)) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_sub_get_borrow @ #?ℤ @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun s x y offset => #(ident.fancy_sub (Z.log2 s) (-offset)) @ (x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_sub_get_borrow_concrete @ ??{ℤ} @ ??{ℤ}) + (#pident.Z_sub_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ}) (fun s x y => #(ident.fancy_sub (Z.log2 s) 0) @ (x, y) when s =? 2^Z.log2 s) (* (Z.sub_with_get_borrow_concrete 2^256) @@ (?c, ?x, ?y << 128) --> (subb 128) @@ (c, x, y) @@ -1514,29 +1474,32 @@ In the RHS, the follow notation applies: (Z.sub_with_get_borrow_concrete 2^256) @@ (?c, ?x, ?y) --> (subb 0) @@ (c, y, x) *) ; make_rewrite - (#pident.Z_sub_with_get_borrow_concrete @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ})) - (fun s b x offset y => #(ident.fancy_subb (Z.log2 s) offset) @ (b, x, y) when s =? 2^Z.log2 s) + (#pident.Z_sub_with_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftl @ ??{ℤ} @ #?ℤ)) + (fun s b x y offset => #(ident.fancy_subb (Z.log2 s) offset) @ (b, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_sub_with_get_borrow_concrete @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ})) - (fun s b x offset y => #(ident.fancy_subb (Z.log2 s) (-offset)) @ (b, x, y) when s =? 2^Z.log2 s) + (#pident.Z_sub_with_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun s b x y offset => #(ident.fancy_subb (Z.log2 s) (-offset)) @ (b, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_sub_with_get_borrow_concrete @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) + (#pident.Z_sub_with_get_borrow @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun s b x y => #(ident.fancy_subb (Z.log2 s) 0) @ (b, x, y) when s =? 2^Z.log2 s) (*(Z.rshi_concrete 2^256 ?n) @@ (?c, ?x, ?y) --> (rshi n) @@ (x, y)*) ; make_rewrite - (#pident.Z_rshi_concrete @ ??{ℤ} @ ??{ℤ}) - (fun '((s, n)%core) x y => #(ident.fancy_rshi (Z.log2 s) n) @ (x, y) when s =? 2^Z.log2 s) + (#pident.Z_rshi @ #?ℤ @ ??{ℤ} @ ??{ℤ} @ #?ℤ) + (fun s x y n => #(ident.fancy_rshi (Z.log2 s) n) @ (x, y) when s =? 2^Z.log2 s) (* Z.zselect @@ (Z.cc_m_concrete 2^256 ?c, ?x, ?y) --> selm @@ (c, x, y) Z.zselect @@ (?c &' 1, ?x, ?y) --> sell @@ (c, x, y) Z.zselect @@ (?c, ?x, ?y) --> selc @@ (c, x, y) *) ; make_rewrite - (#pident.Z_zselect @ (#pident.Z_cc_m_concrete @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ}) + (#pident.Z_zselect @ (#pident.Z_cc_m @ #?ℤ @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ}) (fun s c x y => #(ident.fancy_selm (Z.log2 s)) @ (c, x, y) when s =? 2^Z.log2 s) ; make_rewrite - (#pident.Z_zselect @ (#pident.Z_land @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ}) + (#pident.Z_zselect @ (#pident.Z_land @ #?ℤ @ ??{ℤ}) @ ??{ℤ} @ ??{ℤ}) (fun mask c x y => #ident.fancy_sell @ (c, x, y) when mask =? 1) + ; make_rewrite + (#pident.Z_zselect @ (#pident.Z_land @ ??{ℤ} @ #?ℤ) @ ??{ℤ} @ ??{ℤ}) + (fun c mask x y => #ident.fancy_sell @ (c, x, y) when mask =? 1) ; make_rewrite (#pident.Z_zselect @ ??{ℤ} @ ??{ℤ} @ ??{ℤ}) (fun c x y => #ident.fancy_selc @ (c, x, y)) @@ -1552,45 +1515,72 @@ Z.mul @@ (?x >> 128, ?y >> 128) --> mulhh @@ (x, y) *) (* literal on left *) ; make_rewrite - (#?ℤ * (#pident.Z_land @ ??{ℤ})) + (#?ℤ * (#pident.Z_land @ ??{ℤ} @ #?ℤ)) + (fun x y mask => let s := (2*Z.log2_up mask)%Z in x <---- invert_low s x; #(ident.fancy_mulll s) @ (##x, y) when (mask =? 2^(s/2)-1)) + ; make_rewrite + (#?ℤ * (#pident.Z_land @ #?ℤ @ ??{ℤ})) (fun x mask y => let s := (2*Z.log2_up mask)%Z in x <---- invert_low s x; #(ident.fancy_mulll s) @ (##x, y) when (mask =? 2^(s/2)-1)) ; make_rewrite - (#?ℤ * (#pident.Z_shiftr @ ??{ℤ})) - (fun x offset y => let s := (2*offset)%Z in x <---- invert_low s x; #(ident.fancy_mullh s) @ (##x, y)) + (#?ℤ * (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun x y offset => let s := (2*offset)%Z in x <---- invert_low s x; #(ident.fancy_mullh s) @ (##x, y)) ; make_rewrite - (#?ℤ * (#pident.Z_land @ ??{ℤ})) + (#?ℤ * (#pident.Z_land @ #?ℤ @ ??{ℤ})) (fun x mask y => let s := (2*Z.log2_up mask)%Z in x <---- invert_high s x; #(ident.fancy_mulhl s) @ (##x, y) when mask =? 2^(s/2)-1) ; make_rewrite - (#?ℤ * (#pident.Z_shiftr @ ??{ℤ})) - (fun x offset y => let s := (2*offset)%Z in x <---- invert_high s x; #(ident.fancy_mulhh s) @ (##x, y)) + (#?ℤ * (#pident.Z_land @ ??{ℤ} @ #?ℤ)) + (fun x y mask => let s := (2*Z.log2_up mask)%Z in x <---- invert_high s x; #(ident.fancy_mulhl s) @ (##x, y) when mask =? 2^(s/2)-1) + ; make_rewrite + (#?ℤ * (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun x y offset => let s := (2*offset)%Z in x <---- invert_high s x; #(ident.fancy_mulhh s) @ (##x, y)) (* literal on right *) ; make_rewrite - ((#pident.Z_land @ ??{ℤ}) * #?ℤ) + ((#pident.Z_land @ #?ℤ @ ??{ℤ}) * #?ℤ) (fun mask x y => let s := (2*Z.log2_up mask)%Z in y <---- invert_low s y; #(ident.fancy_mulll s) @ (x, ##y) when (mask =? 2^(s/2)-1)) ; make_rewrite - ((#pident.Z_land @ ??{ℤ}) * #?ℤ) + ((#pident.Z_land @ ??{ℤ} @ #?ℤ) * #?ℤ) + (fun x mask y => let s := (2*Z.log2_up mask)%Z in y <---- invert_low s y; #(ident.fancy_mulll s) @ (x, ##y) when (mask =? 2^(s/2)-1)) + ; make_rewrite + ((#pident.Z_land @ #?ℤ @ ??{ℤ}) * #?ℤ) (fun mask x y => let s := (2*Z.log2_up mask)%Z in y <---- invert_high s y; #(ident.fancy_mullh s) @ (x, ##y) when mask =? 2^(s/2)-1) ; make_rewrite - ((#pident.Z_shiftr @ ??{ℤ}) * #?ℤ) - (fun offset x y => let s := (2*offset)%Z in y <---- invert_low s y; #(ident.fancy_mulhl s) @ (x, ##y)) + ((#pident.Z_land @ ??{ℤ} @ #?ℤ) * #?ℤ) + (fun x mask y => let s := (2*Z.log2_up mask)%Z in y <---- invert_high s y; #(ident.fancy_mullh s) @ (x, ##y) when mask =? 2^(s/2)-1) + ; make_rewrite + ((#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) * #?ℤ) + (fun x offset y => let s := (2*offset)%Z in y <---- invert_low s y; #(ident.fancy_mulhl s) @ (x, ##y)) ; make_rewrite - ((#pident.Z_shiftr @ ??{ℤ}) * #?ℤ) - (fun offset x y => let s := (2*offset)%Z in y <---- invert_high s y; #(ident.fancy_mulhh s) @ (x, ##y)) + ((#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) * #?ℤ) + (fun x offset y => let s := (2*offset)%Z in y <---- invert_high s y; #(ident.fancy_mulhh s) @ (x, ##y)) (* no literal *) ; make_rewrite - ((#pident.Z_land @ ??{ℤ}) * (#pident.Z_land @ ??{ℤ})) + ((#pident.Z_land @ #?ℤ @ ??{ℤ}) * (#pident.Z_land @ #?ℤ @ ??{ℤ})) (fun mask1 x mask2 y => let s := (2*Z.log2_up mask1)%Z in #(ident.fancy_mulll s) @ (x, y) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1)) ; make_rewrite - ((#pident.Z_land @ ??{ℤ}) * (#pident.Z_shiftr @ ??{ℤ})) - (fun mask x offset y => let s := (2*offset)%Z in #(ident.fancy_mullh s) @ (x, y) when mask =? 2^(s/2)-1) + ((#pident.Z_land @ ??{ℤ} @ #?ℤ) * (#pident.Z_land @ #?ℤ @ ??{ℤ})) + (fun x mask1 mask2 y => let s := (2*Z.log2_up mask1)%Z in #(ident.fancy_mulll s) @ (x, y) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1)) + ; make_rewrite + ((#pident.Z_land @ #?ℤ @ ??{ℤ}) * (#pident.Z_land @ ??{ℤ} @ #?ℤ)) + (fun mask1 x y mask2 => let s := (2*Z.log2_up mask1)%Z in #(ident.fancy_mulll s) @ (x, y) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1)) + ; make_rewrite + ((#pident.Z_land @ ??{ℤ} @ #?ℤ) * (#pident.Z_land @ ??{ℤ} @ #?ℤ)) + (fun x mask1 y mask2 => let s := (2*Z.log2_up mask1)%Z in #(ident.fancy_mulll s) @ (x, y) when (mask1 =? 2^(s/2)-1) && (mask2 =? 2^(s/2)-1)) + ; make_rewrite + ((#pident.Z_land @ #?ℤ @ ??{ℤ}) * (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun mask x y offset => let s := (2*offset)%Z in #(ident.fancy_mullh s) @ (x, y) when mask =? 2^(s/2)-1) + ; make_rewrite + ((#pident.Z_land @ ??{ℤ} @ #?ℤ) * (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun x mask y offset => let s := (2*offset)%Z in #(ident.fancy_mullh s) @ (x, y) when mask =? 2^(s/2)-1) + ; make_rewrite + ((#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) * (#pident.Z_land @ #?ℤ @ ??{ℤ})) + (fun x offset mask y => let s := (2*offset)%Z in #(ident.fancy_mulhl s) @ (x, y) when mask =? 2^(s/2)-1) ; make_rewrite - ((#pident.Z_shiftr @ ??{ℤ}) * (#pident.Z_land @ ??{ℤ})) - (fun offset x mask y => let s := (2*offset)%Z in #(ident.fancy_mulhl s) @ (x, y) when mask =? 2^(s/2)-1) + ((#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) * (#pident.Z_land @ ??{ℤ} @ #?ℤ)) + (fun x offset y mask => let s := (2*offset)%Z in #(ident.fancy_mulhl s) @ (x, y) when mask =? 2^(s/2)-1) ; make_rewrite - ((#pident.Z_shiftr @ ??{ℤ}) * (#pident.Z_shiftr @ ??{ℤ})) - (fun offset1 x offset2 y => let s := (2*offset1)%Z in #(ident.fancy_mulhh s) @ (x, y) when offset1 =? offset2) + ((#pident.Z_shiftr @ ??{ℤ} @ #?ℤ) * (#pident.Z_shiftr @ ??{ℤ} @ #?ℤ)) + (fun x offset1 y offset2 => let s := (2*offset1)%Z in #(ident.fancy_mulhh s) @ (x, y) when offset1 =? offset2) ]%list%pattern%cps%option%under_lets%Z%bool. diff --git a/src/Experiments/NewPipeline/SlowPrimeSynthesisExamples.v b/src/Experiments/NewPipeline/SlowPrimeSynthesisExamples.v index 4cfefcabb..3d0506a66 100644 --- a/src/Experiments/NewPipeline/SlowPrimeSynthesisExamples.v +++ b/src/Experiments/NewPipeline/SlowPrimeSynthesisExamples.v @@ -44,6 +44,14 @@ Module X25519_64. SuchThat (ropp_correctT n s c machine_wordsize base_51_opp) As base_51_opp_correct. Proof. Time solve_ropp machine_wordsize. Time Qed. + Derive base_51_to_bytes + SuchThat (rto_bytes_correctT n s c machine_wordsize base_51_to_bytes) + As base_51_to_bytes_correct. + Proof. Time solve_rto_bytes machine_wordsize. Time Qed. + Derive base_51_from_bytes + SuchThat (rfrom_bytes_correctT n s c machine_wordsize base_51_from_bytes) + As base_51_from_bytes_correct. + Proof. Time solve_rfrom_bytes machine_wordsize. Time Qed. Derive base_51_encode SuchThat (rencode_correctT n s c machine_wordsize base_51_encode) As base_51_encode_correct. @@ -60,7 +68,7 @@ Module X25519_64. : check_args n s c machine_wordsize (ErrorT.Success tt) = ErrorT.Success tt. Proof. vm_compute; reflexivity. Qed. - Definition base_51_good : GoodT n s c + Definition base_51_good : GoodT n s c machine_wordsize := Good n s c machine_wordsize base_51_curve_good base_51_carry_mul_correct @@ -71,7 +79,9 @@ Module X25519_64. base_51_opp_correct base_51_zero_correct base_51_one_correct - base_51_encode_correct. + base_51_encode_correct + base_51_to_bytes_correct + base_51_from_bytes_correct. Print Assumptions base_51_good. Import PrintingNotations. @@ -141,7 +151,7 @@ fun var : type -> Type => *) Compute Compilers.ToString.C.ToFunctionString - "fecarry_mul" base_51_carry_mul + "" "fecarry_mul" base_51_carry_mul None (Some loose_bounds, (Some loose_bounds, tt)). (* void fecarry_mul(uint64_t[5] x1, uint64_t[5] x2, uint64_t[5] x3) { @@ -160,7 +170,7 @@ void fecarry_mul(uint64_t[5] x1, uint64_t[5] x2, uint64_t[5] x3) { } *) Compute Compilers.ToString.C.ToFunctionString - "fesub" base_51_sub + "" "fesub" base_51_sub None (Some tight_bounds, (Some tight_bounds, tt)). (* void fesub(uint64_t[5] x1, uint64_t[5] x2, uint64_t[5] x3) { @@ -203,6 +213,14 @@ Module X25519_32. SuchThat (ropp_correctT n s c machine_wordsize base_25p5_opp) As base_25p5_opp_correct. Proof. Time solve_ropp machine_wordsize. Time Qed. + Derive base_25p5_to_bytes + SuchThat (rto_bytes_correctT n s c machine_wordsize base_25p5_to_bytes) + As base_25p5_to_bytes_correct. + Proof. Time solve_rto_bytes machine_wordsize. Time Qed. + Derive base_25p5_from_bytes + SuchThat (rfrom_bytes_correctT n s c machine_wordsize base_25p5_from_bytes) + As base_25p5_from_bytes_correct. + Proof. Time solve_rfrom_bytes machine_wordsize. Time Qed. Derive base_25p5_encode SuchThat (rencode_correctT n s c machine_wordsize base_25p5_encode) As base_25p5_encode_correct. @@ -219,7 +237,7 @@ Module X25519_32. : check_args n s c machine_wordsize (ErrorT.Success tt) = ErrorT.Success tt. Proof. vm_compute; reflexivity. Qed. - Definition base_25p5_good : GoodT n s c + Definition base_25p5_good : GoodT n s c machine_wordsize := Good n s c machine_wordsize base_25p5_curve_good base_25p5_carry_mul_correct @@ -230,7 +248,9 @@ Module X25519_32. base_25p5_opp_correct base_25p5_zero_correct base_25p5_one_correct - base_25p5_encode_correct. + base_25p5_encode_correct + base_25p5_to_bytes_correct + base_25p5_from_bytes_correct. Print Assumptions base_25p5_good. Import PrintingNotations. @@ -416,11 +436,11 @@ Module P192_64. Set Printing Depth 100000. Local Notation "'mul64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, _)%core) @ (#(Z_mul_split_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_add_get_carry_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc64' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_add_with_get_carry_concrete 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adx64' '(' c ',' x ',' y ')'" := (#(Z_cast bool) @ (#Z_add_with_carry @ c @ x @ y))%expr (at level 50) : expr_scope. @@ -486,11 +506,11 @@ Module P192_32. Set Printing Depth 100000. Local Notation "'mul32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, _)%core) @ (#(Z_mul_split_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_add_get_carry_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc32' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_add_with_get_carry_concrete 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. Print mulmod. (* @@ -684,15 +704,15 @@ Module P384_32. Set Printing Depth 100000. Local Notation "'mul32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, _)%core) @ (#(Z_mul_split_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_add_get_carry_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc32' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_add_with_get_carry_concrete 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'sub32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_sub_get_borrow_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_sub_get_borrow @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'sbb32' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_sub_with_get_borrow_concrete 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_sub_with_get_borrow @ #(ident.Literal (t:=base.type.Z) 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. Print mulmod. @@ -716,11 +736,11 @@ Module P256_32. Set Printing Width 100000. Local Notation "'mul32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, _)%core) @ (#(Z_mul_split_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add32' '(' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_add_get_carry_concrete 4294967296) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 4294967296) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc32' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint32, bool)%core) @ (#(Z_add_with_get_carry_concrete 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint32, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 4294967296) @ c @ x @ y))%expr (at level 50) : expr_scope. (* Print is too slow *) diff --git a/src/Experiments/NewPipeline/Toplevel1.v b/src/Experiments/NewPipeline/Toplevel1.v index a1a97b873..27d1b5806 100644 --- a/src/Experiments/NewPipeline/Toplevel1.v +++ b/src/Experiments/NewPipeline/Toplevel1.v @@ -2,6 +2,7 @@ Require Import Coq.ZArith.ZArith Coq.micromega.Lia. Require Import Coq.derive.Derive. Require Import Coq.Bool.Bool. Require Import Coq.Strings.String. +Require Import Coq.MSets.MSetPositive. Require Import Coq.Lists.List. Require Crypto.Util.Strings.String. Require Import Crypto.Util.Strings.Decimal. @@ -286,132 +287,6 @@ Notation "x" := (expr.Var x) (only printing, at level 9) : expr_scope. Axiom admit_pf : False. Notation admit := (match admit_pf with end). -Ltac cache_reify _ := - intros; - etransitivity; - [ - | repeat match goal with |- _ = ?f' ?x => is_var x; apply (f_equal (fun f => f _)) end; - Reify_rhs (); - reflexivity ]; - subst_evars; - reflexivity. - -Create HintDb reify_gen_cache. - -Derive carry_mul_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (f g : list Z) - (n : nat) - (s : Z) - (c : list (Z * Z)) - (idxs : list nat), - Interp (t:=reify_type_of carry_mulmod) - carry_mul_gen limbwidth_num limbwidth_den s c n idxs f g - = carry_mulmod limbwidth_num limbwidth_den s c n idxs f g) - As carry_mul_gen_correct. -Proof. Time cache_reify (). Time Qed. -Hint Extern 1 (_ = carry_mulmod _ _ _ _ _ _ _ _) => simple apply carry_mul_gen_correct : reify_gen_cache. - -Derive carry_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (f : list Z) - (n : nat) - (s : Z) - (c : list (Z * Z)) - (idxs : list nat), - Interp (t:=reify_type_of carrymod) - carry_gen limbwidth_num limbwidth_den s c n idxs f - = carrymod limbwidth_num limbwidth_den s c n idxs f) - As carry_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = carrymod _ _ _ _ _ _ _) => simple apply carry_gen_correct : reify_gen_cache. - -Derive encode_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (v : Z) - (n : nat) - (s : Z) - (c : list (Z * Z)), - Interp (t:=reify_type_of encodemod) - encode_gen limbwidth_num limbwidth_den s c n v - = encodemod limbwidth_num limbwidth_den s c n v) - As encode_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = encodemod _ _ _ _ _ _) => simple apply encode_gen_correct : reify_gen_cache. - -Derive add_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (f g : list Z) - (n : nat), - Interp (t:=reify_type_of addmod) - add_gen limbwidth_num limbwidth_den n f g - = addmod limbwidth_num limbwidth_den n f g) - As add_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = addmod _ _ _ _ _) => simple apply add_gen_correct : reify_gen_cache. -Derive sub_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (n : nat) - (s : Z) - (c : list (Z * Z)) - (coef : Z) - (f g : list Z), - Interp (t:=reify_type_of submod) - sub_gen limbwidth_num limbwidth_den s c n coef f g - = submod limbwidth_num limbwidth_den s c n coef f g) - As sub_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = submod _ _ _ _ _ _ _ _) => simple apply sub_gen_correct : reify_gen_cache. - -Derive opp_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (n : nat) - (s : Z) - (c : list (Z * Z)) - (coef : Z) - (f : list Z), - Interp (t:=reify_type_of oppmod) - opp_gen limbwidth_num limbwidth_den s c n coef f - = oppmod limbwidth_num limbwidth_den s c n coef f) - As opp_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = oppmod _ _ _ _ _ _ _) => simple apply opp_gen_correct : reify_gen_cache. - -Definition zeromod limbwidth_num limbwidth_den n s c := encodemod limbwidth_num limbwidth_den n s c 0. -Definition onemod limbwidth_num limbwidth_den n s c := encodemod limbwidth_num limbwidth_den n s c 1. - -Derive zero_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (n : nat) - (s : Z) - (c : list (Z * Z)), - Interp (t:=reify_type_of zeromod) - zero_gen limbwidth_num limbwidth_den s c n - = zeromod limbwidth_num limbwidth_den s c n) - As zero_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = zeromod _ _ _ _ _) => simple apply zero_gen_correct : reify_gen_cache. - -Derive one_gen - SuchThat (forall (limbwidth_num limbwidth_den : Z) - (n : nat) - (s : Z) - (c : list (Z * Z)), - Interp (t:=reify_type_of onemod) - one_gen limbwidth_num limbwidth_den s c n - = onemod limbwidth_num limbwidth_den s c n) - As one_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = onemod _ _ _ _ _) => simple apply one_gen_correct : reify_gen_cache. - -Derive id_gen - SuchThat (forall (ls : list Z), - Interp (t:=reify_type_of (@id (list Z))) - id_gen ls - = id ls) - As id_gen_correct. -Proof. cache_reify (). Qed. -Hint Extern 1 (_ = id _) => simple apply id_gen_correct : reify_gen_cache. Module Pipeline. Import GeneralizeVar. @@ -426,59 +301,13 @@ Module Pipeline. | Value_not_ltZ (descr : string) (lhs rhs : Z) | Values_not_provably_distinctZ (descr : string) (lhs rhs : Z) | Values_not_provably_equalZ (descr : string) (lhs rhs : Z) - | Stringification_failed {t} (e : @Compilers.defaults.Expr t). + | Values_not_provably_equal_listZ (descr : string) (lhs rhs : list Z) + | Stringification_failed {t} (e : @Compilers.defaults.Expr t) (err : string). Notation ErrorT := (ErrorT ErrorMessage). Section show. Local Open Scope string_scope. - Definition show_prim_zrange_opt_interp {t:base.type.base} - : Show (ZRange.type.base.option.interp t) - := match t return Show (ZRange.type.base.option.interp t) with - | base.type.unit => _ - | base.type.Z => _ - | base.type.nat => _ - | base.type.bool => _ - end. - Global Existing Instance show_prim_zrange_opt_interp. - Fixpoint show_base_zrange_opt_interp {t} : Show (ZRange.type.base.option.interp t) - := fun parens - => match t return ZRange.type.base.option.interp t -> string with - | base.type.type_base t - => fun v : ZRange.type.base.option.interp t - => @show_prim_zrange_opt_interp t parens v - | base.type.prod A B - => fun '(a, b) - => "(" ++ @show_base_zrange_opt_interp A false a - ++ ", " ++ @show_base_zrange_opt_interp B true b - ++ ")" - | base.type.list A - => fun v : option (list (ZRange.type.option.interp A)) - => show parens v - end. - Global Existing Instance show_base_zrange_opt_interp. - Definition show_zrange_opt_interp {t} : Show (ZRange.type.option.interp t) - := fun parens - => match t return ZRange.type.option.interp t -> string with - | type.base t - => fun v : ZRange.type.base.option.interp t - => show parens v - | type.arrow s d => fun _ => "λ" - end. - Global Existing Instance show_zrange_opt_interp. - Fixpoint show_for_each_lhs_of_arrow {base_type} (f : type.type base_type -> Type) (show_f : forall t, Show (f t)) (t : type.type base_type) (p : bool) : type.for_each_lhs_of_arrow f t -> string - := match t return type.for_each_lhs_of_arrow f t -> string with - | type.base t => fun (tt : unit) => show p tt - | type.arrow s d - => fun '((x, xs) : f s * type.for_each_lhs_of_arrow f d) - => let _ : Show (f s) := show_f s in - let _ : Show (type.for_each_lhs_of_arrow f d) := @show_for_each_lhs_of_arrow base_type f show_f d in - show p (x, xs) - end. - Global Instance: forall {base_type f show_f t}, Show (type.for_each_lhs_of_arrow f t) := @show_for_each_lhs_of_arrow. - - Local Notation NewLine := (String "010" "") (only parsing). - Fixpoint find_too_loose_base_bounds {t} : ZRange.type.base.option.interp t -> ZRange.type.base.option.interp t-> bool * list (nat * nat) * list (zrange * zrange) := match t return ZRange.type.base.option.interp t -> ZRange.type.option.interp t-> bool * list (nat * nat) * list (zrange * zrange) with @@ -536,7 +365,7 @@ Module Pipeline. : string := let '(none_some, lens, bs) := find_too_loose_bounds b1 b2 in String.concat - NewLine + String.NewLine ((if none_some then "Found None where Some was expected"::nil else nil) ++ (List.map (A:=nat*nat) @@ -556,12 +385,14 @@ Module Pipeline. => ((["Computed bounds " ++ show true computed_bounds ++ " are not tight enough (expected bounds not looser than " ++ show true expected_bounds ++ ")."]%string) ++ [explain_too_loose_bounds (t:=type.base _) computed_bounds expected_bounds] ++ match ToString.C.ToFunctionLines - "f" syntax_tree None arg_bounds with - | Some E_lines + "" "f" syntax_tree None arg_bounds ZRange.type.base.option.None with + | inl (E_lines, types_used) => ["When doing bounds analysis on the syntax tree:"] ++ E_lines ++ [""] - ++ ["with input bounds " ++ show true arg_bounds ++ "." ++ NewLine]%string - | None => ["(Unprintible syntax tree used in bounds analysis)" ++ NewLine]%string + ++ ["with input bounds " ++ show true arg_bounds ++ "." ++ String.NewLine]%string + | inr errs + => (["(Unprintible syntax tree used in bounds analysis)" ++ String.NewLine]%string) + ++ ["Stringification failed on the syntax tree:"] ++ show_lines false syntax_tree ++ [errs] end)%list | Bounds_analysis_failed => ["Bounds analysis failed."] | Type_too_complicated_for_cps t @@ -573,13 +404,14 @@ Module Pipeline. | Value_not_ltZ descr lhs rhs => ["Value not < (" ++ descr ++ ") : expected " ++ show false lhs ++ " < " ++ show false rhs] | Values_not_provably_distinctZ descr lhs rhs - => ["Values not provalby distinct (" ++ descr ++ ") : expected " ++ show true lhs ++ " ≠ " ++ show true rhs] + => ["Values not provably distinct (" ++ descr ++ ") : expected " ++ show true lhs ++ " ≠ " ++ show true rhs] | Values_not_provably_equalZ descr lhs rhs - => ["Values not provalby equal (" ++ descr ++ ") : expected " ++ show true lhs ++ " = " ++ show true rhs] - | Stringification_failed t e => ["Stringification failed on the syntax tree:"] ++ show_lines false e + | Values_not_provably_equal_listZ descr lhs rhs + => ["Values not provably equal (" ++ descr ++ ") : expected " ++ show true lhs ++ " = " ++ show true rhs] + | Stringification_failed t e err => ["Stringification failed on the syntax tree:"] ++ show_lines false e ++ [err] end. Local Instance show_ErrorMessage : Show ErrorMessage - := fun parens err => String.concat NewLine (show_lines parens err). + := fun parens err => String.concat String.NewLine (show_lines parens err). End show. Definition invert_result {T} (v : ErrorT T) @@ -630,6 +462,7 @@ Module Pipeline. end. Definition BoundsPipelineToStrings + (type_prefix : string) (name : string) (with_dead_code_elimination : bool := true) (with_subst01 : bool) @@ -639,7 +472,7 @@ Module Pipeline. (E : Expr t) arg_bounds out_bounds - : ErrorT (list string) + : ErrorT (list string * ToString.C.ident_infos) := let E := BoundsPipeline (*with_dead_code_elimination*) with_subst01 @@ -648,15 +481,16 @@ Module Pipeline. E arg_bounds out_bounds in match E with | Success E' => let E := ToString.C.ToFunctionLines - name E' None arg_bounds in + type_prefix name E' None arg_bounds out_bounds in match E with - | Some E => Success E - | None => Error (Stringification_failed E') + | inl E => Success E + | inr err => Error (Stringification_failed E' err) end | Error err => Error err end. Definition BoundsPipelineToString + (type_prefix : string) (name : string) (with_dead_code_elimination : bool := true) (with_subst01 : bool) @@ -666,16 +500,16 @@ Module Pipeline. (E : Expr t) arg_bounds out_bounds - : ErrorT string + : ErrorT (string * ToString.C.ident_infos) := let E := BoundsPipelineToStrings - name + type_prefix name (*with_dead_code_elimination*) with_subst01 translate_to_fancy relax_zrange E arg_bounds out_bounds in match E with - | Success E => Success (ToString.C.LinesToString E) + | Success (E, types_used) => Success (ToString.C.LinesToString E, types_used) | Error err => Error err end. @@ -802,6 +636,205 @@ Proof. try (rewrite <- Z.log2_up_le_pow2_full in *; omega). Qed. +Ltac cache_reify _ := + intros; + etransitivity; + [ + | repeat match goal with |- _ = ?f' ?x => is_var x; apply (f_equal (fun f => f _)) end; + Reify_rhs (); + reflexivity ]; + subst_evars; + reflexivity. + +Create HintDb reify_gen_cache. + +Derive carry_mul_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (f g : list Z) + (n : nat) + (s : Z) + (c : list (Z * Z)) + (idxs : list nat), + Interp (t:=reify_type_of carry_mulmod) + carry_mul_gen limbwidth_num limbwidth_den s c n idxs f g + = carry_mulmod limbwidth_num limbwidth_den s c n idxs f g) + As carry_mul_gen_correct. +Proof. Time cache_reify (). Time Qed. +Hint Extern 1 (_ = carry_mulmod _ _ _ _ _ _ _ _) => simple apply carry_mul_gen_correct : reify_gen_cache. + +Derive carry_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (f : list Z) + (n : nat) + (s : Z) + (c : list (Z * Z)) + (idxs : list nat), + Interp (t:=reify_type_of carrymod) + carry_gen limbwidth_num limbwidth_den s c n idxs f + = carrymod limbwidth_num limbwidth_den s c n idxs f) + As carry_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = carrymod _ _ _ _ _ _ _) => simple apply carry_gen_correct : reify_gen_cache. + +Derive encode_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (v : Z) + (n : nat) + (s : Z) + (c : list (Z * Z)), + Interp (t:=reify_type_of encodemod) + encode_gen limbwidth_num limbwidth_den s c n v + = encodemod limbwidth_num limbwidth_den s c n v) + As encode_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = encodemod _ _ _ _ _ _) => simple apply encode_gen_correct : reify_gen_cache. + +Derive add_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (f g : list Z) + (n : nat), + Interp (t:=reify_type_of addmod) + add_gen limbwidth_num limbwidth_den n f g + = addmod limbwidth_num limbwidth_den n f g) + As add_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = addmod _ _ _ _ _) => simple apply add_gen_correct : reify_gen_cache. + +Derive sub_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (s : Z) + (c : list (Z * Z)) + (coef : Z) + (f g : list Z), + Interp (t:=reify_type_of submod) + sub_gen limbwidth_num limbwidth_den s c n coef f g + = submod limbwidth_num limbwidth_den s c n coef f g) + As sub_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = submod _ _ _ _ _ _ _ _) => simple apply sub_gen_correct : reify_gen_cache. +Derive opp_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (s : Z) + (c : list (Z * Z)) + (coef : Z) + (f : list Z), + Interp (t:=reify_type_of oppmod) + opp_gen limbwidth_num limbwidth_den s c n coef f + = oppmod limbwidth_num limbwidth_den s c n coef f) + As opp_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = oppmod _ _ _ _ _ _ _) => simple apply opp_gen_correct : reify_gen_cache. +Definition zeromod limbwidth_num limbwidth_den s c n := encodemod limbwidth_num limbwidth_den s c n 0. +Definition onemod limbwidth_num limbwidth_den s c n := encodemod limbwidth_num limbwidth_den s c n 1. +Definition primemod limbwidth_num limbwidth_den s c n := encodemod limbwidth_num limbwidth_den s c n (s - Associational.eval c). +Derive zero_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (s : Z) + (c : list (Z * Z)), + Interp (t:=reify_type_of zeromod) + zero_gen limbwidth_num limbwidth_den s c n + = zeromod limbwidth_num limbwidth_den s c n) + As zero_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = zeromod _ _ _ _ _) => simple apply zero_gen_correct : reify_gen_cache. + +Derive one_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (s : Z) + (c : list (Z * Z)), + Interp (t:=reify_type_of onemod) + one_gen limbwidth_num limbwidth_den s c n + = onemod limbwidth_num limbwidth_den s c n) + As one_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = onemod _ _ _ _ _) => simple apply one_gen_correct : reify_gen_cache. + +Derive prime_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (s : Z) + (c : list (Z * Z)), + Interp (t:=reify_type_of primemod) + prime_gen limbwidth_num limbwidth_den s c n + = primemod limbwidth_num limbwidth_den s c n) + As prime_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = primemod _ _ _ _ _) => simple apply prime_gen_correct : reify_gen_cache. + +Derive id_gen + SuchThat (forall (ls : list Z), + Interp (t:=reify_type_of (@id (list Z))) + id_gen ls + = id ls) + As id_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = id _) => simple apply id_gen_correct : reify_gen_cache. + +Derive to_bytes_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (bitwidth : Z) + (m_enc : list Z) + (f : list Z), + Interp (t:=reify_type_of to_bytesmod) + to_bytes_gen limbwidth_num limbwidth_den n bitwidth m_enc f + = to_bytesmod limbwidth_num limbwidth_den n bitwidth m_enc f) + As to_bytes_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = to_bytesmod _ _ _ _ _ _) => simple apply to_bytes_gen_correct : reify_gen_cache. + +Derive from_bytes_gen + SuchThat (forall (limbwidth_num limbwidth_den : Z) + (n : nat) + (f : list Z), + Interp (t:=reify_type_of from_bytesmod) + from_bytes_gen limbwidth_num limbwidth_den n f + = from_bytesmod limbwidth_num limbwidth_den n f) + As from_bytes_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = from_bytesmod _ _ _ _) => simple apply from_bytes_gen_correct : reify_gen_cache. + +Derive mulx_gen + SuchThat (forall (s x y : Z), + Interp (t:=reify_type_of mulx) + mulx_gen s x y + = mulx s x y) + As mulx_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = mulx _ _ _) => simple apply mulx_gen_correct : reify_gen_cache. + +Derive addcarryx_gen + SuchThat (forall (s c x y : Z), + Interp (t:=reify_type_of addcarryx) + addcarryx_gen s c x y + = addcarryx s c x y) + As addcarryx_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = addcarryx _ _ _ _) => simple apply addcarryx_gen_correct : reify_gen_cache. + +Derive subborrowx_gen + SuchThat (forall (s c x y : Z), + Interp (t:=reify_type_of subborrowx) + subborrowx_gen s c x y + = subborrowx s c x y) + As subborrowx_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = subborrowx _ _ _ _) => simple apply subborrowx_gen_correct : reify_gen_cache. + +Derive cmovznz_gen + SuchThat (forall (bitwidth cond z nz : Z), + Interp (t:=reify_type_of cmovznz) + cmovznz_gen bitwidth cond z nz + = cmovznz bitwidth cond z nz) + As cmovznz_gen_correct. +Proof. cache_reify (). Qed. +Hint Extern 1 (_ = cmovznz _ _ _ _) => simple apply cmovznz_gen_correct : reify_gen_cache. + + (** XXX TODO: Translate Jade's python script *) Module Import UnsaturatedSolinas. Section rcarry_mul. @@ -813,42 +846,66 @@ Module Import UnsaturatedSolinas. Let limbwidth := (Z.log2_up (s - Associational.eval c) / Z.of_nat n)%Q. Let idxs := (seq 0 n ++ [0; 1])%list%nat. Let coef := 2. + Let n_bytes := bytes_n (Qnum limbwidth) (Qden limbwidth) n. + Let prime_upperbound_list : list Z + := encode (weight (Qnum limbwidth) (Qden limbwidth)) n s c (s-1). + Let prime_bytes_upperbound_list : list Z + := encode (weight 8 1) n_bytes s c (s-1). Let tight_upperbounds : list Z := List.map (fun v : Z => Qceiling (11/10 * v)) - (encode (weight (Qnum limbwidth) (Qden limbwidth)) n s c (s-1)). + prime_upperbound_list. Definition prime_bound : ZRange.type.option.interp (base.type.Z) := Some r[0~>(s - Associational.eval c - 1)]%zrange. + Definition prime_bytes_bounds : ZRange.type.option.interp (base.type.list (base.type.Z)) + := Some (List.map (fun v => Some r[0 ~> v]%zrange) prime_bytes_upperbound_list). + + Definition m_enc : list Z + := encode (weight (Qnum limbwidth) (Qden limbwidth)) n s c (s-Associational.eval c). Definition relax_zrange_of_machine_wordsize := relax_zrange_gen [machine_wordsize; 2 * machine_wordsize]%Z. + Definition relax_zrange_of_machine_wordsize_with_bytes + := relax_zrange_gen [1; 8; machine_wordsize; 2 * machine_wordsize]%Z. + Let relax_zrange := relax_zrange_of_machine_wordsize. + Let relax_zrange_with_bytes := relax_zrange_of_machine_wordsize_with_bytes. Definition tight_bounds : list (ZRange.type.option.interp base.type.Z) := List.map (fun u => Some r[0~>u]%zrange) tight_upperbounds. Definition loose_bounds : list (ZRange.type.option.interp base.type.Z) := List.map (fun u => Some r[0 ~> 3*u]%zrange) tight_upperbounds. + (** Note: If you change the name or type signature of this function, you will need to update the code in CLI.v *) Definition check_args {T} (res : Pipeline.ErrorT T) : Pipeline.ErrorT T - := if negb (Qle_bool 1 limbwidth)%Q - then Error (Pipeline.Value_not_leQ "1 ≤ limbwidth" 1%Q limbwidth) - else if (negb (0 if b:bool then Error e else k) + res + [(negb (Qle_bool 1 limbwidth)%Q, Pipeline.Value_not_leQ "limbwidth < 1" 1%Q limbwidth); + ((negb (0 d) _). - Notation BoundsPipeline rop in_bounds out_bounds - := (Pipeline.BoundsPipeline + Notation BoundsPipelineToStrings prefix name rop in_bounds out_bounds + := ((prefix ++ name)%string, + Pipeline.BoundsPipelineToStrings + prefix (prefix ++ name)%string (*false*) true None relax_zrange rop%Expr in_bounds out_bounds). @@ -867,9 +924,32 @@ Module Import UnsaturatedSolinas. Hrop rv) (only parsing). + Notation BoundsPipelineToStrings_with_bytes_no_subst01 prefix name rop in_bounds out_bounds + := ((prefix ++ name)%string, + Pipeline.BoundsPipelineToStrings + prefix (prefix ++ name)%string + (*false*) false None + relax_zrange_with_bytes + rop%Expr in_bounds out_bounds). + + Notation BoundsPipeline_with_bytes_no_subst01_correct in_bounds out_bounds op + := (fun rv (rop : Expr (reify_type_of op)) Hrop + => @Pipeline.BoundsPipeline_correct_trans + (*false*) false None + relax_zrange_with_bytes + (relax_zrange_gen_good _) + _ + rop + in_bounds + out_bounds + op + Hrop rv) + (only parsing). + (* N.B. We only need [rcarry_mul] if we want to extract the Pipeline; otherwise we can just use [rcarry_mul_correct] *) - Definition rcarry_mul - := BoundsPipeline + Definition srcarry_mul prefix + := BoundsPipelineToStrings + prefix "carry_mul" (carry_mul_gen @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify s @ GallinaReify.Reify c @ GallinaReify.Reify n @ GallinaReify.Reify idxs) (Some loose_bounds, (Some loose_bounds, tt)) @@ -881,8 +961,9 @@ Module Import UnsaturatedSolinas. (Some tight_bounds) (carry_mulmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) s c n idxs). - Definition rcarry - := BoundsPipeline + Definition srcarry prefix + := BoundsPipelineToStrings + prefix "carry" (carry_gen @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify s @ GallinaReify.Reify c @ GallinaReify.Reify n @ GallinaReify.Reify idxs) (Some loose_bounds, tt) @@ -894,8 +975,9 @@ Module Import UnsaturatedSolinas. (Some tight_bounds) (carrymod (Qnum limbwidth) (Z.pos (Qden limbwidth)) s c n idxs). - Definition rrelax - := BoundsPipeline + Definition srrelax prefix + := BoundsPipelineToStrings + prefix "relax" id_gen (Some tight_bounds, tt) (Some loose_bounds). @@ -906,8 +988,9 @@ Module Import UnsaturatedSolinas. (Some loose_bounds) (@id (list Z)). - Definition radd - := BoundsPipeline + Definition sradd prefix + := BoundsPipelineToStrings + prefix "add" (add_gen @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify n) (Some tight_bounds, (Some tight_bounds, tt)) @@ -919,8 +1002,9 @@ Module Import UnsaturatedSolinas. (Some loose_bounds) (addmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n). - Definition rsub - := BoundsPipeline + Definition srsub prefix + := BoundsPipelineToStrings + prefix "sub" (sub_gen @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify s @ GallinaReify.Reify c @ GallinaReify.Reify n @ GallinaReify.Reify coef) (Some tight_bounds, (Some tight_bounds, tt)) @@ -932,8 +1016,9 @@ Module Import UnsaturatedSolinas. (Some loose_bounds) (submod (Qnum limbwidth) (Z.pos (Qden limbwidth)) s c n coef). - Definition ropp - := BoundsPipeline + Definition sropp prefix + := BoundsPipelineToStrings + prefix "opp" (opp_gen @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify s @ GallinaReify.Reify c @ GallinaReify.Reify n @ GallinaReify.Reify coef) (Some tight_bounds, tt) @@ -945,6 +1030,34 @@ Module Import UnsaturatedSolinas. (Some loose_bounds) (oppmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) s c n coef). + Definition srto_bytes prefix + := BoundsPipelineToStrings_with_bytes_no_subst01 + prefix "to_bytes" + (to_bytes_gen + @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify n @ GallinaReify.Reify machine_wordsize @ GallinaReify.Reify m_enc) + (Some tight_bounds, tt) + prime_bytes_bounds. + + Definition rto_bytes_correct + := BoundsPipeline_with_bytes_no_subst01_correct + (Some tight_bounds, tt) + prime_bytes_bounds + (to_bytesmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n machine_wordsize m_enc). + + Definition srfrom_bytes prefix + := BoundsPipelineToStrings_with_bytes_no_subst01 + prefix "from_bytes" + (from_bytes_gen + @ GallinaReify.Reify (Qnum limbwidth) @ GallinaReify.Reify (Z.pos (Qden limbwidth)) @ GallinaReify.Reify n) + (prime_bytes_bounds, tt) + (Some tight_bounds). + + Definition rfrom_bytes_correct + := BoundsPipeline_with_bytes_no_subst01_correct + (prime_bytes_bounds, tt) + (Some tight_bounds) + (from_bytesmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n). + Definition rencode_correct := BoundsPipeline_correct (prime_bound, tt) @@ -963,6 +1076,62 @@ Module Import UnsaturatedSolinas. (Some tight_bounds) (onemod (Qnum limbwidth) (Z.pos (Qden limbwidth)) s c n). + Definition srmulx prefix (s : Z) + := BoundsPipelineToStrings_with_bytes_no_subst01 + prefix ("mulx_u" ++ decimal_string_of_Z s) + (mulx_gen + @ GallinaReify.Reify s) + (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt))%zrange + (Some r[0~>2^s-1], Some r[0~>2^s-1])%zrange. + + Definition srmulx_correct (s : Z) + := BoundsPipeline_with_bytes_no_subst01_correct + (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt))%zrange + (Some r[0~>2^s-1], Some r[0~>2^s-1])%zrange + (mulx s). + + Definition sraddcarryx prefix (s : Z) + := BoundsPipelineToStrings_with_bytes_no_subst01 + prefix ("addcarryx_u" ++ decimal_string_of_Z s) + (addcarryx_gen + @ GallinaReify.Reify s) + (Some r[0~>1], (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt)))%zrange + (Some r[0~>2^s-1], Some r[0~>1])%zrange. + + Definition sraddcarryx_correct (s : Z) + := BoundsPipeline_with_bytes_no_subst01_correct + (Some r[0~>1], (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt)))%zrange + (Some r[0~>2^s-1], Some r[0~>1])%zrange + (addcarryx s). + + Definition srsubborrowx prefix (s : Z) + := BoundsPipelineToStrings_with_bytes_no_subst01 + prefix ("subborrowx_u" ++ decimal_string_of_Z s) + (subborrowx_gen + @ GallinaReify.Reify s) + (Some r[0~>1], (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt)))%zrange + (Some r[0~>2^s-1], Some r[0~>1])%zrange. + + Definition srsubborrowx_correct (s : Z) + := BoundsPipeline_with_bytes_no_subst01_correct + (Some r[0~>1], (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt)))%zrange + (Some r[0~>2^s-1], Some r[0~>1])%zrange + (subborrowx s). + + Definition srcmovznz prefix (s : Z) + := BoundsPipelineToStrings_with_bytes_no_subst01 + prefix ("cmovznz_u" ++ decimal_string_of_Z s) + (cmovznz_gen + @ GallinaReify.Reify s) + (Some r[0~>1], (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt)))%zrange + (Some r[0~>2^s-1])%zrange. + + Definition srcmovznz_correct (s : Z) + := BoundsPipeline_with_bytes_no_subst01_correct + (Some r[0~>1], (Some r[0~>2^s-1], (Some r[0~>2^s-1], tt)))%zrange + (Some r[0~>2^s-1])%zrange + (cmovznz s). + (* we need to strip off [Hrv : ... = Pipeline.Success rv] and related arguments *) Definition rcarry_mul_correctT rv : Prop := type_of_strip_3arrow (@rcarry_mul_correct rv). @@ -976,6 +1145,10 @@ Module Import UnsaturatedSolinas. := type_of_strip_3arrow (@rsub_correct rv). Definition ropp_correctT rv : Prop := type_of_strip_3arrow (@ropp_correct rv). + Definition rto_bytes_correctT rv : Prop + := type_of_strip_3arrow (@rto_bytes_correct rv). + Definition rfrom_bytes_correctT rv : Prop + := type_of_strip_3arrow (@rfrom_bytes_correct rv). Definition rencode_correctT rv : Prop := type_of_strip_3arrow (@rencode_correct rv). Definition rzero_correctT rv : Prop @@ -994,10 +1167,13 @@ Module Import UnsaturatedSolinas. {roppv} (Hroppv : ropp_correctT roppv) {rzerov} (Hrzerov : rzero_correctT rzerov) {ronev} (Hronev : rone_correctT ronev) - {rencodev} (Hrencodev : rencode_correctT rencodev). + {rencodev} (Hrencodev : rencode_correctT rencodev) + {rto_bytesv} (Hto_bytesv : rto_bytes_correctT rto_bytesv) + {rfrom_bytesv} (Hfrom_bytesv : rfrom_bytes_correctT rfrom_bytesv). Local Ltac use_curve_good_t := - repeat first [ progress rewrite ?map_length, ?Z.mul_0_r, ?Pos.mul_1_r, ?Z.mul_1_r in * + repeat first [ assumption + | progress rewrite ?map_length, ?Z.mul_0_r, ?Pos.mul_1_r, ?Z.mul_1_r in * | reflexivity | lia | rewrite interp_reify_list, ?map_map @@ -1018,11 +1194,16 @@ Module Import UnsaturatedSolinas. /\ n <> 0%nat /\ List.length tight_bounds = n /\ List.length loose_bounds = n - /\ 0 < Qden limbwidth <= Qnum limbwidth. + /\ 0 < Qden limbwidth <= Qnum limbwidth + /\ s = weight (Qnum limbwidth) (QDen limbwidth) n + /\ map (Z.land (Z.ones machine_wordsize)) m_enc = m_enc + /\ eval (weight (Qnum limbwidth) (QDen limbwidth)) n m_enc = s - Associational.eval c + /\ Datatypes.length m_enc = n + /\ 0 < Associational.eval c < s. Proof. clear -curve_good. - cbv [check_args] in curve_good. - cbv [tight_bounds loose_bounds prime_bound] in *. + cbv [check_args fold_right] in curve_good. + cbv [tight_bounds loose_bounds prime_bound m_enc] in *. break_innermost_match_hyps; try discriminate. rewrite negb_false_iff in *. Z.ltb_to_lt. @@ -1033,8 +1214,17 @@ Module Import UnsaturatedSolinas. rewrite ?map_length, ?Z.mul_0_r, ?Pos.mul_1_r, ?Z.mul_1_r in *. specialize_by lia. repeat match goal with H := _ |- _ => subst H end. + repeat match goal with + | [ H : list_beq _ _ _ _ = true |- _ ] => apply internal_list_dec_bl in H; [ | intros; Z.ltb_to_lt; omega.. ] + end. repeat apply conj. - { destruct (s - Associational.eval c); cbn; lia. } + { destruct (s - Associational.eval c) eqn:?; cbn; lia. } + { use_curve_good_t. } + { use_curve_good_t. } + { use_curve_good_t. } + { use_curve_good_t. } + { use_curve_good_t. } + { use_curve_good_t. } { use_curve_good_t. } { use_curve_good_t. } { use_curve_good_t. } @@ -1046,6 +1236,7 @@ Module Import UnsaturatedSolinas. { use_curve_good_t. } Qed. + (** TODO: Find a better place to put the spec for [to_bytes] *) Definition GoodT : Prop := @Ring.GoodT (Qnum limbwidth) @@ -1060,59 +1251,155 @@ Module Import UnsaturatedSolinas. (Interp roppv) (Interp rzerov) (Interp ronev) - (Interp rencodev). + (Interp rencodev) + /\ (let to_bytesT := (base.type.list base.type.Z -> base.type.list base.type.Z)%etype in + forall f + (Hf : type.andb_bool_for_each_lhs_of_arrow (t:=to_bytesT) (@ZRange.type.option.is_bounded_by) (Some tight_bounds, tt) f = true), + ((ZRange.type.base.option.is_bounded_by prime_bytes_bounds (type.app_curried (Interp rto_bytesv) f) = true + /\ (forall cast_outside_of_range, type.app_curried (expr.Interp (@ident.gen_interp cast_outside_of_range) rto_bytesv) f + = type.app_curried (t:=to_bytesT) (to_bytesmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n machine_wordsize m_enc) f)) + /\ (Positional.eval (weight 8 1) n_bytes (type.app_curried (t:=to_bytesT) (to_bytesmod (Qnum limbwidth) (Z.pos (Qden limbwidth)) n machine_wordsize m_enc) f)) = (Positional.eval (weight (Qnum limbwidth) (Z.pos (Qden limbwidth))) n (fst f) mod m))). Theorem Good : GoodT. Proof. pose proof use_curve_good; destruct_head'_and; destruct_head_hnf' ex. - eapply Ring.Good; - lazymatch goal with - | [ H : ?P ?rop |- context[expr.Interp _ ?rop] ] - => intros; - let H1 := fresh in - let H2 := fresh in - unshelve edestruct H as [H1 H2]; [ .. | solve [ split; [ eapply H1 | eapply H2 ] ] ]; - solve [ exact tt | eassumption | reflexivity ] - | _ => idtac - end; - repeat first [ assumption - | intros; apply eval_carry_mulmod - | intros; apply eval_carrymod - | intros; apply eval_addmod - | intros; apply eval_submod - | intros; apply eval_oppmod - | intros; apply eval_encodemod - | apply conj ]. - Qed. + split. + { eapply Ring.Good; + lazymatch goal with + | [ H : ?P ?rop |- context[expr.Interp _ ?rop] ] + => intros; + let H1 := fresh in + let H2 := fresh in + unshelve edestruct H as [H1 H2]; [ .. | solve [ split; [ eapply H1 | eapply H2 ] ] ]; + solve [ exact tt | eassumption | reflexivity ] + | _ => idtac + end; + repeat first [ assumption + | intros; apply eval_carry_mulmod + | intros; apply eval_carrymod + | intros; apply eval_addmod + | intros; apply eval_submod + | intros; apply eval_oppmod + | intros; apply eval_encodemod + | apply conj ]. } + { cbv zeta; intros f Hf; split. + { apply Hto_bytesv; assumption. } + { cbn [type.for_each_lhs_of_arrow type_base type.andb_bool_for_each_lhs_of_arrow ZRange.type.option.is_bounded_by fst snd] in *. + rewrite Bool.andb_true_iff in *; split_and'. + etransitivity; [ apply eval_to_bytesmod | f_equal; (eassumption || (symmetry; eassumption)) ]; + auto; try omega. + { erewrite Ring.length_is_bounded_by by eassumption; assumption. } + { lazymatch goal with + | [ H : eval _ _ _ = ?x |- _ <= _ < 2 * ?x ] => rewrite <- H + end. + cbv [m_enc tight_bounds tight_upperbounds prime_upperbound_list] in H15 |- *. + Lemma fold_andb_map_snoc A B f x xs y ys + : @fold_andb_map A B f (xs ++ [x]) (ys ++ [y]) = @fold_andb_map A B f xs ys && f x y. + Proof. + clear. + revert ys; induction xs as [|x' xs'], ys as [|y' ys']; cbn; + rewrite ?Bool.andb_true_r, ?Bool.andb_false_r; + try (destruct ys'; cbn; rewrite Bool.andb_false_r); + try (destruct xs'; cbn; rewrite Bool.andb_false_r); + try reflexivity. + rewrite IHxs', Bool.andb_assoc; reflexivity. + Qed. + Lemma eval_is_bounded_by wt n' bounds bounds' f + (H : ZRange.type.base.option.is_bounded_by (t:=base.type.list base.type.Z) bounds f = true) + (Hb : bounds = Some (List.map (@Some _) bounds')) + (Hblen : length bounds' = n') + (Hwt : forall i, List.In i (seq 0 n') -> 0 <= wt i) + : eval wt n' (List.map lower bounds') <= eval wt n' f <= eval wt n' (List.map upper bounds'). + Proof. + clear -H Hb Hblen Hwt. + setoid_rewrite in_seq in Hwt. + subst bounds. + pose proof H as H'; apply fold_andb_map_length in H'. + revert dependent bounds'; intro bounds'. + revert dependent f; intro f. + rewrite <- (List.rev_involutive bounds'), <- (List.rev_involutive f); + generalize (List.rev bounds') (List.rev f); clear bounds' f; intros bounds f; revert bounds f. + induction n' as [|n IHn], bounds as [|b bounds], f as [|f fs]; intros; + cbn [length rev map] in *; distr_length. + { rewrite !map_app in *; cbn [map] in *. + erewrite !eval_snoc by (distr_length; eauto). + cbn [ZRange.type.base.option.is_bounded_by ZRange.type.base.is_bounded_by] in *. + specialize_by (intros; auto with omega). + specialize (Hwt n); specialize_by omega. + repeat first [ progress Bool.split_andb + | rewrite Nat.add_1_r in * + | rewrite fold_andb_map_snoc in * + | rewrite Nat.succ_inj_wd in * + | progress Z.ltb_to_lt + | progress cbn [In seq] in * + | match goal with + | [ H : length _ = ?v |- _ ] => rewrite H in * + | [ H : ?v = length _ |- _ ] => rewrite <- H in * + end ]. + split; apply Z.add_le_mono; try apply IHn; auto; distr_length; nia. } + Qed. + eapply eval_is_bounded_by with (wt:=weight (Qnum limbwidth) (QDen limbwidth)) in H15. + 2:rewrite <- (map_map _ (@Some _)); reflexivity. + 2:distr_length; reflexivity. + rewrite ?map_map in *. + cbn [lower upper] in *. + split. + { etransitivity; [ erewrite <- eval_zeros; [ | apply weight_0, wprops ] | apply H15 ]. + apply Z.eq_le_incl; f_equal. + admit. + omega. } + { eapply Z.le_lt_trans; [ apply H15 | ]. + assert (Hlen : length (encode (weight (Qnum limbwidth) (QDen limbwidth)) n s c (s - 1)) = n) by distr_length. + revert Hlen. + generalize ((encode (weight (Qnum limbwidth) (QDen limbwidth)) n s c (s - 1))). + intro ls. + clear. + revert ls. + clearbody limbwidth. + induction n as [|n' IHn'], ls as [|l ls]; cbn [length]; intros; try omega. + admit. + cbn [map]. + admit. } + admit. } + Admitted. End make_ring. Section for_stringification. - Local Open Scope string_scope. - - Let ToFunLines t name E arg_bounds - := (name, - match E with - | Success E' - => let E := @ToString.C.ToFunctionLines - name t E' None arg_bounds in - match E with - | Some E => Success E - | None => Error (Pipeline.Stringification_failed E') - end - | Error err => Error err - end). + Definition aggregate_infos {A B C} (ls : list (A * ErrorT B (C * ToString.C.ident_infos))) : ToString.C.ident_infos + := fold_right + ToString.C.ident_info_union + ToString.C.ident_info_empty + (List.map + (fun '(_, res) => match res with + | Success (_, infos) => infos + | Error _ => ToString.C.ident_info_empty + end) + ls). + + Definition extra_synthesis (function_name_prefix : string) (infos : ToString.C.ident_infos) + : list (string * Pipeline.ErrorT (list string)) * PositiveSet.t + := let ls_addcarryx := List.flat_map + (fun lg_split:positive => [sraddcarryx function_name_prefix lg_split; srsubborrowx function_name_prefix lg_split]) + (PositiveSet.elements (ToString.C.addcarryx_lg_splits infos)) in + let ls_mulx := List.map + (fun lg_split:positive => srmulx function_name_prefix lg_split) + (PositiveSet.elements (ToString.C.mulx_lg_splits infos)) in + let ls_cmov := List.map + (fun bitwidth:positive => srcmovznz function_name_prefix bitwidth) + (PositiveSet.elements (ToString.C.cmovznz_bitwidths infos)) in + let ls := ls_addcarryx ++ ls_mulx ++ ls_cmov in + let infos := aggregate_infos ls in + (List.map (fun '(name, res) => (name, (res <- res; Success (fst res))%error)) ls, + ToString.C.bitwidths_used infos). (** Note: If you change the name or type signature of this function, you will need to update the code in CLI.v *) - Definition Synthesize (function_name_prefix : string) : list (string * Pipeline.ErrorT (list string)) - := let loose_bounds := Some loose_bounds in - let tight_bounds := Some tight_bounds in - let fe op := (function_name_prefix ++ op)%string in - [(ToFunLines _ (fe "carry_mul") rcarry_mul (loose_bounds, (loose_bounds, tt))); - (ToFunLines _ (fe "carry") rcarry (loose_bounds, tt)); - (ToFunLines _ (fe "add") radd (tight_bounds, (tight_bounds, tt))); - (ToFunLines _ (fe "sub") rsub (tight_bounds, (tight_bounds, tt))); - (ToFunLines _ (fe "opp") ropp (tight_bounds, tt))]. + Definition Synthesize (function_name_prefix : string) : list (string * Pipeline.ErrorT (list string)) * PositiveSet.t (* types used *) + := let ls := List.map (fun sr => sr function_name_prefix) [srcarry_mul; srcarry; sradd; srsub; sropp; srto_bytes; srfrom_bytes] in + let infos := aggregate_infos ls in + let '(extra_ls, extra_bit_widths) := extra_synthesis function_name_prefix infos in + (extra_ls ++ List.map (fun '(name, res) => (name, (res <- res; Success (fst res))%error)) ls, + PositiveSet.union extra_bit_widths (ToString.C.bitwidths_used infos)). End for_stringification. End rcarry_mul. End UnsaturatedSolinas. @@ -1193,6 +1480,8 @@ Ltac solve_rcarry := solve_rop rcarry_correct. Ltac solve_radd := solve_rop radd_correct. Ltac solve_rsub := solve_rop rsub_correct. Ltac solve_ropp := solve_rop ropp_correct. +Ltac solve_rto_bytes := solve_rop rto_bytes_correct. +Ltac solve_rfrom_bytes := solve_rop rfrom_bytes_correct. Ltac solve_rencode := solve_rop rencode_correct. Ltac solve_rrelax := solve_rop rrelax_correct. Ltac solve_rzero := solve_rop rzero_correct. @@ -1239,16 +1528,20 @@ Module PrintingNotations. Notation "x -₃₂ y" := (#(ident.Z_cast uint32) @ (#ident.Z_sub @ x @ y))%expr (at level 50) : expr_scope. Notation "( out_t )( v >> count )" - := ((#(ident.Z_cast out_t) @ (#(ident.Z_shiftr count) @ v))%expr) + := ((#(ident.Z_cast out_t) @ (#ident.Z_shiftr @ v @ count))%expr) (format "( out_t )( v >> count )") : expr_scope. Notation "( out_t )( v << count )" - := ((#(ident.Z_cast out_t) @ (#(ident.Z_shiftl count) @ v))%expr) + := ((#(ident.Z_cast out_t) @ (#ident.Z_shiftl @ v @ count))%expr) (format "( out_t )( v << count )") : expr_scope. Notation "( range )( v )" := ((#(ident.Z_cast range) @ $v)%expr) (format "( range )( v )") : expr_scope. + Notation "( mask & ( out_t )( v ) )" + := ((#(ident.Z_cast out_t) @ (#ident.Z_land @ #(ident.Literal (t:=base.type.Z) mask) @ v))%expr) + (format "( mask & ( out_t )( v ) )") + : expr_scope. Notation "( ( out_t )( v ) & mask )" - := ((#(ident.Z_cast out_t) @ (#(ident.Z_land mask) @ v))%expr) + := ((#(ident.Z_cast out_t) @ (#ident.Z_land @ v @ #(ident.Literal (t:=base.type.Z) mask)))%expr) (format "( ( out_t )( v ) & mask )") : expr_scope. @@ -1327,17 +1620,17 @@ Module PrintingNotations. (* TODO: come up with a better notation for arithmetic with carries that still distinguishes it from arithmetic without carries? *) Local Notation "'TwoPow256'" := 115792089237316195423570985008687907853269984665640564039457584007913129639936 (only parsing). - Notation "'ADD_256' ( x , y )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#(ident.Z_add_get_carry_concrete TwoPow256) @ x @ y))%expr : expr_scope. - Notation "'ADD_128' ( x , y )" := (#(ident.Z_cast2 (uint128, bool)%core) @ (#(ident.Z_add_get_carry_concrete TwoPow256) @ x @ y))%expr : expr_scope. - Notation "'ADDC_256' ( x , y , z )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#(ident.Z_add_with_get_carry_concrete TwoPow256) @ x @ y @ z))%expr : expr_scope. - Notation "'ADDC_128' ( x , y , z )" := (#(ident.Z_cast2 (uint128, bool)%core) @ (#(ident.Z_add_with_get_carry_concrete TwoPow256) @ x @ y @ z))%expr : expr_scope. - Notation "'SUB_256' ( x , y )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#(ident.Z_sub_get_borrow_concrete TwoPow256) @ x @ y))%expr : expr_scope. - Notation "'SUBB_256' ( x , y , z )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#(ident.Z_sub_with_get_borrow_concrete TwoPow256) @ x @ y @ z))%expr : expr_scope. + Notation "'ADD_256' ( x , y )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#ident.Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) TwoPow256) @ x @ y))%expr : expr_scope. + Notation "'ADD_128' ( x , y )" := (#(ident.Z_cast2 (uint128, bool)%core) @ (#ident.Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) TwoPow256) @ x @ y))%expr : expr_scope. + Notation "'ADDC_256' ( x , y , z )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#ident.Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) TwoPow256) @ x @ y @ z))%expr : expr_scope. + Notation "'ADDC_128' ( x , y , z )" := (#(ident.Z_cast2 (uint128, bool)%core) @ (#ident.Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) TwoPow256) @ x @ y @ z))%expr : expr_scope. + Notation "'SUB_256' ( x , y )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#ident.Z_sub_get_borrow @ #(ident.Literal (t:=base.type.Z) TwoPow256) @ x @ y))%expr : expr_scope. + Notation "'SUBB_256' ( x , y , z )" := (#(ident.Z_cast2 (uint256, bool)%core) @ (#ident.Z_sub_with_get_borrow @ #(ident.Literal (t:=base.type.Z) TwoPow256) @ x @ y @ z))%expr : expr_scope. Notation "'ADDM' ( x , y , z )" := (#(ident.Z_cast uint256) @ (#ident.Z_add_modulo @ x @ y @ z))%expr : expr_scope. - Notation "'RSHI' ( x , y , z )" := (#(ident.Z_cast _) @ (#(ident.Z_rshi_concrete _ z) @ x @ y))%expr : expr_scope. + Notation "'RSHI' ( x , y , z )" := (#(ident.Z_cast _) @ (#ident.Z_rshi @ _ @ x @ y @ z))%expr : expr_scope. Notation "'SELC' ( x , y , z )" := (#(ident.Z_cast uint256) @ (ident.Z_zselect @ x @ y @ z))%expr : expr_scope. - Notation "'SELM' ( x , y , z )" := (#(ident.Z_cast uint256) @ (ident.Z_zselect @ (#(Z_cast bool) @ (Z_cc_m_concrete _) @ x) @ y @ z))%expr : expr_scope. - Notation "'SELL' ( x , y , z )" := (#(ident.Z_cast uint256) @ (#ident.Z_zselect @ (#(Z_cast bool) @ (#(Z_land 1) @ x)) @ y @ z))%expr : expr_scope. + Notation "'SELM' ( x , y , z )" := (#(ident.Z_cast uint256) @ (ident.Z_zselect @ (#(Z_cast bool) @ (#Z_cc_m @ _) @ x) @ y @ z))%expr : expr_scope. + Notation "'SELL' ( x , y , z )" := (#(ident.Z_cast uint256) @ (#ident.Z_zselect @ (#(Z_cast bool) @ (#Z_land @ #(ident.Literal (t:=base.type.Z 1)) @ x)) @ y @ z))%expr : expr_scope. End PrintingNotations. (* @@ -1459,8 +1752,10 @@ Module SaturatedSolinas. then Error (Pipeline.Value_not_ltZ "0 < machine_wordsize" 0 machine_wordsize) else res. - Notation BoundsPipeline rop in_bounds out_bounds - := (Pipeline.BoundsPipeline + Notation BoundsPipelineToStrings prefix name rop in_bounds out_bounds + := ((prefix ++ name)%string, + Pipeline.BoundsPipelineToStrings + prefix (prefix ++ name)%string (*false*) false None relax_zrange rop%Expr in_bounds out_bounds). @@ -1485,8 +1780,9 @@ Module SaturatedSolinas. (Some boundsn) (mulmod' s c machine_wordsize n nreductions). - Definition rmulmod - := BoundsPipeline + Definition srmulmod prefix + := BoundsPipelineToStrings + prefix "mulmod" (mulmod_gen @ GallinaReify.Reify s @ GallinaReify.Reify c @ GallinaReify.Reify machine_wordsize @ GallinaReify.Reify n @ GallinaReify.Reify nreductions) (Some boundsn, (Some boundsn, tt)) (Some boundsn). @@ -1496,28 +1792,41 @@ Module SaturatedSolinas. := type_of_strip_3arrow (@rmulmod_correct rv). Section for_stringification. - Local Open Scope string_scope. - - Let ToFunLines t name E arg_bounds - := (name, - match E with - | Success E' - => let E := @ToString.C.ToFunctionLines - name t E' None arg_bounds in - match E with - | Some E => Success E - | None => Error (Pipeline.Stringification_failed E') - end - | Error err => Error err - end). + Definition aggregate_infos {A B C} (ls : list (A * ErrorT B (C * ToString.C.ident_infos))) : ToString.C.ident_infos + := fold_right + ToString.C.ident_info_union + ToString.C.ident_info_empty + (List.map + (fun '(_, res) => match res with + | Success (_, infos) => infos + | Error _ => ToString.C.ident_info_empty + end) + ls). + + Definition extra_synthesis (function_name_prefix : string) (infos : ToString.C.ident_infos) + : list (string * Pipeline.ErrorT (list string)) * PositiveSet.t + := let ls_addcarryx := List.flat_map + (fun lg_split:positive => [sraddcarryx machine_wordsize function_name_prefix lg_split; srsubborrowx machine_wordsize function_name_prefix lg_split]) + (PositiveSet.elements (ToString.C.addcarryx_lg_splits infos)) in + let ls_mulx := List.map + (fun lg_split:positive => srmulx machine_wordsize function_name_prefix lg_split) + (PositiveSet.elements (ToString.C.mulx_lg_splits infos)) in + let ls_cmov := List.map + (fun bitwidth:positive => srcmovznz machine_wordsize function_name_prefix bitwidth) + (PositiveSet.elements (ToString.C.cmovznz_bitwidths infos)) in + let ls := ls_addcarryx ++ ls_mulx ++ ls_cmov in + let infos := aggregate_infos ls in + (List.map (fun '(name, res) => (name, (res <- res; Success (fst res))%error)) ls, + ToString.C.bitwidths_used infos). (** Note: If you change the name or type signature of this function, you will need to update the code in CLI.v *) - Definition Synthesize (function_name_prefix : string) : list (string * Pipeline.ErrorT (list string)) - := let loose_bounds := Some loose_bounds in - let tight_bounds := Some tight_bounds in - let fe op := (function_name_prefix ++ op)%string in - [(ToFunLines _ (fe "mulmod") rmulmod (Some boundsn, (Some boundsn, tt)))]. + Definition Synthesize (function_name_prefix : string) : list (string * Pipeline.ErrorT (list string)) * PositiveSet.t (* types used *) + := let ls := List.map (fun sr => sr function_name_prefix) [srmulmod] in + let infos := aggregate_infos ls in + let '(extra_ls, extra_bit_widths) := extra_synthesis function_name_prefix infos in + (extra_ls ++ List.map (fun '(name, res) => (name, (res <- res; Success (fst res))%error)) ls, + PositiveSet.union extra_bit_widths (ToString.C.bitwidths_used infos)). End for_stringification. End rmulmod. End SaturatedSolinas. @@ -2318,3 +2627,64 @@ Time Compute exact r) (None, (Some (repeat (@None _) 5), tt)) ZRange.type.base.option.None). + +Compute + (Pipeline.BoundsPipeline + true None (relax_zrange_gen [64; 128]) + ltac:(let r := Reify (fun f => carry_mulmod 51 1 (2^255) [(1,19)] 5 (seq 0 5 ++ [0; 1])%list%nat f f) in + exact r) + (Some (repeat (@None _) 5), tt) + ZRange.type.base.option.None). + +Compute + (Pipeline.BoundsPipelineToString + "fiat_" "fiat_mulx_u64" + true None (relax_zrange_gen [64; 128]) + ltac:(let r := Reify (mulx 64) in + exact r) + (Some r[0~>2^64-1], (Some r[0~>2^64-1], tt))%zrange + (Some r[0~>2^64-1], Some r[0~>2^64-1])%zrange). + +Compute + (Pipeline.BoundsPipelineToString + "fiat_" "fiat_addcarryx_u64" + true None (relax_zrange_gen [1; 64; 128]) + ltac:(let r := Reify (addcarryx 64) in + exact r) + (Some r[0~>1], (Some r[0~>2^64-1], (Some r[0~>2^64-1], tt)))%zrange + (Some r[0~>2^64-1], Some r[0~>1])%zrange). + +Compute + (Pipeline.BoundsPipelineToString + "fiat_" "fiat_addcarryx_u51" + true None (relax_zrange_gen [1; 64; 128]) + ltac:(let r := Reify (addcarryx 51) in + exact r) + (Some r[0~>1], (Some r[0~>2^51-1], (Some r[0~>2^51-1], tt)))%zrange + (Some r[0~>2^51-1], Some r[0~>1])%zrange). + +Compute + (Pipeline.BoundsPipelineToString + "fiat_" "fiat_subborrowx_u64" + true None (relax_zrange_gen [1; 64; 128]) + ltac:(let r := Reify (subborrowx 64) in + exact r) + (Some r[0~>1], (Some r[0~>2^64-1], (Some r[0~>2^64-1], tt)))%zrange + (Some r[0~>2^64-1], Some r[0~>1])%zrange). +Compute + (Pipeline.BoundsPipelineToString + "fiat_" "fiat_subborrowx_u51" + true None (relax_zrange_gen [1; 64; 128]) + ltac:(let r := Reify (subborrowx 51) in + exact r) + (Some r[0~>1], (Some r[0~>2^51-1], (Some r[0~>2^51-1], tt)))%zrange + (Some r[0~>2^51-1], Some r[0~>1])%zrange). + +Compute + (Pipeline.BoundsPipelineToString + "fiat_" "fiat_cmovznz64" + true None (relax_zrange_gen [1; 64; 128]) + ltac:(let r := Reify (cmovznz 64) in + exact r) + (Some r[0~>1], (Some r[0~>2^64-1], (Some r[0~>2^64-1], tt)))%zrange + (Some r[0~>2^64-1])%zrange). diff --git a/src/Experiments/NewPipeline/Toplevel2.v b/src/Experiments/NewPipeline/Toplevel2.v index 34a758d87..35c25f915 100644 --- a/src/Experiments/NewPipeline/Toplevel2.v +++ b/src/Experiments/NewPipeline/Toplevel2.v @@ -110,6 +110,14 @@ Module X25519_64. SuchThat (ropp_correctT n s c machine_wordsize base_51_opp) As base_51_opp_correct. Proof. Time solve_ropp machine_wordsize. Time Qed. + Derive base_51_to_bytes + SuchThat (rto_bytes_correctT n s c machine_wordsize base_51_to_bytes) + As base_51_to_bytes_correct. + Proof. Time solve_rto_bytes machine_wordsize. Time Qed. + Derive base_51_from_bytes + SuchThat (rfrom_bytes_correctT n s c machine_wordsize base_51_from_bytes) + As base_51_from_bytes_correct. + Proof. Time solve_rfrom_bytes machine_wordsize. Time Qed. Derive base_51_encode SuchThat (rencode_correctT n s c machine_wordsize base_51_encode) As base_51_encode_correct. @@ -126,7 +134,7 @@ Module X25519_64. : check_args n s c machine_wordsize (Success tt) = Success tt. Proof. vm_compute; reflexivity. Qed. - Definition base_51_good : GoodT n s c + Definition base_51_good : GoodT n s c machine_wordsize := Good n s c machine_wordsize base_51_curve_good base_51_carry_mul_correct @@ -137,12 +145,16 @@ Module X25519_64. base_51_opp_correct base_51_zero_correct base_51_one_correct - base_51_encode_correct. + base_51_encode_correct + base_51_to_bytes_correct + base_51_from_bytes_correct. Print Assumptions base_51_good. Import PrintingNotations. Set Printing Width 80. Open Scope string_scope. + Local Notation prime_bytes_bounds := (prime_bytes_bounds n s c). + Print base_51_to_bytes. Print base_51_carry_mul. (*base_51_carry_mul = fun var : type -> Type => @@ -207,7 +219,7 @@ fun var : type -> Type => *) Compute ToString.C.ToFunctionString - "fecarry_mul" base_51_carry_mul + "" "fecarry_mul" base_51_carry_mul None (Some loose_bounds, (Some loose_bounds, tt)). (* void fecarry_mul(uint64_t[5] x1, uint64_t[5] x2, uint64_t[5] x3) { @@ -226,7 +238,7 @@ void fecarry_mul(uint64_t[5] x1, uint64_t[5] x2, uint64_t[5] x3) { } *) Compute ToString.C.ToFunctionString - "fesub" base_51_sub + "" "fesub" base_51_sub None (Some tight_bounds, (Some tight_bounds, tt)). (* void fesub(uint64_t[5] x1, uint64_t[5] x2, uint64_t[5] x3) { @@ -255,27 +267,27 @@ Module P224_64. Set Printing Depth 100000. Local Notation "'mul128' '(' x ',' y ')'" := - (#(Z_cast2 (uint128, _)%core) @ (#(Z_mul_split_concrete 340282366920938463463374607431768211456) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint128, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 340282366920938463463374607431768211456) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add128' '(' x ',' y ')'" := - (#(Z_cast2 (uint128, bool)%core) @ (#(Z_add_get_carry_concrete 340282366920938463463374607431768211456) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint128, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 340282366920938463463374607431768211456) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc128' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint128, bool)%core) @ (#(Z_add_with_get_carry_concrete 340282366920938463463374607431768211456) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint128, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 340282366920938463463374607431768211456) @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'sub128' '(' x ',' y ')'" := - (#(Z_cast2 (uint128, bool)%core) @ (#(Z_sub_get_borrow_concrete 340282366920938463463374607431768211456) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint128, bool)%core) @ (#Z_sub_get_borrow @ #(ident.Literal (t:=base.type.Z) 340282366920938463463374607431768211456) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'sbb128' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint128, bool)%core) @ (#(Z_sub_with_get_borrow_concrete 340282366920938463463374607431768211456) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint128, bool)%core) @ (#Z_sub_with_get_borrow @ #(ident.Literal (t:=base.type.Z) 340282366920938463463374607431768211456) @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'mul64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, _)%core) @ (#(Z_mul_split_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_add_get_carry_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc64' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_add_with_get_carry_concrete 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adx64' '(' c ',' x ',' y ')'" := (#(Z_cast bool) @ (#Z_add_with_carry @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'sub64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_sub_get_borrow_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_sub_get_borrow @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'sbb64' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_sub_with_get_borrow_concrete 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_sub_with_get_borrow @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. Set Printing Width 1000000. Print mulmod. End P224_64. @@ -296,11 +308,11 @@ Module P192_64. Set Printing Depth 100000. Local Notation "'mul64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, _)%core) @ (#(Z_mul_split_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, _)%core) @ (#Z_mul_split @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'add64' '(' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_add_get_carry_concrete 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_add_get_carry @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adc64' '(' c ',' x ',' y ')'" := - (#(Z_cast2 (uint64, bool)%core) @ (#(Z_add_with_get_carry_concrete 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. + (#(Z_cast2 (uint64, bool)%core) @ (#Z_add_with_get_carry @ #(ident.Literal (t:=base.type.Z) 18446744073709551616) @ c @ x @ y))%expr (at level 50) : expr_scope. Local Notation "'adx64' '(' c ',' x ',' y ')'" := (#(Z_cast bool) @ (#Z_add_with_carry @ c @ x @ y))%expr (at level 50) : expr_scope. -- cgit v1.2.3